diff --git a/theories/heap_lang/adequacy.v b/theories/heap_lang/adequacy.v
index 04d705959f50d321fda4e5a1598e24e9416495f1..45dc7c76b1988acd370cf7c317dea999b2fc3391 100644
--- a/theories/heap_lang/adequacy.v
+++ b/theories/heap_lang/adequacy.v
@@ -22,6 +22,6 @@ Proof.
   iMod (gen_heap_init σ.(heap)) as (?) "Hh".
   iMod (proph_map_init κs σ.(used_proph)) as (?) "Hp".
   iModIntro.
-  iExists (fun σ κs => (gen_heap_ctx σ.(heap) ∗ proph_map_ctx κs σ.(used_proph))%I). iFrame.
+  iExists (λ σ κs, (gen_heap_ctx σ.(heap) ∗ proph_map_ctx κs σ.(used_proph))%I). iFrame.
   iApply (Hwp (HeapG _ _ _ _)).
 Qed.
diff --git a/theories/heap_lang/lang.v b/theories/heap_lang/lang.v
index 77f96f1bfd12d6b5166fb0df52580526995828f1..168df652407deee50e9c7ce89aebaa9386ac40ea 100644
--- a/theories/heap_lang/lang.v
+++ b/theories/heap_lang/lang.v
@@ -176,7 +176,6 @@ Record state : Type := {
   heap: gmap loc val;
   used_proph: gset proph;
 }.
-Implicit Type σ : state.
 
 (** Equality and other typeclass stuff *)
 Lemma to_of_val v : to_val (of_val v) = Some v.
diff --git a/theories/heap_lang/lib/coin_flip.v b/theories/heap_lang/lib/coin_flip.v
index 804f25440dbf1b362132fd1ded5e5d9052e6e46d..a490ca0b24d77c16d3cfe9796ffb6bf7e19e3752 100644
--- a/theories/heap_lang/lib/coin_flip.v
+++ b/theories/heap_lang/lib/coin_flip.v
@@ -1,9 +1,7 @@
-From iris.heap_lang Require Export lifting notation.
 From iris.base_logic.lib Require Export invariants.
 From iris.program_logic Require Export atomic.
 From iris.proofmode Require Import tactics.
 From iris.heap_lang Require Import proofmode notation par.
-From iris.bi.lib Require Import fractional.
 Set Default Proof Using "Type".
 
 (** Nondeterminism and Speculation:
@@ -11,31 +9,33 @@ Set Default Proof Using "Type".
     Logical Relations for Fine-Grained Concurrency by Turon et al. (POPL'13) *)
 
 Definition rand: val :=
-    λ: "_",
-       let: "y" := ref #false in
-       Fork ("y" <- #true) ;;
-       !"y".
+  λ: "_",
+    let: "y" := ref #false in
+    Fork ("y" <- #true) ;;
+    !"y".
 
 Definition earlyChoice: val :=
   λ: "x",
-  let: "r" := rand #() in
-  "x" <- #0 ;;
-      "r".
+    let: "r" := rand #() in
+    "x" <- #0 ;;
+    "r".
 
 Definition lateChoice: val :=
   λ: "x",
-  let: "p" := NewProph in
-  "x" <- #0 ;;
-      let: "r" := rand #() in
-      resolve_proph: "p" to: "r" ;;
-      "r".
+    let: "p" := NewProph in
+    "x" <- #0 ;;
+    let: "r" := rand #() in
+    resolve_proph: "p" to: "r" ;;
+    "r".
 
 Section coinflip.
-  Context `{!heapG Σ} (N: namespace).
+  Context `{!heapG Σ}.
+
+  Local Definition N := nroot .@ "coin".
 
   Lemma rand_spec :
     {{{ True }}} rand #() {{{ (b : bool), RET #b; True }}}.
-  Proof using N.
+  Proof.
     iIntros (Φ) "_ HP".
     wp_lam. wp_alloc l as "Hl". wp_lam.
     iMod (inv_alloc N _ (∃ (b: bool), l ↦ #b)%I with "[Hl]") as "#Hinv"; first by eauto.
@@ -50,7 +50,7 @@ Section coinflip.
         earlyChoice #x
         @ ⊤
     <<< ∃ (b: bool), x ↦ #0, RET #b >>>.
-  Proof using N.
+  Proof.
     iApply wp_atomic_intro. iIntros (Φ) "AU". wp_lam.
     wp_bind (rand _)%E. iApply rand_spec; first done.
     iIntros (b) "!> _". wp_let.
@@ -62,7 +62,7 @@ Section coinflip.
     iModIntro. wp_seq. done.
   Qed.
 
-  Definition val_to_bool v : bool :=
+  Definition val_to_bool (v : option val) : bool :=
     match v with
     | Some (LitV (LitBool b)) => b
     | _                       => true
@@ -73,7 +73,7 @@ Section coinflip.
         lateChoice #x
         @ ⊤
     <<< ∃ (b: bool), x ↦ #0, RET #b >>>.
-  Proof using N.
+  Proof.
     iApply wp_atomic_intro. iIntros (Φ) "AU". wp_lam.
     wp_apply wp_new_proph; first done.
     iIntros (v p) "Hp".
diff --git a/theories/heap_lang/proph_map.v b/theories/heap_lang/proph_map.v
index 3a32f3487a54b9fd602486d02429a2715a5db968..da2f43d25628e525012a81730deeefc72a023332 100644
--- a/theories/heap_lang/proph_map.v
+++ b/theories/heap_lang/proph_map.v
@@ -8,7 +8,6 @@ Definition proph_map (P V : Type) `{Countable P} := gmap P (option V).
 Definition proph_val_list (P V : Type) := list (P * V).
 
 Section first_resolve.
-
   Context {P V : Type} `{Countable P}.
   Implicit Type pvs : proph_val_list P V.
   Implicit Type p : P.
@@ -132,7 +131,6 @@ Section to_proph_map.
 End to_proph_map.
 
 
-
 Lemma proph_map_init `{proph_mapPreG P V PVS} pvs ps :
   (|==> ∃ _ : proph_mapG P V PVS, proph_map_ctx pvs ps)%I.
 Proof.
diff --git a/theories/heap_lang/tactics.v b/theories/heap_lang/tactics.v
index 5cf879d9d6995bc2ee7140a5a12cdd67c0c7bc30..d0cca5401f823ae2782e46e5d2e42adf42eb9fbb 100644
--- a/theories/heap_lang/tactics.v
+++ b/theories/heap_lang/tactics.v
@@ -97,11 +97,11 @@ Ltac of_expr e :=
      let e0 := of_expr e0 in let e1 := of_expr e1 in let e2 := of_expr e2 in
      constr:(CAS e0 e1 e2)
   | heap_lang.FAA ?e1 ?e2 =>
-    let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(FAA e1 e2)
+     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(FAA e1 e2)
   | heap_lang.NewProph =>
-    constr:(NewProph)
+     constr:(NewProph)
   | heap_lang.ResolveProph ?e1 ?e2 =>
-    let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(ResolveProph e1 e2)
+     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(ResolveProph e1 e2)
   | to_expr ?e => e
   | of_val ?v => constr:(Val v (of_val v) (to_of_val v))
   | language.of_val ?v => constr:(Val v (of_val v) (to_of_val v))
diff --git a/theories/heap_lang/total_adequacy.v b/theories/heap_lang/total_adequacy.v
index 0abe7f1f539c678e006654331890c77fb9b42380..5b61b23da89e6bc777c2a236ddb3bbce32c874d8 100644
--- a/theories/heap_lang/total_adequacy.v
+++ b/theories/heap_lang/total_adequacy.v
@@ -12,6 +12,6 @@ Proof.
   iMod (gen_heap_init σ.(heap)) as (?) "Hh".
   iMod (proph_map_init [] σ.(used_proph)) as (?) "Hp".
   iModIntro.
-  iExists (fun σ κs => (gen_heap_ctx σ.(heap) ∗ proph_map_ctx κs σ.(used_proph))%I). iFrame.
+  iExists (λ σ κs, (gen_heap_ctx σ.(heap) ∗ proph_map_ctx κs σ.(used_proph))%I). iFrame.
   iApply (Hwp (HeapG _ _ _ _)).
 Qed.
diff --git a/theories/program_logic/adequacy.v b/theories/program_logic/adequacy.v
index c48f569d69777c02753eb1ce5f21397de750a17c..ee2b5375b0f81c7672f0a7658f5348bed8253ea9 100644
--- a/theories/program_logic/adequacy.v
+++ b/theories/program_logic/adequacy.v
@@ -187,15 +187,15 @@ Proof.
     iMod wsat_alloc as (Hinv) "[Hw HE]". specialize (Hwp _ κs).
     rewrite {1}uPred_fupd_eq in Hwp; iMod (Hwp with "[$Hw $HE]") as ">(Hw & HE & Hwp)".
     iDestruct "Hwp" as (Istate) "[HI Hwp]".
-    iApply (@wptp_result _ _ (IrisG _ _ _ Hinv Istate) _ _ _ _ _ []); last first.
-    rewrite app_nil_r. all: eauto with iFrame.
+    iApply (@wptp_result _ _ (IrisG _ _ _ Hinv Istate) _ _ _ _ _ []); first by eauto.
+    rewrite app_nil_r. eauto with iFrame.
   - destruct s; last done. intros t2 σ2 e2 _ [n [κs ?]]%erased_steps_nsteps ?.
     eapply (soundness (M:=iResUR Σ) _ (S (S n))).
     iMod wsat_alloc as (Hinv) "[Hw HE]". specialize (Hwp _ κs).
     rewrite uPred_fupd_eq in Hwp; iMod (Hwp with "[$Hw $HE]") as ">(Hw & HE & Hwp)".
     iDestruct "Hwp" as (Istate) "[HI Hwp]".
-    iApply (@wptp_safe _ _ (IrisG _ _ _ Hinv Istate) _ _ _ []); last first.
-    rewrite app_nil_r. all: eauto with iFrame.
+    iApply (@wptp_safe _ _ (IrisG _ _ _ Hinv Istate) _ _ _ []); [by eauto..|].
+    rewrite app_nil_r. eauto with iFrame.
 Qed.
 
 Theorem wp_adequacy Σ Λ `{invPreG Σ} s e σ φ :
diff --git a/theories/program_logic/language.v b/theories/program_logic/language.v
index 969eb63bfd372cb952629e1b5bdb57fbcd2237a2..de88479fcd85137796b98b0ecfa07c6d939759a9 100644
--- a/theories/program_logic/language.v
+++ b/theories/program_logic/language.v
@@ -110,12 +110,12 @@ Section language.
     option_list x ++ xs.
 
   Inductive nsteps : nat → cfg Λ → list (observation Λ) → cfg Λ → Prop :=
-  | nsteps_refl ρ :
-      nsteps 0 ρ [] ρ
-  | nsteps_l n ρ1 ρ2 ρ3 κ κs :
-      step ρ1 κ ρ2 →
-      nsteps n ρ2 κs ρ3 →
-      nsteps (S n) ρ1 (cons_obs κ κs) ρ3.
+    | nsteps_refl ρ :
+       nsteps 0 ρ [] ρ
+    | nsteps_l n ρ1 ρ2 ρ3 κ κs :
+       step ρ1 κ ρ2 →
+       nsteps n ρ2 κs ρ3 →
+       nsteps (S n) ρ1 (cons_obs κ κs) ρ3.
   Hint Constructors nsteps.
 
   Definition erased_step (ρ1 ρ2 : cfg Λ) := ∃ κ, step ρ1 κ ρ2.
@@ -174,7 +174,7 @@ Section language.
   Lemma erased_step_Permutation (t1 t1' t2 : list (expr Λ)) σ1 σ2 :
     t1 ≡ₚ t1' → erased_step (t1,σ1) (t2,σ2) → ∃ t2', t2 ≡ₚ t2' ∧ erased_step (t1',σ1) (t2',σ2).
   Proof.
-    intros* Heq [? Hs]. pose proof (step_Permutation _ _ _ _ _ _ Heq Hs). firstorder.
+    intros Heq [? Hs]. pose proof (step_Permutation _ _ _ _ _ _ Heq Hs). firstorder.
   Qed.
 
   Record pure_step (e1 e2 : expr Λ) := {
diff --git a/theories/program_logic/total_weakestpre.v b/theories/program_logic/total_weakestpre.v
index 6b47d468a1cabbc554b34743785467e1e06062ff..c27629d3e1ccc5fbc070a377f6f71a1b7018d03c 100644
--- a/theories/program_logic/total_weakestpre.v
+++ b/theories/program_logic/total_weakestpre.v
@@ -190,8 +190,7 @@ Proof.
   rewrite wp_unfold twp_unfold /wp_pre /twp_pre. destruct (to_val e) as [v|]=>//.
   iIntros (σ1 κ κs) "Hσ". iMod ("H" with "Hσ") as "[% H]". iIntros "!>". iSplitR.
   { destruct s; last done. eauto using reducible_no_obs_reducible. }
-  iIntros (e2 σ2 efs Hstep). iMod ("H" $! _ _ _ _ Hstep) as "(% & Hst & H & Hfork)".
-  subst κ. iFrame "Hst".
+  iIntros (e2 σ2 efs Hstep). iMod ("H" $! _ _ _ _ Hstep) as "(-> & $ & H & Hfork)".
   iApply step_fupd_intro; [set_solver+|]. iNext.
   iSplitL "H". by iApply "IH". iApply (@big_sepL_impl with "[$Hfork]").
   iIntros "!#" (k e' _) "H". by iApply "IH".