diff --git a/_CoqProject b/_CoqProject
index 1ff1ee86ecf91feccb32341d58118264afb0171a..9df1d533e43175cc854e9ae6e0f5134a3691e953 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -66,13 +66,9 @@ program_logic/model.v
 program_logic/adequacy.v
 program_logic/lifting.v
 program_logic/invariants.v
-program_logic/viewshifts.v
-program_logic/wsat.v
 program_logic/ownership.v
 program_logic/weakestpre.v
-program_logic/weakestpre_fix.v
 program_logic/pviewshifts.v
-program_logic/resources.v
 program_logic/hoare.v
 program_logic/language.v
 program_logic/ectx_language.v
@@ -86,6 +82,7 @@ program_logic/sts.v
 program_logic/namespaces.v
 program_logic/boxes.v
 program_logic/counter_examples.v
+program_logic/iris.v
 heap_lang/lang.v
 heap_lang/tactics.v
 heap_lang/wp_tactics.v
@@ -105,7 +102,6 @@ heap_lang/lib/barrier/protocol.v
 heap_lang/lib/barrier/proof.v
 heap_lang/proofmode.v
 tests/heap_lang.v
-tests/program_logic.v
 tests/one_shot.v
 tests/joining_existentials.v
 tests/proofmode.v
@@ -122,6 +118,5 @@ proofmode/notation.v
 proofmode/invariants.v
 proofmode/weakestpre.v
 proofmode/ghost_ownership.v
-proofmode/sts.v
 proofmode/classes.v
 proofmode/class_instances.v
diff --git a/heap_lang/derived.v b/heap_lang/derived.v
index 7344167581e707950ee0f3e8f0c7e4f6f7edd8ee..3783733b53a2b7fb68a1a14a4eee53b7e51ed12b 100644
--- a/heap_lang/derived.v
+++ b/heap_lang/derived.v
@@ -12,9 +12,9 @@ Notation Skip := (Seq (Lit LitUnit) (Lit LitUnit)).
 Notation Match e0 x1 e1 x2 e2 := (Case e0 (Lam x1 e1) (Lam x2 e2)).
 
 Section derived.
-Context {Σ : iFunctor}.
-Implicit Types P Q : iProp heap_lang Σ.
-Implicit Types Φ : val → iProp heap_lang Σ.
+Context `{irisG heap_lang Σ}.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val → iProp Σ.
 
 (** Proof rules for the sugar *)
 Lemma wp_lam E x ef e Φ :
diff --git a/heap_lang/heap.v b/heap_lang/heap.v
index adbcaa6760cf1cb3a74a23e3e8cfe9b6eab36595..095d82d5eba86e7dfa28382afbfa0ff331428b15 100644
--- a/heap_lang/heap.v
+++ b/heap_lang/heap.v
@@ -1,5 +1,5 @@
 From iris.heap_lang Require Export lifting.
-From iris.algebra Require Import upred_big_op frac dec_agree.
+From iris.algebra Require Import upred_big_op gmap frac dec_agree.
 From iris.program_logic Require Export invariants ghost_ownership.
 From iris.program_logic Require Import ownership auth.
 From iris.proofmode Require Import weakestpre.
@@ -13,7 +13,8 @@ Definition heapUR : ucmraT := gmapUR loc (prodR fracR (dec_agreeR val)).
 
 (** The CMRA we need. *)
 Class heapG Σ := HeapG {
-  heap_inG :> authG heap_lang Σ heapUR;
+  heapG_iris_inG :> irisG heap_lang Σ;
+  heap_inG :> authG Σ heapUR;
   heap_name : gname
 }.
 (** The Functor we need. *)
@@ -25,16 +26,16 @@ Definition of_heap : heapUR → state := omap (maybe DecAgree ∘ snd).
 Section definitions.
   Context `{heapG Σ}.
 
-  Definition heap_mapsto_def (l : loc) (q : Qp) (v: val) : iPropG heap_lang Σ :=
+  Definition heap_mapsto_def (l : loc) (q : Qp) (v: val) : iProp Σ :=
     auth_own heap_name {[ l := (q, DecAgree v) ]}.
   Definition heap_mapsto_aux : { x | x = @heap_mapsto_def }. by eexists. Qed.
   Definition heap_mapsto := proj1_sig heap_mapsto_aux.
   Definition heap_mapsto_eq : @heap_mapsto = @heap_mapsto_def :=
     proj2_sig heap_mapsto_aux.
 
-  Definition heap_inv (h : heapUR) : iPropG heap_lang Σ :=
+  Definition heap_inv (h : heapUR) : iProp Σ :=
     ownP (of_heap h).
-  Definition heap_ctx : iPropG heap_lang Σ :=
+  Definition heap_ctx : iProp Σ :=
     auth_ctx heap_name heapN heap_inv.
 
   Global Instance heap_inv_proper : Proper ((≡) ==> (⊣⊢)) heap_inv.
@@ -44,9 +45,7 @@ Section definitions.
 End definitions.
 
 Typeclasses Opaque heap_ctx heap_mapsto.
-Instance: Params (@heap_inv) 1.
-Instance: Params (@heap_mapsto) 4.
-Instance: Params (@heap_ctx) 2.
+Instance: Params (@heap_inv) 2.
 
 Notation "l ↦{ q } v" := (heap_mapsto l q v)
   (at level 20, q at level 50, format "l  ↦{ q }  v") : uPred_scope.
@@ -54,8 +53,8 @@ Notation "l ↦ v" := (heap_mapsto l 1 v) (at level 20) : uPred_scope.
 
 Section heap.
   Context {Σ : gFunctors}.
-  Implicit Types P Q : iPropG heap_lang Σ.
-  Implicit Types Φ : val → iPropG heap_lang Σ.
+  Implicit Types P Q : iProp Σ.
+  Implicit Types Φ : val → iProp Σ.
   Implicit Types σ : state.
   Implicit Types h g : heapUR.
 
@@ -103,15 +102,14 @@ Section heap.
   Hint Resolve heap_store_valid.
 
   (** Allocation *)
-  Lemma heap_alloc E σ :
-    authG heap_lang Σ heapUR → nclose heapN ⊆ E →
+  Lemma heap_alloc `{irisG heap_lang Σ, authG Σ heapUR} E σ :
     ownP σ ={E}=> ∃ _ : heapG Σ, heap_ctx ∧ [★ map] l↦v ∈ σ, l ↦ v.
   Proof.
     intros. rewrite -{1}(from_to_heap σ). etrans.
     { rewrite [ownP _]later_intro.
       apply (auth_alloc (ownP ∘ of_heap) heapN E); auto using to_heap_valid. }
     apply pvs_mono, exist_elim=> γ.
-    rewrite -(exist_intro (HeapG _ _ γ)) /heap_ctx; apply and_mono_r.
+    rewrite -(exist_intro (HeapG _ _ _ γ)) /heap_ctx; apply and_mono_r.
     rewrite heap_mapsto_eq /heap_mapsto_def /heap_name.
     induction σ as [|l v σ Hl IH] using map_ind.
     { rewrite big_sepM_empty; apply True_intro. }
@@ -157,22 +155,21 @@ Section heap.
   Proof. by rewrite heap_mapsto_op_half. Qed.
 
   (** Weakest precondition *)
-  (* FIXME: try to reduce usage of wp_pvs. We're losing view shifts here. *)
   Lemma wp_alloc E e v Φ :
     to_val e = Some v → nclose heapN ⊆ E →
     heap_ctx ★ ▷ (∀ l, l ↦ v ={E}=★ Φ (LitV (LitLoc l))) ⊢ WP Alloc e @ E {{ Φ }}.
   Proof.
     iIntros (<-%of_to_val ?) "[#Hinv HΦ]". rewrite /heap_ctx.
-    iPvs (auth_empty heap_name) as "Hheap".
-    iApply wp_pvs; iApply (auth_fsa heap_inv (wp_fsa _)); eauto with fsaV.
-    iFrame "Hinv Hheap". iIntros (h). rewrite left_id.
-    iIntros "[% Hheap]". rewrite /heap_inv.
-    iApply wp_alloc_pst. iFrame "Hheap". iNext.
-    iIntros (l) "[% Hheap]"; iPvsIntro; iExists {[ l := (1%Qp, DecAgree v) ]}.
-    rewrite -of_heap_insert -(insert_singleton_op h); last by apply of_heap_None.
-    iFrame "Hheap". iSplitR; first iPureIntro.
-    { by apply alloc_unit_singleton_local_update; first apply of_heap_None. }
-    iIntros "Hheap". iApply "HΦ". by rewrite heap_mapsto_eq /heap_mapsto_def.
+    iVs (auth_empty heap_name) as "Hh".
+    iVs (auth_open with "[Hh]") as (h) "[Hv [Hh Hclose]]"; eauto.
+    rewrite left_id /heap_inv. iDestruct "Hv" as %?.
+    iApply wp_alloc_pst. iFrame "Hh". iNext.
+    iIntros (l) "[% Hh]"; iVsIntro.
+    iVs ("Hclose" $! {[ l := (1%Qp, DecAgree v) ]} with "[Hh]").
+    { rewrite -of_heap_insert -(insert_singleton_op h); last by apply of_heap_None.
+      iFrame "Hh". iPureIntro.
+      by apply alloc_unit_singleton_local_update; first apply of_heap_None. }
+    iApply "HΦ". by rewrite heap_mapsto_eq /heap_mapsto_def.
   Qed.
 
   Lemma wp_load E l q v Φ :
@@ -180,14 +177,15 @@ Section heap.
     heap_ctx ★ ▷ l ↦{q} v ★ ▷ (l ↦{q} v ={E}=★ Φ v)
     ⊢ WP Load (Lit (LitLoc l)) @ E {{ Φ }}.
   Proof.
-    iIntros (?) "[#Hh [Hl HΦ]]".
+    iIntros (?) "[#Hinv [Hl HΦ]]".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
-    iApply wp_pvs; iApply (auth_fsa heap_inv (wp_fsa _)); eauto with fsaV.
-    iFrame "Hh Hl". iIntros (h) "[% Hl]". rewrite /heap_inv.
+    iVs (auth_open with "[Hl]") as (h) "[% [Hl Hclose]]"; eauto.
+    rewrite /heap_inv.
     iApply (wp_load_pst _ (<[l:=v]>(of_heap h)));first by rewrite lookup_insert.
     rewrite of_heap_singleton_op //. iFrame "Hl".
-    iIntros "> Hown". iPvsIntro. iExists _; iSplit; first done.
-    rewrite of_heap_singleton_op //. by iFrame.
+    iIntros "> Hown". iVsIntro. iVs ("Hclose" with "* [Hown]").
+    { iSplit; first done. rewrite of_heap_singleton_op //. by iFrame. }
+    by iApply "HΦ".
   Qed.
 
   Lemma wp_store E l v' e v Φ :
@@ -195,15 +193,18 @@ Section heap.
     heap_ctx ★ ▷ l ↦ v' ★ ▷ (l ↦ v ={E}=★ Φ (LitV LitUnit))
     ⊢ WP Store (Lit (LitLoc l)) e @ E {{ Φ }}.
   Proof.
-    iIntros (<-%of_to_val ?) "[#Hh [Hl HΦ]]".
+    iIntros (<-%of_to_val ?) "[#Hinv [Hl HΦ]]".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
-    iApply wp_pvs; iApply (auth_fsa heap_inv (wp_fsa _)); eauto with fsaV.
-    iFrame "Hh Hl". iIntros (h) "[% Hl]". rewrite /heap_inv.
+    iVs (auth_open with "[Hl]") as (h) "[% [Hl Hclose]]"; eauto.
+    rewrite /heap_inv.
     iApply (wp_store_pst _ (<[l:=v']>(of_heap h))); rewrite ?lookup_insert //.
     rewrite insert_insert !of_heap_singleton_op; eauto. iFrame "Hl".
-    iIntros "> Hown". iPvsIntro. iExists {[l := (1%Qp, DecAgree v)]}; iSplit.
-    { iPureIntro; by apply singleton_local_update, exclusive_local_update. }
-    rewrite of_heap_singleton_op //; eauto. by iFrame.
+    iIntros "> Hl". iVsIntro.
+    iVs ("Hclose" $! {[l := (1%Qp, DecAgree v)]} with "[Hl]").
+    { iSplit.
+      - iPureIntro; by apply singleton_local_update, exclusive_local_update.
+      - rewrite of_heap_singleton_op //; eauto. }
+    by iApply "HΦ".
   Qed.
 
   Lemma wp_cas_fail E l q v' e1 v1 e2 v2 Φ :
@@ -213,12 +214,13 @@ Section heap.
   Proof.
     iIntros (<-%of_to_val <-%of_to_val ??) "[#Hh [Hl HΦ]]".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
-    iApply wp_pvs; iApply (auth_fsa heap_inv (wp_fsa _)); eauto with fsaV.
-    iFrame "Hh Hl". iIntros (h) "[% Hl]". rewrite /heap_inv.
+    iVs (auth_open with "[Hl]") as (h) "[% [Hl Hclose]]"; eauto.
+    rewrite /heap_inv.
     iApply (wp_cas_fail_pst _ (<[l:=v']>(of_heap h))); rewrite ?lookup_insert //.
     rewrite of_heap_singleton_op //. iFrame "Hl".
-    iIntros "> Hown". iPvsIntro. iExists _; iSplit; first done.
-    rewrite of_heap_singleton_op //. by iFrame.
+    iIntros "> Hown". iVsIntro. iVs ("Hclose" with "* [Hown]").
+    { iSplit; first done. rewrite of_heap_singleton_op //. by iFrame. }
+    by iApply "HΦ".
   Qed.
 
   Lemma wp_cas_suc E l e1 v1 e2 v2 Φ :
@@ -228,12 +230,15 @@ Section heap.
   Proof.
     iIntros (<-%of_to_val <-%of_to_val ?) "[#Hh [Hl HΦ]]".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
-    iApply wp_pvs; iApply (auth_fsa heap_inv (wp_fsa _)); eauto with fsaV.
-    iFrame "Hh Hl". iIntros (h) "[% Hl]". rewrite /heap_inv.
+    iVs (auth_open with "[Hl]") as (h) "[% [Hl Hclose]]"; eauto.
+    rewrite /heap_inv.
     iApply (wp_cas_suc_pst _ (<[l:=v1]>(of_heap h))); rewrite ?lookup_insert //.
     rewrite insert_insert !of_heap_singleton_op; eauto. iFrame "Hl".
-    iIntros "> Hown". iPvsIntro. iExists {[l := (1%Qp, DecAgree v2)]}; iSplit.
-    { iPureIntro; by apply singleton_local_update, exclusive_local_update. }
-    rewrite of_heap_singleton_op //; eauto. by iFrame.
+    iIntros "> Hl". iVsIntro.
+    iVs ("Hclose" $! {[l := (1%Qp, DecAgree v2)]} with "[Hl]").
+    { iSplit.
+      - iPureIntro; by apply singleton_local_update, exclusive_local_update.
+      - rewrite of_heap_singleton_op //; eauto. }
+    by iApply "HΦ".
   Qed.
 End heap.
diff --git a/heap_lang/lib/assert.v b/heap_lang/lib/assert.v
index 3cc335765624ebfc9f5598ccb39d62b5ec3d1ae9..4e62f488b3535afc4aafb4d28d83f4cc739f0717 100644
--- a/heap_lang/lib/assert.v
+++ b/heap_lang/lib/assert.v
@@ -7,7 +7,7 @@ Definition assert : val :=
 Notation "'assert:' e" := (assert (λ: <>, e))%E (at level 99) : expr_scope.
 Global Opaque assert.
 
-Lemma wp_assert {Σ} E (Φ : val → iProp heap_lang Σ) e `{!Closed [] e} :
+Lemma wp_assert `{heapG Σ} E (Φ : val → iProp Σ) e `{!Closed [] e} :
   WP e @ E {{ v, v = #true ∧ ▷ Φ #() }} ⊢ WP assert: e @ E {{ Φ }}.
 Proof.
   iIntros "HΦ". rewrite /assert. wp_let. wp_seq.
diff --git a/heap_lang/lib/barrier/proof.v b/heap_lang/lib/barrier/proof.v
index ba12947ad01f07a030a232f256a9fc1f1cdd93dd..4827e526b9fcc836e6102a3240cf80efdde95e8c 100644
--- a/heap_lang/lib/barrier/proof.v
+++ b/heap_lang/lib/barrier/proof.v
@@ -1,57 +1,54 @@
 From iris.prelude Require Import functions.
 From iris.algebra Require Import upred_big_op.
-From iris.program_logic Require Import saved_prop.
+From iris.program_logic Require Import saved_prop sts.
 From iris.heap_lang Require Import proofmode.
-From iris.proofmode Require Import sts.
 From iris.heap_lang.lib.barrier Require Export barrier.
 From iris.heap_lang.lib.barrier Require Import protocol.
-Import uPred.
 
 (** The CMRAs we need. *)
 (* Not bundling heapG, as it may be shared with other users. *)
 Class barrierG Σ := BarrierG {
-  barrier_stsG :> stsG heap_lang Σ sts;
-  barrier_savedPropG :> savedPropG heap_lang Σ idCF;
+  barrier_stsG :> stsG Σ sts;
+  barrier_savedPropG :> savedPropG Σ idCF;
 }.
 (** The Functors we need. *)
 Definition barrierGF : gFunctorList := [stsGF sts; savedPropGF idCF].
 (* Show and register that they match. *)
-Instance inGF_barrierG `{H : inGFs heap_lang Σ barrierGF} : barrierG Σ.
+Instance inGF_barrierG `{H : inGFs Σ barrierGF} : barrierG Σ.
 Proof. destruct H as (?&?&?). split; apply _. Qed.
 
 (** Now we come to the Iris part of the proof. *)
 Section proof.
 Context `{!heapG Σ, !barrierG Σ} (N : namespace).
 Implicit Types I : gset gname.
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Definition ress (P : iProp) (I : gset gname) : iProp :=
-  (∃ Ψ : gname → iProp,
+Definition ress (P : iProp Σ) (I : gset gname) : iProp Σ :=
+  (∃ Ψ : gname → iProp Σ,
     ▷ (P -★ [★ set] i ∈ I, Ψ i) ★ [★ set] i ∈ I, saved_prop_own i (Ψ i))%I.
 
 Coercion state_to_val (s : state) : val :=
   match s with State Low _ => #0 | State High _ => #1 end.
 Arguments state_to_val !_ / : simpl nomatch.
 
-Definition state_to_prop (s : state) (P : iProp) : iProp :=
+Definition state_to_prop (s : state) (P : iProp Σ) : iProp Σ :=
   match s with State Low _ => P | State High _ => True%I end.
 Arguments state_to_prop !_ _ / : simpl nomatch.
 
-Definition barrier_inv (l : loc) (P : iProp) (s : state) : iProp :=
+Definition barrier_inv (l : loc) (P : iProp Σ) (s : state) : iProp Σ :=
   (l ↦ s ★ ress (state_to_prop s P) (state_I s))%I.
 
-Definition barrier_ctx (γ : gname) (l : loc) (P : iProp) : iProp :=
+Definition barrier_ctx (γ : gname) (l : loc) (P : iProp Σ) : iProp Σ :=
   (■ (heapN ⊥ N) ★ heap_ctx ★ sts_ctx γ N (barrier_inv l P))%I.
 
-Definition send (l : loc) (P : iProp) : iProp :=
+Definition send (l : loc) (P : iProp Σ) : iProp Σ :=
   (∃ γ, barrier_ctx γ l P ★ sts_ownS γ low_states {[ Send ]})%I.
 
-Definition recv (l : loc) (R : iProp) : iProp :=
+Definition recv (l : loc) (R : iProp Σ) : iProp Σ :=
   (∃ γ P Q i,
     barrier_ctx γ l P ★ sts_ownS γ (i_states i) {[ Change i ]} ★
     saved_prop_own i Q ★ ▷ (Q -★ R))%I.
 
-Global Instance barrier_ctx_persistent (γ : gname) (l : loc) (P : iProp) :
+Global Instance barrier_ctx_persistent (γ : gname) (l : loc) (P : iProp Σ) :
   PersistentP (barrier_ctx γ l P).
 Proof. apply _. Qed.
 
@@ -93,15 +90,15 @@ Proof.
 Qed.
 
 (** Actual proofs *)
-Lemma newbarrier_spec (P : iProp) (Φ : val → iProp) :
+Lemma newbarrier_spec (P : iProp Σ) (Φ : val → iProp Σ) :
   heapN ⊥ N →
   heap_ctx ★ (∀ l, recv l P ★ send l P -★ Φ #l) ⊢ WP newbarrier #() {{ Φ }}.
 Proof.
   iIntros (HN) "[#? HΦ]".
   rewrite /newbarrier. wp_seq. wp_alloc l as "Hl".
   iApply ("HΦ" with "|==>[-]").
-  iPvs (saved_prop_alloc (F:=idCF) _ P) as (γ) "#?".
-  iPvs (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]")
+  iVs (saved_prop_alloc (F:=idCF) P) as (γ) "#?".
+  iVs (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]")
     as (γ') "[#? Hγ']"; eauto.
   { iNext. rewrite /barrier_inv /=. iFrame.
     iExists (const P). rewrite !big_sepS_singleton /=. eauto. }
@@ -110,59 +107,57 @@ Proof.
   iAssert (sts_ownS γ' (i_states γ) {[Change γ]}
     ★ sts_ownS γ' low_states {[Send]})%I with "|==>[-]" as "[Hr Hs]".
   { iApply sts_ownS_op; eauto using i_states_closed, low_states_closed.
-    + set_solver.
-    + iApply (sts_own_weaken with "Hγ'");
+    - set_solver.
+    - iApply (sts_own_weaken with "Hγ'");
         auto using sts.closed_op, i_states_closed, low_states_closed;
         abstract set_solver. }
-  iPvsIntro. rewrite /recv /send. iSplitL "Hr".
+  iVsIntro. rewrite /recv /send. iSplitL "Hr".
   - iExists γ', P, P, γ. iFrame. auto.
   - auto.
 Qed.
 
-Lemma signal_spec l P (Φ : val → iProp) :
+Lemma signal_spec l P (Φ : val → iProp Σ) :
   send l P ★ P ★ Φ #() ⊢ WP signal #l {{ Φ }}.
 Proof.
   rewrite /signal /send /barrier_ctx.
   iIntros "(Hs&HP&HΦ)"; iDestruct "Hs" as (γ) "[#(%&Hh&Hsts) Hγ]". wp_let.
-  iSts γ as [p I]; iDestruct "Hγ" as "[Hl Hr]".
-  wp_store. iPvsIntro. destruct p; [|done].
-  iExists (State High I), (∅ : set token).
+  iVs (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
+    as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
+  destruct p; [|done]. wp_store. iFrame "HΦ".
+  iVs ("Hclose" $! (State High I) (∅ : set token) with "[-]"); last done.
   iSplit; [iPureIntro; by eauto using signal_step|].
-  iSplitR "HΦ"; [iNext|by auto].
-  rewrite {2}/barrier_inv /ress /=; iFrame "Hl".
+  iNext. rewrite {2}/barrier_inv /ress /=; iFrame "Hl".
   iDestruct "Hr" as (Ψ) "[Hr Hsp]"; iExists Ψ; iFrame "Hsp".
   iIntros "> _"; by iApply "Hr".
 Qed.
 
-Lemma wait_spec l P (Φ : val → iProp) :
+Lemma wait_spec l P (Φ : val → iProp Σ) :
   recv l P ★ (P -★ Φ #()) ⊢ WP wait #l {{ Φ }}.
 Proof.
   rename P into R; rewrite /recv /barrier_ctx.
   iIntros "[Hr HΦ]"; iDestruct "Hr" as (γ P Q i) "(#(%&Hh&Hsts)&Hγ&#HQ&HQR)".
   iLöb as "IH". wp_rec. wp_focus (! _)%E.
-  iSts γ as [p I]; iDestruct "Hγ" as "[Hl Hr]".
-  wp_load. iPvsIntro. destruct p.
-  - (* a Low state. The comparison fails, and we recurse. *)
-    iExists (State Low I), {[ Change i ]}; iSplit; [done|iSplitL "Hl Hr"].
-    { iNext. rewrite {2}/barrier_inv /=. by iFrame. }
-    iIntros "Hγ".
+  iVs (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
+    as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
+  wp_load. destruct p.
+  - iVs ("Hclose" $! (State Low I) {[ Change i ]} with "[Hl Hr]") as "Hγ".
+    { iSplit; first done. iNext. rewrite {2}/barrier_inv /=. by iFrame. }
     iAssert (sts_ownS γ (i_states i) {[Change i]})%I with "|==>[Hγ]" as "Hγ".
     { iApply (sts_own_weaken with "Hγ"); eauto using i_states_closed. }
-    wp_op=> ?; simplify_eq; wp_if. iApply ("IH" with "Hγ [HQR] HΦ"). auto.
+    iVsIntro. wp_op=> ?; simplify_eq; wp_if.
+    iApply ("IH" with "Hγ [HQR] HΦ"). auto.
   - (* a High state: the comparison succeeds, and we perform a transition and
     return to the client *)
-    iExists (State High (I ∖ {[ i ]})), (∅ : set token).
-    iSplit; [iPureIntro; by eauto using wait_step|].
     iDestruct "Hr" as (Ψ) "[HΨ Hsp]".
     iDestruct (big_sepS_delete _ _ i with "Hsp") as "[#HΨi Hsp]"; first done.
     iAssert (▷ Ψ i ★ ▷ [★ set] j ∈ I ∖ {[i]}, Ψ j)%I with "[HΨ]" as "[HΨ HΨ']".
     { iNext. iApply (big_sepS_delete _ _ i); first done. by iApply "HΨ". }
-    iSplitL "HΨ' Hl Hsp"; [iNext|].
-    + rewrite {2}/barrier_inv /=; iFrame "Hl".
-      iExists Ψ; iFrame. auto.
-    + iPoseProof (saved_prop_agree i Q (Ψ i) with "[#]") as "Heq"; first by auto.
-      iIntros "_". wp_op=> ?; simplify_eq/=; wp_if.
-      iPvsIntro. iApply "HΦ". iApply "HQR". by iRewrite "Heq".
+    iVs ("Hclose" $! (State High (I ∖ {[ i ]})) (∅ : set token) with "[HΨ' Hl Hsp]").
+    { iSplit; [iPureIntro; by eauto using wait_step|].
+      iNext. rewrite {2}/barrier_inv /=; iFrame "Hl". iExists Ψ; iFrame. auto. }
+    iPoseProof (saved_prop_agree i Q (Ψ i) with "[#]") as "Heq"; first by auto.
+    iVsIntro. wp_op=> ?; simplify_eq/=; wp_if.
+    iVsIntro. iApply "HΦ". iApply "HQR". by iRewrite "Heq".
 Qed.
 
 Lemma recv_split E l P1 P2 :
@@ -170,28 +165,27 @@ Lemma recv_split E l P1 P2 :
 Proof.
   rename P1 into R1; rename P2 into R2. rewrite {1}/recv /barrier_ctx.
   iIntros (?). iDestruct 1 as (γ P Q i) "(#(%&Hh&Hsts)&Hγ&#HQ&HQR)".
-  iApply pvs_trans'.
-  iSts γ as [p I]; iDestruct "Hγ" as "[Hl Hr]".
-  iPvs (saved_prop_alloc_strong _ (R1: ∙%CF iProp) I) as (i1) "[% #Hi1]".
-  iPvs (saved_prop_alloc_strong _ (R2: ∙%CF iProp) (I ∪ {[i1]}))
-    as (i2) "[Hi2' #Hi2]"; iDestruct "Hi2'" as %Hi2; iPvsIntro.
+  iVs (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
+    as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
+  iVs (saved_prop_alloc_strong (R1: ∙%CF (iProp Σ)) I) as (i1) "[% #Hi1]".
+  iVs (saved_prop_alloc_strong (R2: ∙%CF (iProp Σ)) (I ∪ {[i1]}))
+    as (i2) "[Hi2' #Hi2]"; iDestruct "Hi2'" as %Hi2.
   rewrite ->not_elem_of_union, elem_of_singleton in Hi2; destruct Hi2.
-  iExists (State p ({[i1; i2]} ∪ I ∖ {[i]})).
-  iExists ({[Change i1; Change i2 ]}).
-  iSplit; [by eauto using split_step|iSplitL].
-  - iNext. rewrite {2}/barrier_inv /=. iFrame "Hl".
-    iApply (ress_split _ _ _ Q R1 R2); eauto. iFrame; auto.
-  - iIntros "Hγ".
-    iAssert (sts_ownS γ (i_states i1) {[Change i1]}
-      ★ sts_ownS γ (i_states i2) {[Change i2]})%I with "|==>[-]" as "[Hγ1 Hγ2]".
-    { iApply sts_ownS_op; eauto using i_states_closed, low_states_closed.
-      + set_solver.
-      + iApply (sts_own_weaken with "Hγ");
-          eauto using sts.closed_op, i_states_closed.
-        abstract set_solver. }
-    iPvsIntro; iSplitL "Hγ1"; rewrite /recv /barrier_ctx.
-    + iExists γ, P, R1, i1. iFrame; auto.
-    + iExists γ, P, R2, i2. iFrame; auto.
+  iVs ("Hclose" $! (State p ({[i1; i2]} ∪ I ∖ {[i]}))
+                   {[Change i1; Change i2 ]} with "[-]") as "Hγ".
+  { iSplit; first by eauto using split_step.
+    iNext. rewrite {2}/barrier_inv /=. iFrame "Hl".
+    iApply (ress_split _ _ _ Q R1 R2); eauto. iFrame; auto. }
+  iAssert (sts_ownS γ (i_states i1) {[Change i1]}
+    ★ sts_ownS γ (i_states i2) {[Change i2]})%I with "|==>[-]" as "[Hγ1 Hγ2]".
+  { iApply sts_ownS_op; eauto using i_states_closed, low_states_closed.
+    - abstract set_solver.
+    - iApply (sts_own_weaken with "Hγ");
+        eauto using sts.closed_op, i_states_closed.
+      abstract set_solver. }
+  iVsIntro; iSplitL "Hγ1"; rewrite /recv /barrier_ctx.
+  - iExists γ, P, R1, i1. iFrame; auto.
+  - iExists γ, P, R2, i2. iFrame; auto.
 Qed.
 
 Lemma recv_weaken l P1 P2 : (P1 -★ P2) ⊢ recv l P1 -★ recv l P2.
@@ -203,9 +197,7 @@ Proof.
 Qed.
 
 Lemma recv_mono l P1 P2 : (P1 ⊢ P2) → recv l P1 ⊢ recv l P2.
-Proof.
-  intros HP%entails_wand. apply wand_entails. rewrite HP. apply recv_weaken.
-Qed.
+Proof. iIntros (HP) "H". iApply (recv_weaken with "[] H"). iApply HP. Qed.
 End proof.
 
 Typeclasses Opaque barrier_ctx send recv.
diff --git a/heap_lang/lib/barrier/protocol.v b/heap_lang/lib/barrier/protocol.v
index 9520887050982b159697e9e92bae286fa2f1fcae..97258bf775a636219142bb1df6ebf34f0ec5bd74 100644
--- a/heap_lang/lib/barrier/protocol.v
+++ b/heap_lang/lib/barrier/protocol.v
@@ -1,5 +1,6 @@
 From iris.algebra Require Export sts.
 From iris.program_logic Require Import ghost_ownership.
+From iris.prelude Require Export gmap.
 
 (** The STS describing the main barrier protocol. Every state has an index-set
     associated with it. These indices are actually [gname], because we use them
diff --git a/heap_lang/lib/barrier/specification.v b/heap_lang/lib/barrier/specification.v
index 0eb0937ec10795af9224c30be01c4734078e31fd..b9ea55b9e956b5bbf3318dd9a33116194c4cc498 100644
--- a/heap_lang/lib/barrier/specification.v
+++ b/heap_lang/lib/barrier/specification.v
@@ -5,13 +5,11 @@ From iris.heap_lang Require Import proofmode.
 Import uPred.
 
 Section spec.
-Context {Σ : gFunctors} `{!heapG Σ} `{!barrierG Σ}. 
-
-Local Notation iProp := (iPropG heap_lang Σ).
+Context `{!heapG Σ} `{!barrierG Σ}.
 
 Lemma barrier_spec (N : namespace) :
   heapN ⊥ N →
-  ∃ recv send : loc → iProp -n> iProp,
+  ∃ recv send : loc → iProp Σ -n> iProp Σ,
     (∀ P, heap_ctx ⊢ {{ True }} newbarrier #()
                      {{ v, ∃ l : loc, v = #l ★ recv l P ★ send l P }}) ∧
     (∀ l P, {{ send l P ★ P }} signal #l {{ _, True }}) ∧
diff --git a/heap_lang/lib/counter.v b/heap_lang/lib/counter.v
index bd7c6a9754e3cf785a0ad743a3286ca27ec50d95..e1677077e65c8c9b33784b1c6cb5c9e6e3baa11d 100644
--- a/heap_lang/lib/counter.v
+++ b/heap_lang/lib/counter.v
@@ -1,9 +1,7 @@
-(* Monotone counter, but using an explicit CMRA instead of auth *)
-From iris.program_logic Require Export global_functor.
+From iris.heap_lang Require Export lang.
+From iris.proofmode Require Import invariants tactics.
 From iris.program_logic Require Import auth.
-From iris.proofmode Require Import invariants ghost_ownership coq_tactics.
 From iris.heap_lang Require Import proofmode notation.
-Import uPred.
 
 Definition newcounter : val := λ: <>, ref #0.
 Definition inc : val :=
@@ -14,18 +12,17 @@ Definition read : val := λ: "l", !"l".
 Global Opaque newcounter inc get.
 
 (** The CMRA we need. *)
-Class counterG Σ := CounterG { counter_tokG :> authG heap_lang Σ mnatUR }.
+Class counterG Σ := CounterG { counter_tokG :> authG Σ mnatUR }.
 Definition counterGF : gFunctorList := [authGF mnatUR].
-Instance inGF_counterG `{H : inGFs heap_lang Σ counterGF} : counterG Σ.
+Instance inGF_counterG `{H : inGFs Σ counterGF} : counterG Σ.
 Proof. destruct H; split; apply _. Qed.
 
 Section proof.
 Context `{!heapG Σ, !counterG Σ} (N : namespace).
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Definition counter_inv (l : loc) (n : mnat) : iProp := (l ↦ #n)%I.
+Definition counter_inv (l : loc) (n : mnat) : iProp Σ := (l ↦ #n)%I.
 
-Definition counter (l : loc) (n : nat) : iProp :=
+Definition counter (l : loc) (n : nat) : iProp Σ :=
   (∃ γ, heapN ⊥ N ∧ heap_ctx ∧
         auth_ctx γ N (counter_inv l) ∧ auth_own γ (n:mnat))%I.
 
@@ -33,53 +30,56 @@ Definition counter (l : loc) (n : nat) : iProp :=
 Global Instance counter_persistent l n : PersistentP (counter l n).
 Proof. apply _. Qed.
 
-Lemma newcounter_spec (R : iProp) Φ :
+Lemma newcounter_spec (R : iProp Σ) Φ :
   heapN ⊥ N →
   heap_ctx ★ (∀ l, counter l 0 -★ Φ #l) ⊢ WP newcounter #() {{ Φ }}.
 Proof.
   iIntros (?) "[#Hh HΦ]". rewrite /newcounter. wp_seq. wp_alloc l as "Hl".
-  iPvs (auth_alloc (counter_inv l) N _ (O:mnat) with "[Hl]")
+  iVs (auth_alloc (counter_inv l) N _ (O:mnat) with "[Hl]")
     as (γ) "[#? Hγ]"; try by auto.
-  iPvsIntro. iApply "HΦ". rewrite /counter; eauto 10.
+  iVsIntro. iApply "HΦ". rewrite /counter; eauto 10.
 Qed.
 
-Lemma inc_spec l j (Φ : val → iProp) :
+Lemma inc_spec l j (Φ : val → iProp Σ) :
   counter l j ★ (counter l (S j) -★ Φ #()) ⊢ WP inc #l {{ Φ }}.
 Proof.
   iIntros "[Hl HΦ]". iLöb as "IH". wp_rec.
   iDestruct "Hl" as (γ) "(% & #? & #Hγ & Hγf)".
   wp_focus (! _)%E.
-  iApply (auth_fsa (counter_inv l) (wp_fsa _) _ N); auto with fsaV.
-  iIntros "{$Hγ $Hγf}"; iIntros (j') "[% Hl] /="; rewrite {2}/counter_inv.
-  wp_load; iPvsIntro; iExists j; iSplit; [done|iIntros "{$Hl} Hγf"].
-  wp_let; wp_op. wp_focus (CAS _ _ _).
-  iApply (auth_fsa (counter_inv l) (wp_fsa _) _ N); auto with fsaV.
-  iIntros "{$Hγ $Hγf}"; iIntros (j'') "[% Hl] /="; rewrite {2}/counter_inv.
+  iVs (auth_open (counter_inv l) with "[Hγf]") as (j') "(% & Hl & Hclose)"; auto.
+  rewrite {2}/counter_inv.
+  wp_load. iVs ("Hclose" $! j with "[Hl]") as "Hγf"; eauto.
+  iVsIntro. wp_let; wp_op. wp_focus (CAS _ _ _).
+  iVs (auth_open (counter_inv l) with "[Hγf]") as (j'') "(% & Hl & Hclose)"; auto.
+  rewrite {2}/counter_inv.
   destruct (decide (j `max` j'' = j `max` j')) as [Hj|Hj].
-  - wp_cas_suc; first (by do 3 f_equal); iPvsIntro.
-    iExists (1 + j `max` j')%nat; iSplit.
-    { iPureIntro. apply mnat_local_update. abstract lia. }
-    rewrite {2}/counter_inv !mnat_op_max (Nat.max_l (S _)); last abstract lia.
-    rewrite Nat2Z.inj_succ -Z.add_1_l.
-    iIntros "{$Hl} Hγf". wp_if.
-    iPvsIntro; iApply "HΦ"; iExists γ; repeat iSplit; eauto.
+  - wp_cas_suc; first (by do 3 f_equal).
+    iVs ("Hclose" $! (1 + j `max` j')%nat with "[Hl]") as "Hγf".
+    { iSplit; [iPureIntro|iNext].
+      { apply mnat_local_update. abstract lia. }
+      rewrite {2}/counter_inv !mnat_op_max (Nat.max_l (S _)); last abstract lia.
+      by rewrite Nat2Z.inj_succ -Z.add_1_l. }
+    iVsIntro. wp_if.
+    iVsIntro; iApply "HΦ"; iExists γ; repeat iSplit; eauto.
     iApply (auth_own_mono with "Hγf"). apply mnat_included. abstract lia.
   - wp_cas_fail; first (rewrite !mnat_op_max; by intros [= ?%Nat2Z.inj]).
-    iPvsIntro. iExists j; iSplit; [done|iIntros "{$Hl} Hγf"].
-    wp_if. iApply ("IH" with "[Hγf] HΦ"). rewrite {3}/counter; eauto 10.
+    iVs ("Hclose" $! j with "[Hl]") as "Hγf"; eauto.
+    iVsIntro. wp_if. iApply ("IH" with "[Hγf] HΦ"). rewrite {3}/counter; eauto 10.
 Qed.
 
-Lemma read_spec l j (Φ : val → iProp) :
+Lemma read_spec l j (Φ : val → iProp Σ) :
   counter l j ★ (∀ i, ■ (j ≤ i)%nat → counter l i -★ Φ #i)
   ⊢ WP read #l {{ Φ }}.
 Proof.
   iIntros "[Hc HΦ]". iDestruct "Hc" as (γ) "(% & #? & #Hγ & Hγf)".
   rewrite /read. wp_let.
-  iApply (auth_fsa (counter_inv l) (wp_fsa _) _ N); auto with fsaV.
-  iIntros "{$Hγ $Hγf}"; iIntros (j') "[% Hl] /=".
-  wp_load; iPvsIntro; iExists (j `max` j'); iSplit.
-  { iPureIntro; apply mnat_local_update; abstract lia. }
-  rewrite !mnat_op_max -Nat.max_assoc Nat.max_idempotent; iIntros "{$Hl} Hγf".
-  iApply ("HΦ" with "[%]"); first abstract lia; rewrite /counter; eauto 10.
+  iVs (auth_open (counter_inv l) with "[Hγf]") as (j') "(% & Hl & Hclose)"; auto.
+  wp_load.
+  iVs ("Hclose" $! (j `max` j') with "[Hl]") as "Hγf".
+  { iSplit; [iPureIntro|iNext].
+    { apply mnat_local_update; abstract lia. }
+    by rewrite !mnat_op_max -Nat.max_assoc Nat.max_idempotent. }
+  iVsIntro. rewrite !mnat_op_max.
+  iApply ("HΦ" with "[%]"); first abstract lia. rewrite /counter; eauto 10.
 Qed.
 End proof.
diff --git a/heap_lang/lib/lock.v b/heap_lang/lib/lock.v
index bc755b1dba7c1d3e8816bef515d3689b8c8a6d8d..37d589a817c61fa4b07e986afd556dcf7c0d2646 100644
--- a/heap_lang/lib/lock.v
+++ b/heap_lang/lib/lock.v
@@ -1,7 +1,7 @@
-From iris.program_logic Require Export global_functor.
-From iris.proofmode Require Import invariants ghost_ownership.
+From iris.heap_lang Require Export lang.
+From iris.proofmode Require Import invariants tactics.
 From iris.heap_lang Require Import proofmode notation.
-Import uPred.
+From iris.algebra Require Import excl.
 
 Definition newlock : val := λ: <>, ref #false.
 Definition acquire : val :=
@@ -12,22 +12,21 @@ Global Opaque newlock acquire release.
 
 (** The CMRA we need. *)
 (* Not bundling heapG, as it may be shared with other users. *)
-Class lockG Σ := LockG { lock_tokG :> inG heap_lang Σ (exclR unitC) }.
+Class lockG Σ := LockG { lock_tokG :> inG Σ (exclR unitC) }.
 Definition lockGF : gFunctorList := [GFunctor (constRF (exclR unitC))].
-Instance inGF_lockG `{H : inGFs heap_lang Σ lockGF} : lockG Σ.
+Instance inGF_lockG `{H : inGFs Σ lockGF} : lockG Σ.
 Proof. destruct H. split. apply: inGF_inG. Qed.
 
 Section proof.
 Context `{!heapG Σ, !lockG Σ} (N : namespace).
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Definition lock_inv (γ : gname) (l : loc) (R : iProp) : iProp :=
+Definition lock_inv (γ : gname) (l : loc) (R : iProp Σ) : iProp Σ :=
   (∃ b : bool, l ↦ #b ★ if b then True else own γ (Excl ()) ★ R)%I.
 
-Definition is_lock (l : loc) (R : iProp) : iProp :=
+Definition is_lock (l : loc) (R : iProp Σ) : iProp Σ :=
   (∃ γ, heapN ⊥ N ∧ heap_ctx ∧ inv N (lock_inv γ l R))%I.
 
-Definition locked (l : loc) (R : iProp) : iProp :=
+Definition locked (l : loc) (R : iProp Σ) : iProp Σ :=
   (∃ γ, heapN ⊥ N ∧ heap_ctx ∧
         inv N (lock_inv γ l R) ∧ own γ (Excl ()))%I.
 
@@ -45,37 +44,36 @@ Proof. apply _. Qed.
 Lemma locked_is_lock l R : locked l R ⊢ is_lock l R.
 Proof. rewrite /is_lock. iDestruct 1 as (γ) "(?&?&?&_)"; eauto. Qed.
 
-Lemma newlock_spec (R : iProp) Φ :
+Lemma newlock_spec (R : iProp Σ) Φ :
   heapN ⊥ N →
   heap_ctx ★ R ★ (∀ l, is_lock l R -★ Φ #l) ⊢ WP newlock #() {{ Φ }}.
 Proof.
   iIntros (?) "(#Hh & HR & HΦ)". rewrite /newlock.
   wp_seq. wp_alloc l as "Hl".
-  iPvs (own_alloc (Excl ())) as (γ) "Hγ"; first done.
-  iPvs (inv_alloc N _ (lock_inv γ l R) with "[-HΦ]") as "#?"; first done.
+  iVs (own_alloc (Excl ())) as (γ) "Hγ"; first done.
+  iVs (inv_alloc N _ (lock_inv γ l R) with "[-HΦ]") as "#?".
   { iIntros ">". iExists false. by iFrame. }
-  iPvsIntro. iApply "HΦ". iExists γ; eauto.
+  iVsIntro. iApply "HΦ". iExists γ; eauto.
 Qed.
 
-Lemma acquire_spec l R (Φ : val → iProp) :
+Lemma acquire_spec l R (Φ : val → iProp Σ) :
   is_lock l R ★ (locked l R -★ R -★ Φ #()) ⊢ WP acquire #l {{ Φ }}.
 Proof.
   iIntros "[Hl HΦ]". iDestruct "Hl" as (γ) "(%&#?&#?)".
   iLöb as "IH". wp_rec. wp_focus (CAS _ _ _)%E.
-  iInv N as ([]) "[Hl HR]".
-  - wp_cas_fail. iPvsIntro; iSplitL "Hl".
-    + iNext. iExists true; eauto.
-    + wp_if. by iApply "IH".
-  - wp_cas_suc. iPvsIntro. iDestruct "HR" as "[Hγ HR]". iSplitL "Hl".
-    + iNext. iExists true; eauto.
-    + wp_if. iApply ("HΦ" with "[-HR] HR"). iExists γ; eauto.
+  iInv N as ([]) "[Hl HR]" "Hclose".
+  - wp_cas_fail. iVs ("Hclose" with "[Hl]"); first (iNext; iExists true; eauto).
+    iVsIntro. wp_if. by iApply "IH".
+  - wp_cas_suc. iDestruct "HR" as "[Hγ HR]".
+    iVs ("Hclose" with "[Hl]"); first (iNext; iExists true; eauto).
+    iVsIntro. wp_if. iApply ("HΦ" with "[-HR] HR"). iExists γ; eauto.
 Qed.
 
-Lemma release_spec R l (Φ : val → iProp) :
+Lemma release_spec R l (Φ : val → iProp Σ) :
   locked l R ★ R ★ Φ #() ⊢ WP release #l {{ Φ }}.
 Proof.
   iIntros "(Hl&HR&HΦ)"; iDestruct "Hl" as (γ) "(% & #? & #? & Hγ)".
-  rewrite /release. wp_let. iInv N as (b) "[Hl _]".
-  wp_store. iPvsIntro. iFrame "HΦ". iNext. iExists false. by iFrame.
+  rewrite /release. wp_let. iInv N as (b) "[Hl _]" "Hclose".
+  wp_store. iFrame "HΦ". iApply "Hclose". iNext. iExists false. by iFrame.
 Qed.
 End proof.
diff --git a/heap_lang/lib/par.v b/heap_lang/lib/par.v
index 1d9cca643ee1a94959fa85bfcb199cfad0ffd5d4..858edf73b5cdd1a796c4a5135ca1d61a573b5ed5 100644
--- a/heap_lang/lib/par.v
+++ b/heap_lang/lib/par.v
@@ -15,16 +15,15 @@ Global Opaque par.
 
 Section proof.
 Context `{!heapG Σ, !spawnG Σ}.
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Lemma par_spec (Ψ1 Ψ2 : val → iProp) e (f1 f2 : val) (Φ : val → iProp) :
+Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp Σ) :
   to_val e = Some (f1,f2)%V →
   (heap_ctx ★ WP f1 #() {{ Ψ1 }} ★ WP f2 #() {{ Ψ2 }} ★
    ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ ▷ Φ (v1,v2)%V)
   ⊢ WP par e {{ Φ }}.
 Proof.
   iIntros (?) "(#Hh&Hf1&Hf2&HΦ)".
-  rewrite /par. wp_value. iPvsIntro. wp_let. wp_proj.
+  rewrite /par. wp_value. iVsIntro. wp_let. wp_proj.
   wp_apply (spawn_spec parN); try wp_done; try solve_ndisj; iFrame "Hf1 Hh".
   iIntros (l) "Hl". wp_let. wp_proj. wp_focus (f2 _).
   iApply wp_wand_l; iFrame "Hf2"; iIntros (v) "H2". wp_let.
@@ -32,8 +31,8 @@ Proof.
   iSpecialize ("HΦ" with "* [-]"); first by iSplitL "H1". by wp_let.
 Qed.
 
-Lemma wp_par (Ψ1 Ψ2 : val → iProp)
-    (e1 e2 : expr) `{!Closed [] e1, Closed [] e2} (Φ : val → iProp) :
+Lemma wp_par (Ψ1 Ψ2 : val → iProp Σ)
+    (e1 e2 : expr) `{!Closed [] e1, Closed [] e2} (Φ : val → iProp Σ) :
   (heap_ctx ★ WP e1 {{ Ψ1 }} ★ WP e2 {{ Ψ2 }} ★
    ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ ▷ Φ (v1,v2)%V)
   ⊢ WP e1 || e2 {{ Φ }}.
diff --git a/heap_lang/lib/spawn.v b/heap_lang/lib/spawn.v
index 4c2617a7f09960df3b8926fb3318df376b1d8b6e..af61860e222ac7d1a4acfd2f4c580a61774e4ca0 100644
--- a/heap_lang/lib/spawn.v
+++ b/heap_lang/lib/spawn.v
@@ -1,7 +1,7 @@
-From iris.program_logic Require Export global_functor.
-From iris.proofmode Require Import invariants ghost_ownership.
+From iris.heap_lang Require Export lang.
+From iris.proofmode Require Import invariants tactics.
 From iris.heap_lang Require Import proofmode notation.
-Import uPred.
+From iris.algebra Require Import excl.
 
 Definition spawn : val :=
   λ: "f",
@@ -17,25 +17,22 @@ Global Opaque spawn join.
 
 (** The CMRA we need. *)
 (* Not bundling heapG, as it may be shared with other users. *)
-Class spawnG Σ := SpawnG {
-  spawn_tokG :> inG heap_lang Σ (exclR unitC);
-}.
+Class spawnG Σ := SpawnG { spawn_tokG :> inG Σ (exclR unitC) }.
 (** The functor we need. *)
 Definition spawnGF : gFunctorList := [GFunctor (constRF (exclR unitC))].
 (* Show and register that they match. *)
-Instance inGF_spawnG `{H : inGFs heap_lang Σ spawnGF} : spawnG Σ.
+Instance inGF_spawnG `{H : inGFs Σ spawnGF} : spawnG Σ.
 Proof. destruct H as (?&?). split. apply: inGF_inG. Qed.
 
 (** Now we come to the Iris part of the proof. *)
 Section proof.
 Context `{!heapG Σ, !spawnG Σ} (N : namespace).
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Definition spawn_inv (γ : gname) (l : loc) (Ψ : val → iProp) : iProp :=
+Definition spawn_inv (γ : gname) (l : loc) (Ψ : val → iProp Σ) : iProp Σ :=
   (∃ lv, l ↦ lv ★ (lv = NONEV ∨
                    ∃ v, lv = SOMEV v ★ (Ψ v ∨ own γ (Excl ()))))%I.
 
-Definition join_handle (l : loc) (Ψ : val → iProp) : iProp :=
+Definition join_handle (l : loc) (Ψ : val → iProp Σ) : iProp Σ :=
   (heapN ⊥ N ★ ∃ γ, heap_ctx ★ own γ (Excl ()) ★
                     inv N (spawn_inv γ l Ψ))%I.
 
@@ -49,7 +46,7 @@ Global Instance join_handle_ne n l :
 Proof. solve_proper. Qed.
 
 (** The main proofs. *)
-Lemma spawn_spec (Ψ : val → iProp) e (f : val) (Φ : val → iProp) :
+Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) (Φ : val → iProp Σ) :
   to_val e = Some f →
   heapN ⊥ N →
   heap_ctx ★ WP f #() {{ Ψ }} ★ (∀ l, join_handle l Ψ -★ Φ #l)
@@ -57,29 +54,27 @@ Lemma spawn_spec (Ψ : val → iProp) e (f : val) (Φ : val → iProp) :
 Proof.
   iIntros (<-%of_to_val ?) "(#Hh & Hf & HΦ)". rewrite /spawn.
   wp_let. wp_alloc l as "Hl". wp_let.
-  iPvs (own_alloc (Excl ())) as (γ) "Hγ"; first done.
-  iPvs (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?"; first done.
+  iVs (own_alloc (Excl ())) as (γ) "Hγ"; first done.
+  iVs (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?".
   { iNext. iExists NONEV. iFrame; eauto. }
   wp_apply wp_fork. iSplitR "Hf".
-  - iPvsIntro. wp_seq. iPvsIntro. iApply "HΦ". rewrite /join_handle. eauto.
+  - iVsIntro. wp_seq. iVsIntro. iApply "HΦ". rewrite /join_handle. eauto.
   - wp_focus (f _). iApply wp_wand_l. iFrame "Hf"; iIntros (v) "Hv".
-    iInv N as (v') "[Hl _]".
-    wp_store. iPvsIntro. iSplit; [iNext|done].
-    iExists (SOMEV v). iFrame. eauto.
+    iInv N as (v') "[Hl _]" "Hclose".
+    wp_store. iApply "Hclose". iNext. iExists (SOMEV v). iFrame. eauto.
 Qed.
 
-Lemma join_spec (Ψ : val → iProp) l (Φ : val → iProp) :
+Lemma join_spec (Ψ : val → iProp Σ) l (Φ : val → iProp Σ) :
   join_handle l Ψ ★ (∀ v, Ψ v -★ Φ v) ⊢ WP join #l {{ Φ }}.
 Proof.
   rewrite /join_handle; iIntros "[[% H] Hv]". iDestruct "H" as (γ) "(#?&Hγ&#?)".
-  iLöb as "IH". wp_rec. wp_focus (! _)%E. iInv N as (v) "[Hl Hinv]".
+  iLöb as "IH". wp_rec. wp_focus (! _)%E. iInv N as (v) "[Hl Hinv]" "Hclose".
   wp_load. iDestruct "Hinv" as "[%|Hinv]"; subst.
-  - iPvsIntro; iSplitL "Hl"; [iNext; iExists _; iFrame; eauto|].
-    wp_match. iApply ("IH" with "Hγ Hv").
+  - iVs ("Hclose" with "[Hl]"); [iNext; iExists _; iFrame; eauto|].
+    iVsIntro. wp_match. iApply ("IH" with "Hγ Hv").
   - iDestruct "Hinv" as (v') "[% [HΨ|Hγ']]"; simplify_eq/=.
-    + iPvsIntro; iSplitL "Hl Hγ".
-      { iNext. iExists _; iFrame; eauto. }
-      wp_match. by iApply "Hv".
+    + iVs ("Hclose" with "[Hl Hγ]"); [iNext; iExists _; iFrame; eauto|].
+      iVsIntro. wp_match. by iApply "Hv".
     + iCombine "Hγ" "Hγ'" as "Hγ". iDestruct (own_valid with "Hγ") as %[].
 Qed.
 End proof.
diff --git a/heap_lang/lib/ticket_lock.v b/heap_lang/lib/ticket_lock.v
index 495208d7b3a12ad28554dee62449ce904a64ae46..19094846864650174969f1cb40bbea13ff5ae571 100644
--- a/heap_lang/lib/ticket_lock.v
+++ b/heap_lang/lib/ticket_lock.v
@@ -30,35 +30,34 @@ Global Opaque newlock acquire release wait_loop.
 
 (** The CMRAs we need. *)
 Class tlockG Σ := TlockG {
-   tlock_G :> authG heap_lang Σ (gset_disjUR nat);
-   tlock_exclG  :> inG heap_lang Σ (exclR unitC)
+   tlock_G :> authG Σ (gset_disjUR nat);
+   tlock_exclG  :> inG Σ (exclR unitC)
 }.
 
 Definition tlockGF : gFunctorList :=
   [authGF (gset_disjUR nat); GFunctor (constRF (exclR unitC))].
-Instance inGF_tlockG `{H : inGFs heap_lang Σ tlockGF} : tlockG Σ.
+Instance inGF_tlockG `{H : inGFs Σ tlockGF} : tlockG Σ.
 Proof. destruct H as (? & ? & ?). split. apply _. apply: inGF_inG. Qed.
 
 Section proof.
 Context `{!heapG Σ, !tlockG Σ} (N : namespace).
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Definition tickets_inv (n: nat) (gs: gset_disjUR nat) : iProp :=
+Definition tickets_inv (n: nat) (gs: gset_disjUR nat) : iProp Σ :=
   (gs = GSet (seq_set 0 n))%I.
 
-Definition lock_inv (γ1 γ2: gname) (l : loc) (R : iProp) : iProp :=
+Definition lock_inv (γ1 γ2: gname) (l : loc) (R : iProp Σ) : iProp Σ :=
   (∃ o n: nat, l ↦ (#o, #n) ★
                auth_inv γ1 (tickets_inv n) ★
                ((own γ2 (Excl ()) ★  R) ∨ auth_own γ1 (GSet {[ o ]})))%I.
 
-Definition is_lock (l: loc) (R: iProp) : iProp :=
+Definition is_lock (l: loc) (R: iProp Σ) : iProp Σ :=
   (∃ γ1 γ2, heapN ⊥ N ∧ heap_ctx ∧ inv N (lock_inv γ1 γ2 l R))%I.
 
-Definition issued (l : loc) (x: nat) (R : iProp) : iProp :=
+Definition issued (l : loc) (x: nat) (R : iProp Σ) : iProp Σ :=
   (∃ γ1 γ2, heapN ⊥ N ∧ heap_ctx ∧ inv N (lock_inv γ1 γ2 l R) ∧
             auth_own γ1 (GSet {[ x ]}))%I.
 
-Definition locked (l : loc) (R : iProp) : iProp :=
+Definition locked (l : loc) (R : iProp Σ) : iProp Σ :=
   (∃ γ1 γ2, heapN ⊥ N ∧ heap_ctx ∧ inv N (lock_inv γ1 γ2 l R) ∧
             own γ2 (Excl ()))%I.
 
@@ -72,15 +71,15 @@ Proof. solve_proper. Qed.
 Global Instance is_lock_persistent l R : PersistentP (is_lock l R).
 Proof. apply _. Qed.
 
-Lemma newlock_spec (R : iProp) Φ :
+Lemma newlock_spec (R : iProp Σ) Φ :
   heapN ⊥ N →
   heap_ctx ★ R ★ (∀ l, is_lock l R -★ Φ #l) ⊢ WP newlock #() {{ Φ }}.
 Proof.
   iIntros (?) "(#Hh & HR & HΦ)". rewrite /newlock.
   wp_seq. wp_alloc l as "Hl".
-  iPvs (own_alloc (Excl ())) as (γ2) "Hγ2"; first done.
-  iPvs (own_alloc_strong (Auth (Excl' ∅) ∅) _ {[ γ2 ]}) as (γ1) "[% Hγ1]"; first done.
-  iPvs (inv_alloc N _ (lock_inv γ1 γ2 l R) with "[-HΦ]"); first done.
+  iVs (own_alloc (Excl ())) as (γ2) "Hγ2"; first done.
+  iVs (own_alloc_strong (Auth (Excl' ∅) ∅) {[ γ2 ]}) as (γ1) "[% Hγ1]"; first done.
+  iVs (inv_alloc N _ (lock_inv γ1 γ2 l R) with "[-HΦ]").
   { iNext. rewrite /lock_inv.
     iExists 0%nat, 0%nat.
     iFrame.
@@ -90,106 +89,94 @@ Proof.
       by iFrame. }
     iLeft.
     by iFrame. }
-  iPvsIntro.
+  iVsIntro.
   iApply "HΦ".
   iExists γ1, γ2.
   iSplit; by auto.
 Qed.
 
-Lemma wait_loop_spec l x R (Φ : val → iProp) :
+Lemma wait_loop_spec l x R (Φ : val → iProp Σ) :
   issued l x R ★ (∀ l, locked l R -★ R -★ Φ #()) ⊢ WP wait_loop #x #l {{ Φ }}.
 Proof.
   iIntros "[Hl HΦ]". iDestruct "Hl" as (γ1 γ2) "(% & #? & #? & Ht)".
   iLöb as "IH". wp_rec. wp_let. wp_focus (! _)%E.
-  iInv N as (o n) "[Hl Ha]".
-  wp_load. iPvsIntro.
-  destruct (decide (x = o)) as [Heq|Hneq].
-  - subst.
-    iDestruct "Ha" as "[Hainv [[Ho HR] | Haown]]".
-    + iSplitL "Hl Hainv Ht".
-      * iNext.
-        iExists o, n.
-        iFrame.
-        by iRight.
-      * wp_proj. wp_let. wp_op=>[_|[]] //.
-        wp_if. iPvsIntro.
-        iApply ("HΦ" with "[-HR] HR"). iExists γ1, γ2; eauto.
+  iInv N as (o n) "[Hl Ha]" "Hclose".
+  wp_load. destruct (decide (x = o)) as [->|Hneq].
+  - iDestruct "Ha" as "[Hainv [[Ho HR] | Haown]]".
+    + iVs ("Hclose" with "[Hl Hainv Ht]").
+      { iNext. iExists o, n. iFrame. eauto. }
+      iVsIntro. wp_proj. wp_let. wp_op=>[_|[]] //.
+      wp_if. iVsIntro.
+      iApply ("HΦ" with "[-HR] HR"). iExists γ1, γ2; eauto.
     + iExFalso. iCombine "Ht" "Haown" as "Haown".
       iDestruct (auth_own_valid with "Haown") as % ?%gset_disj_valid_op.
       set_solver.
-  - iSplitL "Hl Ha".
-    + iNext. iExists o, n. by iFrame.
-    + wp_proj. wp_let. wp_op=>?; first omega.
-      wp_if. by iApply ("IH" with "Ht").
+  - iVs ("Hclose" with "[Hl Ha]").
+    { iNext. iExists o, n. by iFrame. }
+    iVsIntro. wp_proj. wp_let. wp_op=>?; first omega.
+    wp_if. by iApply ("IH" with "Ht").
 Qed.
 
-Lemma acquire_spec l R (Φ : val → iProp) :
+Lemma acquire_spec l R (Φ : val → iProp Σ) :
   is_lock l R ★ (∀ l, locked l R -★ R -★ Φ #()) ⊢ WP acquire #l {{ Φ }}.
 Proof.
   iIntros "[Hl HΦ]". iDestruct "Hl" as (γ1 γ2) "(% & #? & #?)".
   iLöb as "IH". wp_rec. wp_focus (! _)%E.
-  iInv N as (o n) "[Hl Ha]".
-  wp_load. iPvsIntro.
-  iSplitL "Hl Ha".
-  - iNext. iExists o, n. by iFrame.
-  - wp_let. wp_proj. wp_proj. wp_op.
-    wp_focus (CAS _ _ _).
-    iInv N as (o' n') "[Hl [Hainv Haown]]".
-    destruct (decide ((#o', #n') = (#o, #n)))%V
-      as [[= ->%Nat2Z.inj ->%Nat2Z.inj] | Hneq].
-    + wp_cas_suc.
-      iDestruct "Hainv" as (s) "[Ho %]"; subst.
-      iPvs (own_update with "Ho") as "Ho".
-      { eapply auth_update_no_frag, (gset_alloc_empty_local_update n).
-        rewrite elem_of_seq_set; omega. }
-      iDestruct "Ho" as "[Hofull Hofrag]".
-      iSplitL "Hl Haown Hofull".
-      * rewrite gset_disj_union; last by apply (seq_set_S_disjoint 0).
-        rewrite -(seq_set_S_union_L 0).
-        iPvsIntro. iNext.
-        iExists o, (S n)%nat.
-        rewrite Nat2Z.inj_succ -Z.add_1_r.
-        iFrame. iExists (GSet (seq_set 0 (S n))).
-        by iFrame.
-      * iPvsIntro. wp_if. wp_proj.
-        iApply wait_loop_spec.
-        iSplitR "HΦ"; last by done.
-        rewrite /issued /auth_own; eauto 10.
-    + wp_cas_fail.
-      iPvsIntro.
-      iSplitL "Hl Hainv Haown".
-      { iNext. iExists o', n'. by iFrame. }
-      { wp_if. by iApply "IH". }
+  iInv N as (o n) "[Hl Ha]" "Hclose".
+  wp_load. iVs ("Hclose" with "[Hl Ha]").
+  { iNext. iExists o, n. by iFrame. }
+  iVsIntro. wp_let. wp_proj. wp_proj. wp_op.
+  wp_focus (CAS _ _ _).
+  iInv N as (o' n') "[Hl [Hainv Haown]]" "Hclose".
+  destruct (decide ((#o', #n') = (#o, #n)))%V
+    as [[= ->%Nat2Z.inj ->%Nat2Z.inj] | Hneq].
+  - wp_cas_suc.
+    iDestruct "Hainv" as (s) "[Ho %]"; subst.
+    iVs (own_update with "Ho") as "Ho".
+    { eapply auth_update_no_frag, (gset_alloc_empty_local_update n).
+      rewrite elem_of_seq_set; omega. }
+    iDestruct "Ho" as "[Hofull Hofrag]".
+    iVs ("Hclose" with "[Hl Haown Hofull]").
+    { rewrite gset_disj_union; last by apply (seq_set_S_disjoint 0).
+      rewrite -(seq_set_S_union_L 0).
+      iNext. iExists o, (S n)%nat.
+      rewrite Nat2Z.inj_succ -Z.add_1_r.
+      iFrame. iExists (GSet (seq_set 0 (S n))). by iFrame. }
+    iVsIntro. wp_if. wp_proj.
+    iApply wait_loop_spec.
+    iSplitR "HΦ"; last by done.
+    rewrite /issued /auth_own; eauto 10.
+  - wp_cas_fail.
+    iVs ("Hclose" with "[Hl Hainv Haown]").
+    { iNext. iExists o', n'. by iFrame. }
+    iVsIntro. wp_if. by iApply "IH".
 Qed.
 
-Lemma release_spec R l (Φ : val → iProp):
+Lemma release_spec R l (Φ : val → iProp Σ):
   locked l R ★ R ★ Φ #() ⊢ WP release #l {{ Φ }}.
 Proof.
   iIntros "(Hl & HR & HΦ)"; iDestruct "Hl" as (γ1 γ2) "(% & #? & #? & Hγ)".
   iLöb as "IH". wp_rec. wp_focus (! _)%E.
-  iInv N as (o n) "[Hl Hr]".
-  wp_load. iPvsIntro.
-  iSplitL "Hl Hr".
-  - iNext. iExists o, n. by iFrame.
-  - wp_let. wp_focus (CAS _ _ _ ).
-    wp_proj. wp_op. wp_proj.
-    iInv N as (o' n') "[Hl Hr]".
-    destruct (decide ((#o', #n') = (#o, #n)))%V
-      as [[= ->%Nat2Z.inj ->%Nat2Z.inj] | Hneq].
-    + wp_cas_suc.
-      iDestruct "Hr" as "[Hainv [[Ho _] | Hown]]".
-      * iExFalso. iCombine "Hγ" "Ho" as "Ho".
-        iDestruct (own_valid with "#Ho") as %[].
-      * iSplitL "Hl HR Hγ Hainv".
-        { iPvsIntro. iNext. iExists (o + 1)%nat, n%nat.
-          iFrame. rewrite Nat2Z.inj_add.
-          iFrame. iLeft; by iFrame. }
-        { iPvsIntro. by wp_if. }
-    + wp_cas_fail.
-      iPvsIntro.
-      iSplitL "Hl Hr".
-      * iNext. iExists o', n'. by iFrame.
-      * wp_if. by iApply ("IH" with "Hγ HR").
+  iInv N as (o n) "[Hl Hr]" "Hclose".
+  wp_load. iVs ("Hclose" with "[Hl Hr]").
+  { iNext. iExists o, n. by iFrame. }
+  iVsIntro. wp_let. wp_focus (CAS _ _ _ ).
+  wp_proj. wp_op. wp_proj.
+  iInv N as (o' n') "[Hl Hr]" "Hclose".
+  destruct (decide ((#o', #n') = (#o, #n)))%V
+    as [[= ->%Nat2Z.inj ->%Nat2Z.inj] | Hneq].
+  - wp_cas_suc.
+    iDestruct "Hr" as "[Hainv [[Ho _] | Hown]]".
+    + iExFalso. iCombine "Hγ" "Ho" as "Ho".
+      iDestruct (own_valid with "#Ho") as %[].
+    + iVs ("Hclose" with "[Hl HR Hγ Hainv]").
+      { iNext. iExists (o + 1)%nat, n%nat.
+        iFrame. rewrite Nat2Z.inj_add.
+        iFrame. iLeft; by iFrame. }
+      iVsIntro. by wp_if.
+  - wp_cas_fail. iVs ("Hclose" with "[Hl Hr]").
+    { iNext. iExists o', n'. by iFrame. }
+    iVsIntro. wp_if. by iApply ("IH" with "Hγ HR").
 Qed.
 End proof.
 
diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v
index 24eb4d3a6c6fe8c986870c08711745eea10f24e1..2df6ed4baa45a83374bd7a7641724846dc955af2 100644
--- a/heap_lang/lifting.v
+++ b/heap_lang/lifting.v
@@ -7,9 +7,9 @@ Import uPred.
 Local Hint Extern 0 (head_reducible _ _) => do_head_step eauto 2.
 
 Section lifting.
-Context {Σ : iFunctor}.
-Implicit Types P Q : iProp heap_lang Σ.
-Implicit Types Φ : val → iProp heap_lang Σ.
+Context `{irisG heap_lang Σ}.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val → iProp Σ.
 Implicit Types ef : option expr.
 
 (** Bind. This bundles some arguments that wp_ectx_bind leaves as indices. *)
@@ -28,7 +28,7 @@ Lemma wp_alloc_pst E σ v Φ :
   ⊢ WP Alloc (of_val v) @ E {{ Φ }}.
 Proof.
   iIntros "[HP HΦ]".
-  iApply (wp_lift_atomic_head_step (Alloc (of_val v)) σ); eauto with fsaV.
+  iApply (wp_lift_atomic_head_step (Alloc (of_val v)) σ); eauto.
   iFrame "HP". iNext. iIntros (v2 σ2 ef) "[% HP]". inv_head_step.
   match goal with H: _ = of_val v2 |- _ => apply (inj of_val (LitV _)) in H end.
   subst v2. iSplit; last done. iApply "HΦ"; by iSplit.
diff --git a/heap_lang/notation.v b/heap_lang/notation.v
index 05e173432ec7b604ec1ba7fbc5f8a765c6edd02b..45e6932a429dd813e5fda1eb0e564c03f79e73fe 100644
--- a/heap_lang/notation.v
+++ b/heap_lang/notation.v
@@ -1,7 +1,7 @@
 From iris.heap_lang Require Export derived.
 Export heap_lang.
 
-Arguments wp {_ _} _ _%E _.
+Arguments wp {_ _ _} _ _%E _.
 
 Notation "'WP' e @ E {{ Φ } }" := (wp E e%E Φ)
   (at level 20, e, Φ at level 200,
diff --git a/heap_lang/proofmode.v b/heap_lang/proofmode.v
index 1fbd3665160f85e9a7ac34a2621a500d6ea693c7..95b2503e3de27ed278342de51865bbb1503025a6 100644
--- a/heap_lang/proofmode.v
+++ b/heap_lang/proofmode.v
@@ -7,9 +7,9 @@ Ltac wp_strip_later ::= iNext.
 
 Section heap.
 Context `{heapG Σ}.
-Implicit Types P Q : iPropG heap_lang Σ.
-Implicit Types Φ : val → iPropG heap_lang Σ.
-Implicit Types Δ : envs (iResUR heap_lang (globalF Σ)).
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val → iProp Σ.
+Implicit Types Δ : envs (iResUR Σ).
 
 Global Instance into_sep_mapsto l q v :
   IntoSep false (l ↦{q} v) (l ↦{q/2} v) (l ↦{q/2} v).
diff --git a/heap_lang/tactics.v b/heap_lang/tactics.v
index 5d847b13308bb3e97edbe601ec5492411143640f..dbe50971a7de1be0598f64ece4dc4490c5799e76 100644
--- a/heap_lang/tactics.v
+++ b/heap_lang/tactics.v
@@ -229,7 +229,9 @@ Ltac solve_atomic :=
      let e' := W.of_expr e in change (language.atomic (W.to_expr e'));
      apply W.atomic_correct; vm_compute; exact I
   end.
-Hint Extern 0 (language.atomic _) => solve_atomic : fsaV.
+Hint Extern 10 (language.atomic _) => solve_atomic.
+(* For the side-condition of elim_vs_pvs_wp_atomic *)
+Hint Extern 10 (language.atomic _) => solve_atomic : typeclass_instances.
 
 (** Substitution *)
 Ltac simpl_subst :=
diff --git a/heap_lang/wp_tactics.v b/heap_lang/wp_tactics.v
index e84f90afecc010c45cfba8082db4859028aaaebc..9679cbfb525187bbe46edb3635c2625f52b95a64 100644
--- a/heap_lang/wp_tactics.v
+++ b/heap_lang/wp_tactics.v
@@ -25,7 +25,11 @@ Ltac wp_strip_pvs :=
   lazymatch goal with
   | |- _ ⊢ |={?E}=> _ =>
     etrans; [|apply pvs_intro];
-    match goal with |- _ ⊢ wp E _ _ => simpl | _ => fail end
+    match goal with
+    | |- _ ⊢ wp E _ _ => simpl
+    | |- _ ⊢ |={E,_}=> _ => simpl
+    | _ => fail
+    end
   end.
 
 Ltac wp_value_head := etrans; [|eapply wp_value_pvs; wp_done]; lazy beta.
@@ -45,7 +49,7 @@ Ltac wp_finish := intros_revert ltac:(
   | |- _ ⊢ wp ?E (Seq _ _) ?Q =>
      etrans; [|eapply wp_seq; wp_done]; wp_strip_later
   | |- _ ⊢ wp ?E _ ?Q => wp_value_head
-  | |- _ ⊢ |={_}=> _ => wp_strip_pvs
+  | |- _ ⊢ |={_,_}=> _ => wp_strip_pvs
   end).
 
 Tactic Notation "wp_value" :=
diff --git a/program_logic/adequacy.v b/program_logic/adequacy.v
index 2acb580466368a5c18e85b2afb83b2531a58eb96..df8a8fac51c8777dcb6d8370ea957826ca97907b 100644
--- a/program_logic/adequacy.v
+++ b/program_logic/adequacy.v
@@ -1,155 +1,142 @@
-From iris.program_logic Require Export hoare.
-From iris.program_logic Require Import wsat ownership.
-Local Hint Extern 10 (_ ≤ _) => omega.
-Local Hint Extern 100 (_ ⊥ _) => set_solver.
-Local Hint Extern 10 (✓{_} _) =>
-  repeat match goal with
-  | H : wsat _ _ _ _ |- _ => apply wsat_valid in H; last omega
-  end; solve_validN.
+From iris.program_logic Require Export weakestpre.
+From iris.algebra Require Import gmap auth agree gset coPset upred_big_op.
+From iris.program_logic Require Import ownership.
+From iris.proofmode Require Import tactics weakestpre.
+Import uPred.
 
 Section adequacy.
-Context {Λ : language} {Σ : iFunctor}.
+Context `{irisG Λ Σ}.
 Implicit Types e : expr Λ.
-Implicit Types P Q : iProp Λ Σ.
-Implicit Types Φ : val Λ → iProp Λ Σ.
-Implicit Types Φs : list (val Λ → iProp Λ Σ).
-Implicit Types m : iGst Λ Σ.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val Λ → iProp Σ.
+Implicit Types Φs : list (val Λ → iProp Σ).
 
-Notation wptp n := (Forall3 (λ e Φ r, uPred_holds (wp ⊤ e Φ) n r)).
-Lemma wptp_le Φs es rs n n' :
-  ✓{n'} (big_op rs) → wptp n es Φs rs → n' ≤ n → wptp n' es Φs rs.
-Proof. induction 2; constructor; eauto using uPred_closed. Qed.
-Lemma nsteps_wptp Φs k n tσ1 tσ2 rs1 :
-  nsteps step k tσ1 tσ2 →
-  1 < n → wptp (k + n) (tσ1.1) Φs rs1 →
-  wsat (k + n) ⊤ (tσ1.2) (big_op rs1) →
-  ∃ rs2 Φs', wptp n (tσ2.1) (Φs ++ Φs') rs2 ∧ wsat n ⊤ (tσ2.2) (big_op rs2).
+Notation world σ := (wsat ★ ownE ⊤ ★ ownP_auth σ)%I.
+
+Definition wptp (t : list (expr Λ)) := ([★] (flip (wp ⊤) (λ _, True) <$> t))%I.
+
+Lemma wptp_cons e t : wptp (e :: t) ⊣⊢ WP e {{ _, True }} ★ wptp t.
+Proof. done. Qed.
+Lemma wptp_app t1 t2 : wptp (t1 ++ t2) ⊣⊢ wptp t1 ★ wptp t2.
+Proof. by rewrite /wptp fmap_app big_sep_app. Qed.
+Lemma wptp_fork ef : wptp (option_list ef) ⊣⊢ wp_fork ef.
+Proof. destruct ef; by rewrite /wptp /= ?right_id. Qed.
+
+Lemma wp_step e1 σ1 e2 σ2 ef Φ :
+  prim_step e1 σ1 e2 σ2 ef →
+  world σ1 ★ WP e1 {{ Φ }} =r=> ▷ |=r=> ◇ (world σ2 ★ WP e2 {{ Φ }} ★ wp_fork ef).
+Proof.
+  rewrite {1}wp_unfold /wp_pre. iIntros (Hstep) "[(Hw & HE & Hσ) [H|[_ H]]]".
+  { iDestruct "H" as (v) "[% _]". apply val_stuck in Hstep; simplify_eq. }
+  rewrite pvs_eq /pvs_def.
+  iVs ("H" $! σ1 with "Hσ [Hw HE]") as "[H|(Hw & HE & _ & H)]"; first by iFrame.
+  { iVsIntro; iNext. by iExFalso. }
+  iVsIntro; iNext.
+  by iVs ("H" $! e2 σ2 ef with "[%] [Hw HE]") as "[?|($ & $ & $ & $)]";
+    [done|by iFrame|rewrite /uPred_now_True; eauto|].
+Qed.
+
+Lemma wptp_step e1 t1 t2 σ1 σ2 Φ :
+  step (e1 :: t1,σ1) (t2, σ2) →
+  world σ1 ★ WP e1 {{ Φ }} ★ wptp t1
+  =r=> ∃ e2 t2', t2 = e2 :: t2' ★ ▷ |=r=> ◇ (world σ2 ★ WP e2 {{ Φ }} ★ wptp t2').
+Proof.
+  iIntros (Hstep) "(HW & He & Ht)".
+  destruct Hstep as [e1' σ1' e2' σ2' ef [|? t1'] t2' ?? Hstep]; simplify_eq/=.
+  - iExists e2', (t2' ++ option_list ef); iSplitR; first eauto.
+    rewrite wptp_app wptp_fork. iFrame "Ht". iApply wp_step; try iFrame; eauto.
+  - iExists e, (t1' ++ e2' :: t2' ++ option_list ef); iSplitR; first eauto.
+    rewrite !wptp_app !wptp_cons wptp_app wptp_fork.
+    iDestruct "Ht" as "($ & He' & $)"; iFrame "He".
+    iApply wp_step; try iFrame; eauto.
+Qed.
+
+Lemma wptp_steps n e1 t1 t2 σ1 σ2 Φ :
+  nsteps step n (e1 :: t1, σ1) (t2, σ2) →
+  world σ1 ★ WP e1 {{ Φ }} ★ wptp t1 ⊢
+  Nat.iter (S n) (λ P, |=r=> ▷ P) (∃ e2 t2',
+    t2 = e2 :: t2' ★ world σ2 ★ WP e2 {{ Φ }} ★ wptp t2').
 Proof.
-  intros Hsteps Hn; revert Φs rs1.
-  induction Hsteps as [|k ?? tσ3 [e1 σ1 e2 σ2 ef t1 t2 ?? Hstep] Hsteps IH];
-    simplify_eq/=; intros Φs rs.
-  { by intros; exists rs, []; rewrite right_id_L. }
-  intros (Φs1&?&rs1&?&->&->&?&
-    (Φ&Φs2&r&rs2&->&->&Hwp&?)%Forall3_cons_inv_l)%Forall3_app_inv_l ?.
-  rewrite wp_eq in Hwp.
-  destruct (wp_step_inv ⊤ ∅ Φ e1 (k + n) (S (k + n)) σ1 r
-    (big_op (rs1 ++ rs2))) as [_ Hwpstep]; eauto using val_stuck.
-  { by rewrite right_id_L -big_op_cons Permutation_middle. }
-  destruct (Hwpstep e2 σ2 ef) as (r2&r2'&Hwsat&?&?); auto; clear Hwpstep.
-  revert Hwsat; rewrite big_op_app right_id_L=>Hwsat.
-  destruct ef as [e'|].
-  - destruct (IH (Φs1 ++ Φ :: Φs2 ++ [λ _, True%I])
-      (rs1 ++ r2 :: rs2 ++ [r2'])) as (rs'&Φs'&?&?).
-    { apply Forall3_app, Forall3_cons,
-        Forall3_app, Forall3_cons, Forall3_nil; eauto using wptp_le; [|];
-      rewrite wp_eq; eauto. }
-    { by rewrite -Permutation_middle /= (assoc (++))
-        (comm (++)) /= assoc big_op_app. }
-    exists rs', ([λ _, True%I] ++ Φs'); split; auto.
-    by rewrite (assoc _ _ _ Φs') -(assoc _ Φs1).
-  - apply (IH (Φs1 ++ Φ :: Φs2) (rs1 ++ r2 ⋅ r2' :: rs2)).
-    { rewrite /option_list right_id_L.
-      apply Forall3_app, Forall3_cons; eauto using wptp_le.
-      rewrite wp_eq.
-      apply uPred_closed with (k + n);
-        first apply uPred_mono with r2; eauto using cmra_includedN_l. }
-    by rewrite -Permutation_middle /= big_op_app.
+  revert e1 t1 t2 σ1 σ2; simpl; induction n as [|n IH]=> e1 t1 t2 σ1 σ2 /=.
+  { inversion_clear 1; iIntros "?"; eauto 10. }
+  iIntros (Hsteps) "H". inversion_clear Hsteps as [|?? [t1' σ1']].
+  iVs (wptp_step with "H") as (e1' t1'') "[% H]"; first eauto; simplify_eq.
+  iVsIntro; iNext; iVs "H" as "[?|?]"; first (iVsIntro; iNext; by iExFalso).
+  by iApply IH.
 Qed.
-Lemma wp_adequacy_steps P Φ k n e1 t2 σ1 σ2 r1 :
-  (P ⊢ WP e1 {{ Φ }}) →
-  nsteps step k ([e1],σ1) (t2,σ2) →
-  1 < n → wsat (k + n) ⊤ σ1 r1 →
-  P (k + n) r1 →
-  ∃ rs2 Φs', wptp n t2 (Φ :: Φs') rs2 ∧ wsat n ⊤ σ2 (big_op rs2).
+
+Instance rvs_iter_mono n : Proper ((⊢) ==> (⊢)) (Nat.iter n (λ P, |=r=> ▷ P)%I).
+Proof. intros P Q HP. induction n; simpl; do 2?f_equiv; auto. Qed.
+
+Lemma wptp_result n e1 t1 v2 t2 σ1 σ2 φ :
+  nsteps step n (e1 :: t1, σ1) (of_val v2 :: t2, σ2) →
+  world σ1 ★ WP e1 {{ v, ■ φ v }} ★ wptp t1 ⊢
+  Nat.iter (S (S n)) (λ P, |=r=> ▷ P) (■ φ v2).
 Proof.
-  intros Hht ????; apply (nsteps_wptp [Φ] k n ([e1],σ1) (t2,σ2) [r1]);
-    rewrite /big_op ?right_id; auto.
-  constructor; last constructor.
-  apply Hht; by eauto using cmra_included_core.
+  intros. rewrite wptp_steps // (Nat_iter_S_r (S n)); apply rvs_iter_mono.
+  iDestruct 1 as (e2 t2') "(% & (Hw & HE & _) & H & _)"; simplify_eq.
+  iDestruct (wp_value_inv with "H") as "H". rewrite pvs_eq /pvs_def.
+  iVs ("H" with "[Hw HE]") as "[?|(_ & _ & $)]"; [by iFrame| |done].
+  iVsIntro; iNext; by iExFalso.
 Qed.
 
-Lemma wp_adequacy_own Φ e1 t2 σ1 m σ2 :
-  ✓ m →
-  (ownP σ1 ★ ownG m ⊢ WP e1 {{ Φ }}) →
-  rtc step ([e1],σ1) (t2,σ2) →
-  ∃ rs2 Φs', wptp 2 t2 (Φ :: Φs') rs2 ∧ wsat 2 ⊤ σ2 (big_op rs2).
+Lemma wp_safe e σ Φ :
+  world σ ★ WP e {{ Φ }} =r=> ▷ ■ (is_Some (to_val e) ∨ reducible e σ).
 Proof.
-  intros Hv ? [k ?]%rtc_nsteps.
-  eapply wp_adequacy_steps with (r1 := (Res ∅ (Excl' σ1) m)); eauto; [|].
-  { by rewrite Nat.add_comm; apply wsat_init, cmra_valid_validN. }
-  uPred.unseal; exists (Res ∅ (Excl' σ1) ∅), (Res ∅ ∅ m); split_and?.
-  - by rewrite Res_op ?left_id ?right_id.
-  - rewrite /ownP; uPred.unseal; rewrite /uPred_holds //=.
-  - by apply ownG_spec.
+  rewrite wp_unfold /wp_pre. iIntros "[(Hw&HE&Hσ) [H|[_ H]]]".
+  { iDestruct "H" as (v) "[% _]"; eauto 10. }
+  rewrite pvs_eq. iVs ("H" with "* Hσ [-]") as "[H|(?&?&%&?)]"; first by iFrame.
+  iVsIntro; iNext; by iExFalso. eauto 10.
 Qed.
 
-Theorem wp_adequacy_result E φ e v t2 σ1 m σ2 :
-  ✓ m →
-  (ownP σ1 ★ ownG m ⊢ WP e @ E {{ v', ■ φ v' }}) →
-  rtc step ([e], σ1) (of_val v :: t2, σ2) →
-  φ v.
+Lemma wptp_safe n e1 e2 t1 t2 σ1 σ2 Φ :
+  nsteps step n (e1 :: t1, σ1) (t2, σ2) → e2 ∈ t2 →
+  world σ1 ★ WP e1 {{ Φ }} ★ wptp t1 ⊢
+  Nat.iter (S (S n)) (λ P, |=r=> ▷ P) (■ (is_Some (to_val e2) ∨ reducible e2 σ2)).
 Proof.
-  intros Hv ? Hs.
-  destruct (wp_adequacy_own (λ v', ■ φ v')%I e (of_val v :: t2) σ1 m σ2)
-             as (rs2&Qs&Hwptp&?); auto.
-  { by rewrite -(wp_mask_weaken E ⊤). }
-  inversion Hwptp as [|?? r ?? rs Hwp _]; clear Hwptp; subst.
-  move: Hwp. rewrite wp_eq. uPred.unseal=> /wp_value_inv Hwp.
-  rewrite pvs_eq in Hwp.
-  destruct (Hwp 2 ∅ σ2 (big_op rs)) as [r' []]; rewrite ?right_id_L; auto.
+  intros ? He2. rewrite wptp_steps // (Nat_iter_S_r (S n)); apply rvs_iter_mono.
+  iDestruct 1 as (e2' t2') "(% & Hw & H & Htp)"; simplify_eq.
+  apply elem_of_cons in He2 as [<-|?]; first (iApply wp_safe; by iFrame "Hw H").
+  iApply wp_safe. iFrame "Hw".
+  iApply (big_sep_elem_of with "Htp"); apply elem_of_list_fmap; eauto.
 Qed.
+End adequacy.
 
-Lemma ht_adequacy_result E φ e v t2 σ1 m σ2 :
-  ✓ m →
-  {{ ownP σ1 ★ ownG m }} e @ E {{ v', ■ φ v' }} →
-  rtc step ([e], σ1) (of_val v :: t2, σ2) →
-  φ v.
+Theorem adequacy_result `{irisPreG Λ Σ} e v2 t2 σ1 σ2 φ :
+  (∀ `{irisG Λ Σ}, ownP σ1 ⊢ WP e {{ v, ■ φ v }}) →
+  rtc step ([e], σ1) (of_val v2 :: t2, σ2) →
+  φ v2.
 Proof.
-  intros ? Hht. eapply wp_adequacy_result with (E:=E); first done.
-  move:Hht. by rewrite /ht uPred.always_elim=>/uPred.impl_entails.
+  intros Hwp [n ?]%rtc_nsteps.
+  eapply (adequacy (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
+  rewrite Nat_iter_S. iVs (iris_alloc σ1) as (?) "(?&?&?&Hσ)".
+  iVsIntro. iNext. iApply wptp_result; eauto.
+  iFrame. iSplitL. by iApply Hwp. done.
 Qed.
 
-Lemma wp_adequacy_reducible E Φ e1 e2 t2 σ1 m σ2 :
-  ✓ m →
-  (ownP σ1 ★ ownG m ⊢ WP e1 @ E {{ Φ }}) →
+Lemma wp_adequacy_reducible `{irisPreG Λ Σ} e1 e2 t2 σ1 σ2 Φ :
+  (∀ `{irisG Λ Σ}, ownP σ1 ⊢ WP e1 {{ Φ }}) →
   rtc step ([e1], σ1) (t2, σ2) →
   e2 ∈ t2 → (is_Some (to_val e2) ∨ reducible e2 σ2).
 Proof.
-  intros Hv ? Hs [i ?]%elem_of_list_lookup.
-  destruct (wp_adequacy_own Φ e1 t2 σ1 m σ2) as (rs2&Φs&?&?); auto.
-  { by rewrite -(wp_mask_weaken E ⊤). }
-  destruct (Forall3_lookup_l (λ e Φ r, wp ⊤ e Φ 2 r) t2
-    (Φ :: Φs) rs2 i e2) as (Φ'&r2&?&?&Hwp); auto.
-  case He:(to_val e2)=>[v2|]; first by eauto. right.
-  destruct (wp_step_inv ⊤ ∅ Φ' e2 1 2 σ2 r2 (big_op (delete i rs2)));
-    rewrite ?right_id_L ?big_op_delete; auto.
-  by rewrite -wp_eq.
+  intros Hwp [n ?]%rtc_nsteps ?.
+  eapply (adequacy (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
+  rewrite Nat_iter_S. iVs (iris_alloc σ1) as (?) "(Hw & HE & Hσ & Hσf)".
+  iVsIntro. iNext. iApply wptp_safe; eauto.
+  iFrame "Hw HE Hσ". iSplitL. by iApply Hwp. done.
 Qed.
 
-(* Connect this up to the threadpool-semantics. *)
-Theorem wp_adequacy_safe E Φ e1 t2 σ1 m σ2 :
-  ✓ m →
-  (ownP σ1 ★ ownG m ⊢ WP e1 @ E {{ Φ }}) →
+Theorem wp_adequacy_safe `{irisPreG Λ Σ} e1 t2 σ1 σ2 Φ :
+  (∀ `{irisG Λ Σ}, ownP σ1 ⊢ WP e1 {{ Φ }}) →
   rtc step ([e1], σ1) (t2, σ2) →
   Forall (λ e, is_Some (to_val e)) t2 ∨ ∃ t3 σ3, step (t2, σ2) (t3, σ3).
 Proof.
   intros.
   destruct (decide (Forall (λ e, is_Some (to_val e)) t2)) as [|Ht2]; [by left|].
   apply (not_Forall_Exists _), Exists_exists in Ht2; destruct Ht2 as (e2&?&He2).
-  destruct (wp_adequacy_reducible E Φ e1 e2 t2 σ1 m σ2) as [?|(e3&σ3&ef&?)];
+  destruct (wp_adequacy_reducible e1 e2 t2 σ1 σ2 Φ) as [?|(e3&σ3&ef&?)];
     rewrite ?eq_None_not_Some; auto.
   { exfalso. eauto. }
   destruct (elem_of_list_split t2 e2) as (t2'&t2''&->); auto.
   right; exists (t2' ++ e3 :: t2'' ++ option_list ef), σ3; econstructor; eauto.
 Qed.
-
-Lemma ht_adequacy_safe E Φ e1 t2 σ1 m σ2 :
-  ✓ m →
-  {{ ownP σ1 ★ ownG m }} e1 @ E {{ Φ }} →
-  rtc step ([e1], σ1) (t2, σ2) →
-  Forall (λ e, is_Some (to_val e)) t2 ∨ ∃ t3 σ3, step (t2, σ2) (t3, σ3).
-Proof.
-  intros ? Hht. eapply wp_adequacy_safe with (E:=E) (Φ:=Φ); first done.
-  move:Hht. by rewrite /ht uPred.always_elim=>/uPred.impl_entails.
-Qed.
-End adequacy.
diff --git a/program_logic/auth.v b/program_logic/auth.v
index c5b8dc22f92a626a98b40c17db51e5ac4378de9c..1c68c07ffa4640e48df49868093941aea681ae40 100644
--- a/program_logic/auth.v
+++ b/program_logic/auth.v
@@ -1,27 +1,27 @@
 From iris.algebra Require Export auth upred_tactics.
+From iris.algebra Require Import gmap.
 From iris.program_logic Require Export invariants ghost_ownership.
 From iris.proofmode Require Import invariants ghost_ownership.
 Import uPred.
 
 (* The CMRA we need. *)
-Class authG Λ Σ (A : ucmraT) := AuthG {
-  auth_inG :> inG Λ Σ (authR A);
+Class authG Σ (A : ucmraT) := AuthG {
+  auth_inG :> inG Σ (authR A);
   auth_timeless :> CMRADiscrete A;
 }.
 (* The Functor we need. *)
 Definition authGF (A : ucmraT) : gFunctor := GFunctor (constRF (authR A)).
 (* Show and register that they match. *)
-Instance authGF_inGF (A : ucmraT) `{inGF Λ Σ (authGF A)}
-  `{CMRADiscrete A} : authG Λ Σ A.
+Instance authGF_inGF `{inGF Σ (authGF A), CMRADiscrete A} : authG Σ A.
 Proof. split; try apply _. apply: inGF_inG. Qed.
 
 Section definitions.
-  Context `{authG Λ Σ A} (γ : gname).
-  Definition auth_own (a : A) : iPropG Λ Σ :=
+  Context `{irisG Λ Σ, authG Σ A} (γ : gname).
+  Definition auth_own (a : A) : iProp Σ :=
     own γ (◯ a).
-  Definition auth_inv (φ : A → iPropG Λ Σ) : iPropG Λ Σ :=
+  Definition auth_inv (φ : A → iProp Σ) : iProp Σ :=
     (∃ a, own γ (● a) ★ φ a)%I.
-  Definition auth_ctx (N : namespace) (φ : A → iPropG Λ Σ) : iPropG Λ Σ :=
+  Definition auth_ctx (N : namespace) (φ : A → iProp Σ) : iProp Σ :=
     inv N (auth_inv φ).
 
   Global Instance auth_own_ne n : Proper (dist n ==> dist n) auth_own.
@@ -41,15 +41,15 @@ Section definitions.
 End definitions.
 
 Typeclasses Opaque auth_own auth_ctx.
-Instance: Params (@auth_inv) 5.
-Instance: Params (@auth_own) 5.
-Instance: Params (@auth_ctx) 6.
+Instance: Params (@auth_inv) 6.
+Instance: Params (@auth_own) 6.
+Instance: Params (@auth_ctx) 7.
 
 Section auth.
-  Context `{AuthI : authG Λ Σ A}.
-  Context (φ : A → iPropG Λ Σ) {φ_proper : Proper ((≡) ==> (⊣⊢)) φ}.
+  Context `{irisG Λ Σ, authG Σ A}.
+  Context (φ : A → iProp Σ) {φ_proper : Proper ((≡) ==> (⊣⊢)) φ}.
   Implicit Types N : namespace.
-  Implicit Types P Q R : iPropG Λ Σ.
+  Implicit Types P Q R : iProp Σ.
   Implicit Types a b : A.
   Implicit Types γ : gname.
 
@@ -57,7 +57,7 @@ Section auth.
   Proof. by rewrite /auth_own -own_op auth_frag_op. Qed.
 
   Global Instance from_sep_own_authM γ a b :
-  FromSep (auth_own γ (a ⋅ b)) (auth_own γ a) (auth_own γ b) | 90.
+    FromSep (auth_own γ (a ⋅ b)) (auth_own γ a) (auth_own γ b) | 90.
   Proof. by rewrite /FromSep auth_own_op. Qed.
 
   Lemma auth_own_mono γ a b : a ≼ b → auth_own γ b ⊢ auth_own γ a.
@@ -71,51 +71,41 @@ Section auth.
   Proof. by rewrite /auth_own own_valid auth_validI. Qed.
 
   Lemma auth_alloc_strong N E a (G : gset gname) :
-    ✓ a → nclose N ⊆ E →
-    ▷ φ a ={E}=> ∃ γ, ■(γ ∉ G) ∧ auth_ctx γ N φ ∧ auth_own γ a.
+    ✓ a → ▷ φ a ={E}=> ∃ γ, ■ (γ ∉ G) ∧ auth_ctx γ N φ ∧ auth_own γ a.
   Proof.
-    iIntros (??) "Hφ". rewrite /auth_own /auth_ctx.
-    iPvs (own_alloc_strong (Auth (Excl' a) a) _ G) as (γ) "[% Hγ]"; first done.
+    iIntros (?) "Hφ". rewrite /auth_own /auth_ctx.
+    iVs (own_alloc_strong (Auth (Excl' a) a) G) as (γ) "[% Hγ]"; first done.
     iRevert "Hγ"; rewrite auth_both_op; iIntros "[Hγ Hγ']".
-    iPvs (inv_alloc N _ (auth_inv γ φ) with "[-Hγ']"); first done.
+    iVs (inv_alloc N _ (auth_inv γ φ) with "[-Hγ']").
     { iNext. iExists a. by iFrame. }
-    iPvsIntro; iExists γ. iSplit; first by iPureIntro. by iFrame.
+    iVsIntro; iExists γ. iSplit; first by iPureIntro. by iFrame.
   Qed.
 
   Lemma auth_alloc N E a :
-    ✓ a → nclose N ⊆ E →
-    ▷ φ a ={E}=> ∃ γ, auth_ctx γ N φ ∧ auth_own γ a.
+    ✓ a → ▷ φ a ={E}=> ∃ γ, auth_ctx γ N φ ∧ auth_own γ a.
   Proof.
-    iIntros (??) "Hφ".
-    iPvs (auth_alloc_strong N E a ∅ with "Hφ") as (γ) "[_ ?]"; [done..|].
-    by iExists γ.
+    iIntros (?) "Hφ".
+    iVs (auth_alloc_strong N E a ∅ with "Hφ") as (γ) "[_ ?]"; eauto.
   Qed.
 
-  Lemma auth_empty γ E : True ={E}=> auth_own γ ∅.
-  Proof. by rewrite -own_empty. Qed.
+  Lemma auth_empty γ : True =r=> auth_own γ ∅.
+  Proof. by rewrite /auth_own -own_empty. Qed.
 
-  Context {V} (fsa : FSA Λ (globalF Σ) V) `{!FrameShiftAssertion fsaV fsa}.
-
-  Lemma auth_fsa E N (Ψ : V → iPropG Λ Σ) γ a :
-    fsaV → nclose N ⊆ E →
-    auth_ctx γ N φ ★ ▷ auth_own γ a ★ (∀ af,
-      ■ ✓ (a ⋅ af) ★ ▷ φ (a ⋅ af) -★
-      fsa (E ∖ nclose N) (λ x, ∃ b,
-        ■ (a ~l~> b @ Some af) ★ ▷ φ (b ⋅ af) ★ (auth_own γ b -★ Ψ x)))
-     ⊢ fsa E Ψ.
+  Lemma auth_open E N γ a :
+    nclose N ⊆ E →
+    auth_ctx γ N φ ★ ▷ auth_own γ a ={E,E∖N}=> ∃ af,
+      ■ ✓ (a ⋅ af) ★ ▷ φ (a ⋅ af) ★ ∀ b,
+      ■ (a ~l~> b @ Some af) ★ ▷ φ (b ⋅ af) ={E∖N,E}=★ auth_own γ b.
   Proof.
-    iIntros (??) "(#? & Hγf & HΨ)". rewrite /auth_ctx /auth_own.
-    iInv N as (a') "[Hγ Hφ]".
+    iIntros (?) "(#? & Hγf)". rewrite /auth_ctx /auth_own.
+    iInv N as (a') "[Hγ Hφ]" "Hclose".
     iTimeless "Hγ"; iTimeless "Hγf"; iCombine "Hγ" "Hγf" as "Hγ".
-    iDestruct (own_valid with "#Hγ") as "Hvalid".
-    iDestruct (auth_validI _ with "Hvalid") as "[Ha' %]"; simpl; iClear "Hvalid".
-    iDestruct "Ha'" as (af) "Ha'"; iDestruct "Ha'" as %Ha'.
-    rewrite ->(left_id _ _) in Ha'; setoid_subst.
-    iApply pvs_fsa_fsa; iApply fsa_wand_r; iSplitL "HΨ Hφ".
-    { iApply "HΨ"; by iSplit. }
-    iIntros (v); iDestruct 1 as (b) "(% & Hφ & HΨ)".
-    iPvs (own_update with "Hγ") as "[Hγ Hγf]"; first eapply auth_update; eauto.
-    iPvsIntro. iSplitL "Hφ Hγ"; last by iApply "HΨ".
+    iDestruct (own_valid with "#Hγ") as % [[af Ha'] ?]%auth_valid_discrete.
+    simpl in Ha'; rewrite ->(left_id _ _) in Ha'; setoid_subst.
+    iVsIntro. iExists af; iFrame "Hφ"; iSplit; first done.
+    iIntros (b) "[% Hφ]".
+    iVs (own_update with "Hγ") as "[Hγ Hγf]"; first eapply auth_update; eauto.
+    iVs ("Hclose" with "[Hφ Hγ]") as "_"; auto.
     iNext. iExists (b â‹… af). by iFrame.
   Qed.
 End auth.
diff --git a/program_logic/boxes.v b/program_logic/boxes.v
index 793b69c51a7c49d081b691da5ac2ab3ee86c2a81..d926a527276d8a0012b13eb1b0f86157cde7eeed 100644
--- a/program_logic/boxes.v
+++ b/program_logic/boxes.v
@@ -1,49 +1,47 @@
-From iris.algebra Require Import upred_big_op.
+From iris.algebra Require Import gmap agree upred_big_op.
 From iris.program_logic Require Import auth saved_prop.
-From iris.proofmode Require Import tactics invariants ghost_ownership.
+From iris.proofmode Require Import tactics invariants.
 Import uPred.
 
 (** The CMRAs we need. *)
-Class boxG Λ Σ :=
-  boxG_inG :> inG Λ Σ (prodR
+Class boxG Σ :=
+  boxG_inG :> inG Σ (prodR
     (authR (optionUR (exclR boolC)))
-    (optionR (agreeR (laterC (iPrePropG Λ Σ))))).
+    (optionR (agreeR (laterC (iPreProp Σ))))).
 
 Section box_defs.
-  Context `{boxG Λ Σ} (N : namespace).
-  Notation iProp := (iPropG Λ Σ).
+  Context `{irisG Λ Σ, boxG Σ} (N : namespace).
 
   Definition slice_name := gname.
 
   Definition box_own_auth (γ : slice_name) (a : auth (option (excl bool)))
-    := own γ (a, (∅:option (agree (later (iPrePropG Λ Σ))))).
+    := own γ (a, (∅:option (agree (later (iPreProp Σ))))).
 
-  Definition box_own_prop (γ : slice_name) (P : iProp) : iProp :=
+  Definition box_own_prop (γ : slice_name) (P : iProp Σ) : iProp Σ :=
     own γ (∅:auth (option (excl bool)), Some (to_agree (Next (iProp_unfold P)))).
 
-  Definition slice_inv (γ : slice_name) (P : iProp) : iProp :=
+  Definition slice_inv (γ : slice_name) (P : iProp Σ) : iProp Σ :=
     (∃ b, box_own_auth γ (● Excl' b) ★ box_own_prop γ P ★ if b then P else True)%I.
 
-  Definition slice (γ : slice_name) (P : iProp) : iProp :=
+  Definition slice (γ : slice_name) (P : iProp Σ) : iProp Σ :=
     inv N (slice_inv γ P).
 
-  Definition box (f : gmap slice_name bool) (P : iProp) : iProp :=
-    (∃ Φ : slice_name → iProp,
+  Definition box (f : gmap slice_name bool) (P : iProp Σ) : iProp Σ :=
+    (∃ Φ : slice_name → iProp Σ,
       ▷ (P ≡ [★ map] γ ↦ b ∈ f, Φ γ) ★
       [★ map] γ ↦ b ∈ f, box_own_auth γ (◯ Excl' b) ★ box_own_prop γ (Φ γ) ★
                          inv N (slice_inv γ (Φ γ)))%I.
 End box_defs.
 
-Instance: Params (@box_own_auth) 4.
-Instance: Params (@box_own_prop) 4.
-Instance: Params (@slice_inv) 4.
-Instance: Params (@slice) 5.
-Instance: Params (@box) 5.
+Instance: Params (@box_own_auth) 5.
+Instance: Params (@box_own_prop) 5.
+Instance: Params (@slice_inv) 5.
+Instance: Params (@slice) 6.
+Instance: Params (@box) 6.
 
 Section box.
-Context `{boxG Λ Σ} (N : namespace).
-Notation iProp := (iPropG Λ Σ).
-Implicit Types P Q : iProp.
+Context `{irisG Λ Σ, boxG Σ} (N : namespace).
+Implicit Types P Q : iProp Σ.
 
 Global Instance box_own_prop_ne n γ : Proper (dist n ==> dist n) (box_own_prop γ).
 Proof. solve_proper. Qed.
@@ -63,9 +61,9 @@ Proof.
   by iDestruct 1 as % [[[] [=]%leibniz_equiv] ?]%auth_valid_discrete.
 Qed.
 
-Lemma box_own_auth_update E γ b1 b2 b3 :
+Lemma box_own_auth_update γ b1 b2 b3 :
   box_own_auth γ (● Excl' b1) ★ box_own_auth γ (◯ Excl' b2)
-  ={E}=> box_own_auth γ (● Excl' b3) ★ box_own_auth γ (◯ Excl' b3).
+  =r=> box_own_auth γ (● Excl' b3) ★ box_own_auth γ (◯ Excl' b3).
 Proof.
   rewrite /box_own_prop -!own_op own_valid_l prod_validI; iIntros "[[Hb _] Hγ]".
   iDestruct "Hb" as % [[[] [= ->]%leibniz_equiv] ?]%auth_valid_discrete.
@@ -94,14 +92,14 @@ Lemma box_insert f P Q :
     slice N γ Q ★ ▷ box N (<[γ:=false]> f) (Q ★ P).
 Proof.
   iDestruct 1 as (Φ) "[#HeqP Hf]".
-  iPvs (own_alloc_strong (● Excl' false ⋅ ◯ Excl' false,
-    Some (to_agree (Next (iProp_unfold Q)))) _ (dom _ f))
+  iVs (own_alloc_strong (● Excl' false ⋅ ◯ Excl' false,
+    Some (to_agree (Next (iProp_unfold Q)))) (dom _ f))
     as (γ) "[Hdom Hγ]"; first done.
   rewrite pair_split. iDestruct "Hγ" as "[[Hγ Hγ'] #HγQ]".
   iDestruct "Hdom" as % ?%not_elem_of_dom.
-  iPvs (inv_alloc N _ (slice_inv γ Q) with "[Hγ]") as "#Hinv"; first done.
+  iVs (inv_alloc N _ (slice_inv γ Q) with "[Hγ]") as "#Hinv".
   { iNext. iExists false; eauto. }
-  iPvsIntro; iExists γ; repeat iSplit; auto.
+  iVsIntro; iExists γ; repeat iSplit; auto.
   iNext. iExists (<[γ:=Q]> Φ); iSplit.
   - iNext. iRewrite "HeqP". by rewrite big_sepM_fn_insert'.
   - rewrite (big_sepM_fn_insert (λ _ _ P',  _ ★ _ _ P' ★ _ _ (_ _ P')))%I //.
@@ -115,7 +113,8 @@ Lemma box_delete f P Q γ :
 Proof.
   iIntros (?) "[#Hinv H]"; iDestruct "H" as (Φ) "[#HeqP Hf]".
   iExists ([★ map] γ'↦_ ∈ delete γ f, Φ γ')%I.
-  iInv N as (b) "(Hγ & #HγQ &_)"; iPvsIntro; iNext.
+  iInv N as (b) "(Hγ & #HγQ &_)" "Hclose".
+  iApply pvs_trans_frame; iFrame "Hclose"; iVsIntro; iNext.
   iDestruct (big_sepM_delete _ f _ false with "Hf")
     as "[[Hγ' #[HγΦ ?]] ?]"; first done.
   iDestruct (box_own_agree γ Q (Φ γ) with "[#]") as "HeqQ"; first by eauto.
@@ -132,14 +131,14 @@ Lemma box_fill f γ P Q :
   slice N γ Q ★ ▷ Q ★ ▷ box N f P ={N}=> ▷ box N (<[γ:=true]> f) P.
 Proof.
   iIntros (?) "(#Hinv & HQ & H)"; iDestruct "H" as (Φ) "[#HeqP Hf]".
-  iInv N as (b') "(Hγ & #HγQ & _)"; iTimeless "Hγ".
+  iInv N as (b') "(Hγ & #HγQ & _)" "Hclose". iTimeless "Hγ".
   iDestruct (big_sepM_later _ f with "Hf") as "Hf".
   iDestruct (big_sepM_delete _ f _ false with "Hf")
     as "[[Hγ' #[HγΦ Hinv']] ?]"; first done; iTimeless "Hγ'".
-  iPvs (box_own_auth_update _ γ b' false true with "[Hγ Hγ']")
+  iVs (box_own_auth_update γ b' false true with "[Hγ Hγ']")
     as "[Hγ Hγ']"; first by iFrame.
-  iPvsIntro; iNext; iSplitL "Hγ HQ"; first (iExists true; by iFrame "Hγ HQ").
-  iExists Φ; iSplit.
+  iVs ("Hclose" with "[Hγ HQ]"); first (iNext; iExists true; by iFrame).
+  iVsIntro; iNext; iExists Φ; iSplit.
   - by rewrite big_sepM_insert_override.
   - rewrite -insert_delete big_sepM_insert ?lookup_delete //.
     iFrame; eauto.
@@ -150,17 +149,16 @@ Lemma box_empty f P Q γ :
   slice N γ Q ★ ▷ box N f P ={N}=> ▷ Q ★ ▷ box N (<[γ:=false]> f) P.
 Proof.
   iIntros (?) "[#Hinv H]"; iDestruct "H" as (Φ) "[#HeqP Hf]".
-  iInv N as (b) "(Hγ & #HγQ & HQ)"; iTimeless "Hγ".
+  iInv N as (b) "(Hγ & #HγQ & HQ)" "Hclose"; iTimeless "Hγ".
   iDestruct (big_sepM_later _ f with "Hf") as "Hf".
   iDestruct (big_sepM_delete _ f with "Hf")
     as "[[Hγ' #[HγΦ Hinv']] ?]"; first done; iTimeless "Hγ'".
   iDestruct (box_own_auth_agree γ b true with "[#]")
     as "%"; subst; first by iFrame.
   iFrame "HQ".
-  iPvs (box_own_auth_update _ γ with "[Hγ Hγ']")
-    as "[Hγ Hγ']"; first by iFrame.
-  iPvsIntro; iNext; iSplitL "Hγ"; first (iExists false; by repeat iSplit).
-  iExists Φ; iSplit.
+  iVs (box_own_auth_update γ with "[Hγ Hγ']") as "[Hγ Hγ']"; first by iFrame.
+  iVs ("Hclose" with "[Hγ]"); first (iNext; iExists false; by repeat iSplit).
+  iVsIntro; iNext; iExists Φ; iSplit.
   - by rewrite big_sepM_insert_override.
   - rewrite -insert_delete big_sepM_insert ?lookup_delete //.
     iFrame; eauto.
@@ -175,10 +173,9 @@ Proof.
   rewrite big_sepM_fmap; iApply (pvs_big_sepM _ _ f).
   iApply (big_sepM_impl _ _ f); iFrame "Hf".
   iAlways; iIntros (γ b' ?) "[(Hγ' & #$ & #$) HΦ]".
-  iInv N as (b) "[Hγ _]"; iTimeless "Hγ".
-  iPvs (box_own_auth_update _ γ with "[Hγ Hγ']")
-    as "[Hγ $]"; first by iFrame.
-  iPvsIntro; iNext; iExists true. by iFrame.
+  iInv N as (b) "[Hγ _]" "Hclose"; iTimeless "Hγ".
+  iVs (box_own_auth_update γ with "[Hγ Hγ']") as "[Hγ $]"; first by iFrame.
+  iApply "Hclose". iNext; iExists true. by iFrame.
 Qed.
 
 Lemma box_empty_all f P Q :
@@ -191,13 +188,14 @@ Proof.
   { iApply (pvs_big_sepM _ _ f); iApply (big_sepM_impl _ _ f); iFrame "Hf".
     iAlways; iIntros (γ b ?) "(Hγ' & #$ & #$)".
     assert (true = b) as <- by eauto.
-    iInv N as (b) "(Hγ & _ & HΦ)"; iTimeless "Hγ".
+    iInv N as (b) "(Hγ & _ & HΦ)" "Hclose"; iTimeless "Hγ".
     iDestruct (box_own_auth_agree γ b true with "[#]")
       as "%"; subst; first by iFrame.
-    iPvs (box_own_auth_update _ γ true true false with "[Hγ Hγ']")
+    iVs (box_own_auth_update γ true true false with "[Hγ Hγ']")
       as "[Hγ $]"; first by iFrame.
-    iPvsIntro; iNext. iFrame "HΦ". iExists false. iFrame; eauto. }
-  iPvsIntro; iSplitL "HΦ".
+    iVs ("Hclose" with "[Hγ]"); first (iNext; iExists false; iFrame; eauto).
+    by iApply "HΦ". }
+  iVsIntro; iSplitL "HΦ".
   - rewrite eq_iff later_iff big_sepM_later. by iApply "HeqP".
   - iExists Φ; iSplit; by rewrite big_sepM_fmap.
 Qed.
diff --git a/program_logic/ectx_lifting.v b/program_logic/ectx_lifting.v
index 48a49739df9dbe8449b2f3e08a4ecb6a0b7bf7c3..d298f34b219caf202f24e7bebbbc444f9b5a9df1 100644
--- a/program_logic/ectx_lifting.v
+++ b/program_logic/ectx_lifting.v
@@ -5,28 +5,26 @@ From iris.proofmode Require Import weakestpre.
 
 Section wp.
 Context {expr val ectx state} {Λ : EctxLanguage expr val ectx state}.
-Context {Σ : iFunctor}.
-Implicit Types P : iProp (ectx_lang expr) Σ.
-Implicit Types Φ : val → iProp (ectx_lang expr) Σ.
+Context `{irisG (ectx_lang expr) Σ}.
+Implicit Types P : iProp Σ.
+Implicit Types Φ : val → iProp Σ.
 Implicit Types v : val.
 Implicit Types e : expr.
 Hint Resolve head_prim_reducible head_reducible_prim_step.
 
-Notation wp_fork ef := (default True ef (flip (wp ⊤) (λ _, True)))%I.
-
 Lemma wp_ectx_bind {E e} K Φ :
   WP e @ E {{ v, WP fill K (of_val v) @ E {{ Φ }} }} ⊢ WP fill K e @ E {{ Φ }}.
 Proof. apply: weakestpre.wp_bind. Qed.
 
-Lemma wp_lift_head_step E1 E2 Φ e1 :
-  E2 ⊆ E1 → to_val e1 = None →
-  (|={E1,E2}=> ∃ σ1, ■ head_reducible e1 σ1 ∧
-       ▷ ownP σ1 ★ ▷ ∀ e2 σ2 ef, (■ head_step e1 σ1 e2 σ2 ef ∧ ownP σ2)
-                                 ={E2,E1}=★ WP e2 @ E1 {{ Φ }} ★ wp_fork ef)
-  ⊢ WP e1 @ E1 {{ Φ }}.
+Lemma wp_lift_head_step E Φ e1 :
+  to_val e1 = None →
+  (|={E,∅}=> ∃ σ1, ■ head_reducible e1 σ1 ★ ▷ ownP σ1 ★
+     ▷ ∀ e2 σ2 ef, ■ head_step e1 σ1 e2 σ2 ef ★ ownP σ2
+          ={∅,E}=★ WP e2 @ E {{ Φ }} ★ wp_fork ef)
+  ⊢ WP e1 @ E {{ Φ }}.
 Proof.
-  iIntros (??) "H". iApply (wp_lift_step E1 E2); try done.
-  iPvs "H" as (σ1) "(%&Hσ1&Hwp)". set_solver. iPvsIntro. iExists σ1.
+  iIntros (?) "H". iApply (wp_lift_step E); try done.
+  iVs "H" as (σ1) "(%&Hσ1&Hwp)". iVsIntro. iExists σ1.
   iSplit; first by eauto. iFrame. iNext. iIntros (e2 σ2 ef) "[% ?]".
   iApply "Hwp". by eauto.
 Qed.
diff --git a/program_logic/ghost_ownership.v b/program_logic/ghost_ownership.v
index 92c4adc27d02c2e6a0f15e21c96fc69cf59ee700..3706f39fc3424def05dc1baecb13910e5ae178e4 100644
--- a/program_logic/ghost_ownership.v
+++ b/program_logic/ghost_ownership.v
@@ -1,39 +1,37 @@
-From iris.prelude Require Export functions.
-From iris.algebra Require Export iprod.
-From iris.program_logic Require Export pviewshifts global_functor.
-From iris.program_logic Require Import ownership.
+From iris.program_logic Require Export global_functor.
+From iris.algebra Require Import iprod gmap.
 Import uPred.
 
-Definition own_def `{inG Λ Σ A} (γ : gname) (a : A) : iPropG Λ Σ :=
-  ownG (to_globalF γ a).
+Definition own_def `{inG Σ A} (γ : gname) (a : A) : iProp Σ :=
+  uPred_ownM (to_iRes γ a).
 Definition own_aux : { x | x = @own_def }. by eexists. Qed.
-Definition own {Λ Σ A i} := proj1_sig own_aux Λ Σ A i.
+Definition own {Σ A i} := proj1_sig own_aux Σ A i.
 Definition own_eq : @own = @own_def := proj2_sig own_aux.
-Instance: Params (@own) 5.
+Instance: Params (@own) 4.
 Typeclasses Opaque own.
 
 (** Properties about ghost ownership *)
 Section global.
-Context `{i : inG Λ Σ A}.
+Context `{i : inG Σ A}.
 Implicit Types a : A.
 
 (** * Properties of own *)
-Global Instance own_ne γ n : Proper (dist n ==> dist n) (@own Λ Σ A _ γ).
+Global Instance own_ne γ n : Proper (dist n ==> dist n) (@own Σ A _ γ).
 Proof. rewrite !own_eq. solve_proper. Qed.
 Global Instance own_proper γ :
-  Proper ((≡) ==> (⊣⊢)) (@own Λ Σ A _ γ) := ne_proper _.
+  Proper ((≡) ==> (⊣⊢)) (@own Σ A _ γ) := ne_proper _.
 
 Lemma own_op γ a1 a2 : own γ (a1 ⋅ a2) ⊣⊢ own γ a1 ★ own γ a2.
-Proof. by rewrite !own_eq /own_def -ownG_op to_globalF_op. Qed.
-Global Instance own_mono γ : Proper (flip (≼) ==> (⊢)) (@own Λ Σ A _ γ).
+Proof. by rewrite !own_eq /own_def -ownM_op to_iRes_op. Qed.
+Global Instance own_mono γ : Proper (flip (≼) ==> (⊢)) (@own Σ A _ γ).
 Proof. move=>a b [c ->]. rewrite own_op. eauto with I. Qed.
 Lemma own_valid γ a : own γ a ⊢ ✓ a.
 Proof.
-  rewrite !own_eq /own_def ownG_valid /to_globalF.
+  rewrite !own_eq /own_def ownM_valid /to_iRes.
   rewrite iprod_validI (forall_elim (inG_id i)) iprod_lookup_singleton.
   rewrite gmap_validI (forall_elim γ) lookup_singleton option_validI.
   (* implicit arguments differ a bit *)
-  by trans (✓ cmra_transport inG_prf a : iPropG Λ Σ)%I; last destruct inG_prf.
+  by trans (✓ cmra_transport inG_prf a : iProp Σ)%I; last destruct inG_prf.
 Qed.
 Lemma own_valid_r γ a : own γ a ⊢ own γ a ★ ✓ a.
 Proof. apply: uPred.always_entails_r. apply own_valid. Qed.
@@ -46,58 +44,53 @@ Proof. rewrite !own_eq /own_def; apply _. Qed.
 
 (* TODO: This also holds if we just have ✓ a at the current step-idx, as Iris
    assertion. However, the map_updateP_alloc does not suffice to show this. *)
-Lemma own_alloc_strong a E (G : gset gname) :
-  ✓ a → True ={E}=> ∃ γ, ■(γ ∉ G) ∧ own γ a.
+Lemma own_alloc_strong a (G : gset gname) :
+  ✓ a → True =r=> ∃ γ, ■ (γ ∉ G) ∧ own γ a.
 Proof.
   intros Ha.
-  rewrite -(pvs_mono _ _ (∃ m, ■ (∃ γ, γ ∉ G ∧ m = to_globalF γ a) ∧ ownG m)%I).
-  - rewrite ownG_empty.
-    eapply pvs_ownG_updateP, (iprod_singleton_updateP_empty (inG_id i));
+  rewrite -(rvs_mono (∃ m, ■ (∃ γ, γ ∉ G ∧ m = to_iRes γ a) ∧ uPred_ownM m)%I).
+  - rewrite ownM_empty.
+    eapply rvs_ownM_updateP, (iprod_singleton_updateP_empty (inG_id i));
       first (eapply alloc_updateP_strong', cmra_transport_valid, Ha);
       naive_solver.
   - apply exist_elim=>m; apply pure_elim_l=>-[γ [Hfresh ->]].
     by rewrite !own_eq /own_def -(exist_intro γ) pure_equiv // left_id.
 Qed.
-Lemma own_alloc a E : ✓ a → True ={E}=> ∃ γ, own γ a.
+Lemma own_alloc a : ✓ a → True =r=> ∃ γ, own γ a.
 Proof.
-  intros Ha. rewrite (own_alloc_strong a E ∅) //; [].
-  apply pvs_mono, exist_mono=>?. eauto with I.
+  intros Ha. rewrite (own_alloc_strong a ∅) //; [].
+  apply rvs_mono, exist_mono=>?. eauto with I.
 Qed.
 
-Lemma own_updateP P γ a E : a ~~>: P → own γ a ={E}=> ∃ a', ■ P a' ∧ own γ a'.
+Lemma own_updateP P γ a : a ~~>: P → own γ a =r=> ∃ a', ■ P a' ∧ own γ a'.
 Proof.
   intros Ha. rewrite !own_eq.
-  rewrite -(pvs_mono _ _ (∃ m, ■ (∃ a', m = to_globalF γ a' ∧ P a') ∧ ownG m)%I).
-  - eapply pvs_ownG_updateP, iprod_singleton_updateP;
+  rewrite -(rvs_mono (∃ m, ■ (∃ a', m = to_iRes γ a' ∧ P a') ∧ uPred_ownM m)%I).
+  - eapply rvs_ownM_updateP, iprod_singleton_updateP;
       first by (eapply singleton_updateP', cmra_transport_updateP', Ha).
     naive_solver.
   - apply exist_elim=>m; apply pure_elim_l=>-[a' [-> HP]].
     rewrite -(exist_intro a'). by apply and_intro; [apply pure_intro|].
 Qed.
 
-Lemma own_update γ a a' E : a ~~> a' → own γ a ={E}=> own γ a'.
+Lemma own_update γ a a' : a ~~> a' → own γ a =r=> own γ a'.
 Proof.
   intros; rewrite (own_updateP (a' =)); last by apply cmra_update_updateP.
-  by apply pvs_mono, exist_elim=> a''; apply pure_elim_l=> ->.
+  by apply rvs_mono, exist_elim=> a''; apply pure_elim_l=> ->.
 Qed.
 End global.
 
-Arguments own_valid {_ _ _} [_] _ _.
-Arguments own_valid_l {_ _ _} [_] _ _.
-Arguments own_valid_r {_ _ _} [_] _ _.
-Arguments own_updateP {_ _ _} [_] _ _ _ _ _.
-Arguments own_update {_ _ _} [_] _ _ _ _ _.
+Arguments own_valid {_ _} [_] _ _.
+Arguments own_valid_l {_ _} [_] _ _.
+Arguments own_valid_r {_ _} [_] _ _.
+Arguments own_updateP {_ _} [_] _ _ _ _.
+Arguments own_update {_ _} [_] _ _ _ _.
 
-Section global_empty.
-Context `{i : inG Λ Σ (A:ucmraT)}.
-Implicit Types a : A.
-
-Lemma own_empty γ E : True ={E}=> own γ (∅:A).
+Lemma own_empty `{inG Σ (A:ucmraT)} γ : True =r=> own γ ∅.
 Proof.
-  rewrite ownG_empty !own_eq /own_def.
-  apply pvs_ownG_update, iprod_singleton_update_empty.
+  rewrite ownM_empty !own_eq /own_def.
+  apply rvs_ownM_update, iprod_singleton_update_empty.
   apply (alloc_unit_singleton_update (cmra_transport inG_prf ∅)); last done.
   - apply cmra_transport_valid, ucmra_unit_valid.
   - intros x; destruct inG_prf. by rewrite left_id.
 Qed.
-End global_empty.
diff --git a/program_logic/global_functor.v b/program_logic/global_functor.v
index 61b832c85008da818b3b3f5722f265cf2c261bc4..1689e988bd136b3fecaf05c1ef2f9d78d2839984 100644
--- a/program_logic/global_functor.v
+++ b/program_logic/global_functor.v
@@ -1,58 +1,36 @@
-From iris.algebra Require Export iprod.
 From iris.program_logic Require Export model.
+From iris.algebra Require Import iprod gmap.
 
-(** The "default" iFunctor is constructed as the dependent product of
-    a bunch of gFunctor. *)
-Structure gFunctor := GFunctor {
-  gFunctor_F :> rFunctor;
-  gFunctor_contractive : rFunctorContractive gFunctor_F;
-}.
-Arguments GFunctor _ {_}.
-Existing Instance gFunctor_contractive.
-
-(** The global CMRA: Indexed product over a gid i to (gname --fin--> Σ i) *)
-Definition gFunctors := { n : nat & fin n → gFunctor }.
-Definition gid (Σ : gFunctors) := fin (projT1 Σ).
-
-(** Name of one instance of a particular CMRA in the ghost state. *)
-Definition gname := positive.
-Canonical Structure gnameC := leibnizC gname.
-
-Definition globalF (Σ : gFunctors) : iFunctor :=
-  IFunctor (iprodURF (λ i, gmapURF gname (projT2 Σ i))).
-Notation iPropG Λ Σ := (iProp Λ (globalF Σ)).
-Notation iPrePropG Λ Σ := (iPreProp Λ (globalF Σ)).
-
-Class inG (Λ : language) (Σ : gFunctors) (A : cmraT) := InG {
+Class inG (Σ : gFunctors) (A : cmraT) := InG {
   inG_id : gid Σ;
-  inG_prf : A = projT2 Σ inG_id (iPreProp Λ (globalF Σ))
+  inG_prf : A = projT2 Σ inG_id (iPreProp Σ)
 }.
-Arguments inG_id {_ _ _} _.
+Arguments inG_id {_ _} _.
 
-Definition to_globalF `{i : inG Λ Σ A} (γ : gname) (a : A) : iGst Λ (globalF Σ) :=
+Definition to_iRes `{i : inG Σ A} (γ : gname) (a : A) : iResUR Σ :=
   iprod_singleton (inG_id i) {[ γ := cmra_transport inG_prf a ]}.
-Instance: Params (@to_globalF) 5.
-Typeclasses Opaque to_globalF.
+Instance: Params (@to_iRes) 4.
+Typeclasses Opaque to_iRes.
 
-(** * Properties of to_globalC *)
-Section to_globalF.
-Context `{i : inG Λ Σ A}.
+(** * Properties of [to_iRes] *)
+Section to_iRes.
+Context `{inG Σ A}.
 Implicit Types a : A.
 
-Global Instance to_globalF_ne γ n :
-  Proper (dist n ==> dist n) (@to_globalF Λ Σ A _ γ).
+Global Instance to_iRes_ne γ n :
+  Proper (dist n ==> dist n) (@to_iRes Σ A _ γ).
 Proof. by intros a a' Ha; apply iprod_singleton_ne; rewrite Ha. Qed.
-Lemma to_globalF_op γ a1 a2 :
-  to_globalF γ (a1 ⋅ a2) ≡ to_globalF γ a1 ⋅ to_globalF γ a2.
+Lemma to_iRes_op γ a1 a2 :
+  to_iRes γ (a1 ⋅ a2) ≡ to_iRes γ a1 ⋅ to_iRes γ a2.
 Proof.
-  by rewrite /to_globalF iprod_op_singleton op_singleton cmra_transport_op.
+  by rewrite /to_iRes iprod_op_singleton op_singleton cmra_transport_op.
 Qed.
-Global Instance to_globalF_timeless γ a : Timeless a → Timeless (to_globalF γ a).
-Proof. rewrite /to_globalF; apply _. Qed.
-Global Instance to_globalF_persistent γ a :
-  Persistent a → Persistent (to_globalF γ a).
-Proof. rewrite /to_globalF; apply _. Qed.
-End to_globalF.
+Global Instance to_iRes_timeless γ a : Timeless a → Timeless (to_iRes γ a).
+Proof. rewrite /to_iRes; apply _. Qed.
+Global Instance to_iRes_persistent γ a :
+  Persistent a → Persistent (to_iRes γ a).
+Proof. rewrite /to_iRes; apply _. Qed.
+End to_iRes.
 
 (** When instantiating the logic, we want to just plug together the
     requirements exported by the modules we use. To this end, we construct
@@ -106,7 +84,7 @@ Notation "#[ Fs1 ; Fs2 ; .. ; Fsn ]" :=
 (** We need another typeclass to identify the *functor* in the Σ. Basing inG on
    the functor breaks badly because Coq is unable to infer the correct
    typeclasses, it does not unfold the functor. *)
-Class inGF (Λ : language) (Σ : gFunctors) (F : gFunctor) := InGF {
+Class inGF (Σ : gFunctors) (F : gFunctor) := InGF {
   inGF_id : gid Σ;
   inGF_prf : F = projT2 Σ inGF_id;
 }.
@@ -115,25 +93,25 @@ is only triggered if the first two arguments of inGF do not contain evars. Since
 instance search for [inGF] is restrained, instances should always have [inGF] as
 their first argument to avoid loops. For example, the instances [authGF_inGF]
 and [auth_identity] otherwise create a cycle that pops up arbitrarily. *)
-Hint Mode inGF + + - : typeclass_instances.
+Hint Mode inGF + - : typeclass_instances.
 
-Lemma inGF_inG `{inGF Λ Σ F} : inG Λ Σ (F (iPrePropG Λ Σ)).
-Proof. exists inGF_id. by rewrite -inGF_prf. Qed.
-Instance inGF_here {Λ Σ} (F: gFunctor) : inGF Λ (gFunctors.cons F Σ) F.
+Lemma inGF_inG {Σ F} : inGF Σ F → inG Σ (F (iPreProp Σ)).
+Proof. intros. exists inGF_id. by rewrite -inGF_prf. Qed.
+Instance inGF_here {Σ} (F: gFunctor) : inGF (gFunctors.cons F Σ) F.
 Proof. by exists 0%fin. Qed.
-Instance inGF_further {Λ Σ} (F F': gFunctor) :
-  inGF Λ Σ F → inGF Λ (gFunctors.cons F' Σ) F.
+Instance inGF_further {Σ} (F F': gFunctor) :
+  inGF Σ F → inGF (gFunctors.cons F' Σ) F.
 Proof. intros [i ?]. by exists (FS i). Qed.
 
 (** For modules that need more than one functor, we offer a typeclass
     [inGFs] to demand a list of rFunctor to be available. We do
     *not* register any instances that go from there to [inGF], to
     avoid cycles. *)
-Class inGFs (Λ : language) (Σ : gFunctors) (Fs : gFunctorList) :=
-  InGFs : (gFunctorList.fold_right (λ F T, inGF Λ Σ F * T) () Fs)%type.
+Class inGFs (Σ : gFunctors) (Fs : gFunctorList) :=
+  InGFs : (gFunctorList.fold_right (λ F T, inGF Σ F * T) () Fs)%type.
 
-Instance inGFs_nil {Λ Σ} : inGFs Λ Σ [].
+Instance inGFs_nil {Σ} : inGFs Σ [].
 Proof. exact tt. Qed.
-Instance inGFs_cons {Λ Σ} F Fs :
-  inGF Λ Σ F → inGFs Λ Σ Fs → inGFs Λ Σ (gFunctorList.cons F Fs).
+Instance inGFs_cons {Σ} F Fs :
+  inGF Σ F → inGFs Σ Fs → inGFs Σ (gFunctorList.cons F Fs).
 Proof. by split. Qed.
diff --git a/program_logic/hoare.v b/program_logic/hoare.v
index 9423942754099e98506a463c9602b6ae5e97bf6f..39629cdba1008db3110a8790fb16bc9515e7f188 100644
--- a/program_logic/hoare.v
+++ b/program_logic/hoare.v
@@ -1,10 +1,10 @@
-From iris.program_logic Require Export weakestpre viewshifts.
-From iris.proofmode Require Import weakestpre invariants.
+From iris.program_logic Require Export weakestpre. (* viewshifts *)
+From iris.proofmode Require Import weakestpre.
 
-Definition ht {Λ Σ} (E : coPset) (P : iProp Λ Σ)
-    (e : expr Λ) (Φ : val Λ → iProp Λ Σ) : iProp Λ Σ :=
+Definition ht `{irisG Λ Σ} (E : coPset) (P : iProp Σ)
+    (e : expr Λ) (Φ : val Λ → iProp Σ) : iProp Σ :=
   (□ (P → WP e @ E {{ Φ }}))%I.
-Instance: Params (@ht) 3.
+Instance: Params (@ht) 4.
 
 Notation "{{ P } } e @ E {{ Φ } }" := (ht E P e Φ)
   (at level 20, P, e, Φ at level 200,
@@ -33,47 +33,47 @@ Notation "{{ P } } e {{ v , Q } }" := (True ⊢ ht ⊤ P e (λ v, Q))
    format "{{  P  } }  e  {{  v ,  Q  } }") : C_scope.
 
 Section hoare.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P Q : iProp Λ Σ.
-Implicit Types Φ Ψ : val Λ → iProp Λ Σ.
+Context `{irisG Λ Σ}.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ Ψ : val Λ → iProp Σ.
 Implicit Types v : val Λ.
 Import uPred.
 
 Global Instance ht_ne E n :
-  Proper (dist n ==> eq==>pointwise_relation _ (dist n) ==> dist n) (@ht Λ Σ E).
+  Proper (dist n ==> eq==>pointwise_relation _ (dist n) ==> dist n) (ht E).
 Proof. solve_proper. Qed.
 Global Instance ht_proper E :
-  Proper ((≡) ==> eq ==> pointwise_relation _ (≡) ==> (≡)) (@ht Λ Σ E).
+  Proper ((≡) ==> eq ==> pointwise_relation _ (≡) ==> (≡)) (ht E).
 Proof. solve_proper. Qed.
 Lemma ht_mono E P P' Φ Φ' e :
   (P ⊢ P') → (∀ v, Φ' v ⊢ Φ v) → {{ P' }} e @ E {{ Φ' }} ⊢ {{ P }} e @ E {{ Φ }}.
 Proof. by intros; apply always_mono, impl_mono, wp_mono. Qed.
 Global Instance ht_mono' E :
-  Proper (flip (⊢) ==> eq ==> pointwise_relation _ (⊢) ==> (⊢)) (@ht Λ Σ E).
+  Proper (flip (⊢) ==> eq ==> pointwise_relation _ (⊢) ==> (⊢)) (ht E).
 Proof. solve_proper. Qed.
 
 Lemma ht_alt E P Φ e : (P ⊢ WP e @ E {{ Φ }}) → {{ P }} e @ E {{ Φ }}.
 Proof. iIntros (Hwp) "! HP". by iApply Hwp. Qed.
 
-Lemma ht_val E v : {{ True : iProp Λ Σ }} of_val v @ E {{ v', v = v' }}.
+Lemma ht_val E v : {{ True }} of_val v @ E {{ v', v = v' }}.
 Proof. iIntros "! _". by iApply wp_value'. Qed.
 
 Lemma ht_vs E P P' Φ Φ' e :
-  (P ={E}=> P') ∧ {{ P' }} e @ E {{ Φ' }} ∧ (∀ v, Φ' v ={E}=> Φ v)
+  □ (P ={E}=★ P') ∧ {{ P' }} e @ E {{ Φ' }} ∧ □ (∀ v, Φ' v ={E}=★ Φ v)
   ⊢ {{ P }} e @ E {{ Φ }}.
 Proof.
-  iIntros "(#Hvs&#Hwp&#HΦ) ! HP". iPvs ("Hvs" with "HP") as "HP".
+  iIntros "(#Hvs&#Hwp&#HΦ) ! HP". iVs ("Hvs" with "HP") as "HP".
   iApply wp_pvs; iApply wp_wand_r; iSplitL; [by iApply "Hwp"|].
   iIntros (v) "Hv". by iApply "HΦ".
 Qed.
 
 Lemma ht_atomic E1 E2 P P' Φ Φ' e :
-  E2 ⊆ E1 → atomic e →
-  (P ={E1,E2}=> P') ∧ {{ P' }} e @ E2 {{ Φ' }} ∧ (∀ v, Φ' v ={E2,E1}=> Φ v)
+  atomic e →
+  □ (P ={E1,E2}=★ P') ∧ {{ P' }} e @ E2 {{ Φ' }} ∧ □ (∀ v, Φ' v ={E2,E1}=★ Φ v)
   ⊢ {{ P }} e @ E1 {{ Φ }}.
 Proof.
-  iIntros (??) "(#Hvs&#Hwp&#HΦ) ! HP". iApply (wp_atomic _ E2); auto.
-  iPvs ("Hvs" with "HP") as "HP"; first set_solver. iPvsIntro.
+  iIntros (?) "(#Hvs&#Hwp&#HΦ) ! HP". iApply (wp_atomic _ E2); auto.
+  iVs ("Hvs" with "HP") as "HP". iVsIntro.
   iApply wp_wand_r; iSplitL; [by iApply "Hwp"|].
   iIntros (v) "Hv". by iApply "HΦ".
 Qed.
@@ -90,7 +90,7 @@ Qed.
 Lemma ht_mask_weaken E1 E2 P Φ e :
   E1 ⊆ E2 → {{ P }} e @ E1 {{ Φ }} ⊢ {{ P }} e @ E2 {{ Φ }}.
 Proof.
-  iIntros (?) "#Hwp ! HP". iApply (wp_mask_frame_mono E1 E2); try done.
+  iIntros (?) "#Hwp ! HP". iApply (wp_mask_mono E1 E2); try done.
   by iApply "Hwp".
 Qed.
 
@@ -102,26 +102,24 @@ Lemma ht_frame_r E P Φ R e :
   {{ P }} e @ E {{ Φ }} ⊢ {{ P ★ R }} e @ E {{ v, Φ v ★ R }}.
 Proof. iIntros "#Hwp ! [HP $]". by iApply "Hwp". Qed.
 
-Lemma ht_frame_step_l E E1 E2 P R1 R2 R3 e Φ :
-  to_val e = None → E ⊥ E1 → E2 ⊆ E1 →
-  (R1 ={E1,E2}=> ▷ R2) ∧ (R2 ={E2,E1}=> R3) ∧ {{ P }} e @ E {{ Φ }}
-  ⊢ {{ R1 ★ P }} e @ E ∪ E1 {{ λ v, R3 ★ Φ v }}.
+Lemma ht_frame_step_l E1 E2 P R1 R2 e Φ :
+  to_val e = None → E2 ⊆ E1 →
+  □ (R1 ={E1,E2}=★ ▷ |={E2,E1}=> R2) ∧ {{ P }} e @ E2 {{ Φ }}
+  ⊢ {{ R1 ★ P }} e @ E1 {{ λ v, R2 ★ Φ v }}.
 Proof.
-  iIntros (???) "[#Hvs1 [#Hvs2 #Hwp]] ! [HR HP]".
-  iApply (wp_frame_step_l E E1 E2); try done.
-  iSplitL "HR"; [|by iApply "Hwp"].
-  iPvs ("Hvs1" with "HR"); first by set_solver.
-  iPvsIntro. iNext. by iApply "Hvs2".
+  iIntros (??) "[#Hvs #Hwp] ! [HR HP]".
+  iApply (wp_frame_step_l E1 E2); try done.
+  iSplitL "HR"; [by iApply "Hvs"|by iApply "Hwp"].
 Qed.
 
-Lemma ht_frame_step_r E E1 E2 P R1 R2 R3 e Φ :
-  to_val e = None → E ⊥ E1 → E2 ⊆ E1 →
-  (R1 ={E1,E2}=> ▷ R2) ∧ (R2 ={E2,E1}=> R3) ∧ {{ P }} e @ E {{ Φ }}
-  ⊢ {{ P ★ R1 }} e @ (E ∪ E1) {{ λ v, Φ v ★ R3 }}.
+Lemma ht_frame_step_r E1 E2 P R1 R2 e Φ :
+  to_val e = None → E2 ⊆ E1 →
+  □ (R1 ={E1,E2}=★ ▷ |={E2,E1}=> R2) ∧ {{ P }} e @ E2 {{ Φ }}
+  ⊢ {{ P ★ R1 }} e @ E1 {{ λ v, Φ v ★ R2 }}.
 Proof.
-  iIntros (???) "[#Hvs1 [#Hvs2 #Hwp]]".
-  setoid_rewrite (comm _ _ R3); rewrite (comm _ _ R1).
-  iApply (ht_frame_step_l _ _ E2); by repeat iSplit.
+  iIntros (??) "[#Hvs #Hwp] ! [HP HR]".
+  iApply (wp_frame_step_r E1 E2); try done.
+  iSplitR "HR"; [by iApply "Hwp"|by iApply "Hvs"].
 Qed.
 
 Lemma ht_frame_step_l' E P R e Φ :
@@ -139,12 +137,4 @@ Proof.
   iIntros (?) "#Hwp ! [HP HR]".
   iApply wp_frame_step_r'; try done. iFrame "HR". by iApply "Hwp".
 Qed.
-
-Lemma ht_inv N E P Φ R e :
-  atomic e → nclose N ⊆ E →
-  inv N R ★ {{ ▷ R ★ P }} e @ E ∖ nclose N {{ v, ▷ R ★ Φ v }}
-  ⊢ {{ P }} e @ E {{ Φ }}.
-Proof.
-  iIntros (??) "[#? #Hwp] ! HP". iInv N as "HR". iApply "Hwp". by iSplitL "HR".
-Qed.
 End hoare.
diff --git a/program_logic/invariants.v b/program_logic/invariants.v
index 3c4808b2ffc73c4dba7ad98ad890b59be85e4720..a4e49288ced543b68dc5b7a61d9177c020947e7a 100644
--- a/program_logic/invariants.v
+++ b/program_logic/invariants.v
@@ -1,25 +1,27 @@
-From iris.program_logic Require Import ownership.
+From iris.program_logic Require Export pviewshifts.
 From iris.program_logic Require Export namespaces.
+From iris.program_logic Require Import ownership.
+From iris.algebra Require Import gmap.
 From iris.proofmode Require Import pviewshifts.
 Import uPred.
 
 (** Derived forms and lemmas about them. *)
-Definition inv_def {Λ Σ} (N : namespace) (P : iProp Λ Σ) : iProp Λ Σ :=
+Definition inv_def `{irisG Λ Σ} (N : namespace) (P : iProp Σ) : iProp Σ :=
   (∃ i, ■ (i ∈ nclose N) ∧ ownI i P)%I.
 Definition inv_aux : { x | x = @inv_def }. by eexists. Qed.
-Definition inv {Λ Σ} := proj1_sig inv_aux Λ Σ.
+Definition inv {Λ Σ i} := proj1_sig inv_aux Λ Σ i.
 Definition inv_eq : @inv = @inv_def := proj2_sig inv_aux.
-Instance: Params (@inv) 3.
+Instance: Params (@inv) 4.
 Typeclasses Opaque inv.
 
 Section inv.
-Context {Λ : language} {Σ : iFunctor}.
+Context `{irisG Λ Σ}.
 Implicit Types i : positive.
 Implicit Types N : namespace.
-Implicit Types P Q R : iProp Λ Σ.
-Implicit Types Φ : val Λ → iProp Λ Σ.
+Implicit Types P Q R : iProp Σ.
+Implicit Types Φ : val Λ → iProp Σ.
 
-Global Instance inv_contractive N : Contractive (@inv Λ Σ N).
+Global Instance inv_contractive N : Contractive (inv N).
 Proof.
   rewrite inv_eq=> n ???. apply exist_ne=>i. by apply and_ne, ownI_contractive.
 Qed.
@@ -27,42 +29,35 @@ Qed.
 Global Instance inv_persistent N P : PersistentP (inv N P).
 Proof. rewrite inv_eq /inv; apply _. Qed.
 
-Lemma always_inv N P : □ inv N P ⊣⊢ inv N P.
-Proof. by rewrite always_always. Qed.
-
-Lemma inv_alloc N E P : nclose N ⊆ E → ▷ P ={E}=> inv N P.
+Lemma inv_alloc N E P : â–· P ={E}=> inv N P.
 Proof.
-  intros. rewrite -(pvs_mask_weaken N) //.
-  by rewrite inv_eq /inv (pvs_allocI N); last apply nclose_infinite.
+  rewrite inv_eq /inv_def pvs_eq /pvs_def. iIntros "HP [Hw $]".
+  iVs (ownI_alloc (∈ nclose N) P with "[HP Hw]") as (i) "(% & $ & ?)"; auto.
+  - intros Ef. exists (coPpick (nclose N ∖ coPset.of_gset Ef)).
+    rewrite -coPset.elem_of_of_gset comm -elem_of_difference.
+    apply coPpick_elem_of=> Hfin.
+    eapply nclose_infinite, (difference_finite_inv _ _), Hfin.
+    apply of_gset_finite.
+  - by iFrame.
+  - rewrite /uPred_now_True; eauto.
 Qed.
 
-(** Fairly explicit form of opening invariants *)
 Lemma inv_open E N P :
-  nclose N ⊆ E →
-  inv N P ⊢ ∃ E', ■ (E ∖ nclose N ⊆ E' ⊆ E) ★
-                    |={E,E'}=> ▷ P ★ (▷ P ={E',E}=★ True).
+  nclose N ⊆ E → inv N P ={E,E∖N}=> ▷ P ★ (▷ P ={E∖N,E}=★ True).
 Proof.
-  rewrite inv_eq /inv. iDestruct 1 as (i) "[% #Hi]".
-  iExists (E ∖ {[ i ]}). iSplit; first (iPureIntro; set_solver).
-  iPvs (pvs_openI' with "Hi") as "HP"; [set_solver..|].
-  iPvsIntro. iSplitL "HP"; first done. iIntros "HP".
-  iPvs (pvs_closeI' _ _ P with "[HP]"); [set_solver|iSplit; done|set_solver|].
-  done.
+  rewrite inv_eq /inv_def pvs_eq /pvs_def; iDestruct 1 as (i) "[Hi #HiP]".
+  iDestruct "Hi" as % ?%elem_of_subseteq_singleton.
+  rewrite {1 4}(union_difference_L (nclose N) E) // ownE_op; last set_solver.
+  rewrite {1 5}(union_difference_L {[ i ]} (nclose N)) // ownE_op; last set_solver.
+  iIntros "(Hw & [HE $] & $)"; iVsIntro; iRight.
+  iDestruct (ownI_open i P with "[Hw HE]") as "($ & $ & HD)"; first by iFrame.
+  iIntros "HP [Hw $]"; iVsIntro; iRight. iApply ownI_close; by iFrame.
 Qed.
 
-(** Invariants can be opened around any frame-shifting assertion. This is less
-    verbose to apply than [inv_open]. *)
-Lemma inv_fsa {A} (fsa : FSA Λ Σ A) `{!FrameShiftAssertion fsaV fsa} E N P Ψ :
-  fsaV → nclose N ⊆ E →
-  inv N P ★ (▷ P -★ fsa (E ∖ nclose N) (λ a, ▷ P ★ Ψ a)) ⊢ fsa E Ψ.
+Lemma inv_open_timeless E N P `{!TimelessP P} :
+  nclose N ⊆ E → inv N P ={E,E∖N}=> P ★ (P ={E∖N,E}=★ True).
 Proof.
-  iIntros (??) "[Hinv HΨ]".
-  iDestruct (inv_open E N P with "Hinv") as (E') "[% Hvs]"; first done.
-  iApply (fsa_open_close E E'); auto; first set_solver.
-  iPvs "Hvs" as "[HP Hvs]"; first set_solver.
-  (* TODO: How do I do sth. like [iSpecialize "HΨ HP"]? *)
-  iPvsIntro. iApply (fsa_mask_weaken _ (E ∖ N)); first set_solver.
-  iApply fsa_wand_r. iSplitR "Hvs"; first by iApply "HΨ"; simpl.
-  iIntros (v) "[HP HΨ]". by iPvs ("Hvs" with "HP"); first set_solver.
+  iIntros (?) "Hinv". iVs (inv_open with "Hinv") as "[HP Hclose]"; auto.
+  iTimeless "HP"; iVsIntro; iIntros "{$HP} HP". iApply "Hclose"; auto.
 Qed.
 End inv.
diff --git a/program_logic/lifting.v b/program_logic/lifting.v
index 6b98cf087a3aea1de5f51eb74755aa4bd9e2e7f3..19e113e8111816db04d09adb6f8907b326e598b4 100644
--- a/program_logic/lifting.v
+++ b/program_logic/lifting.v
@@ -1,43 +1,28 @@
 From iris.program_logic Require Export weakestpre.
-From iris.program_logic Require Import wsat ownership.
+From iris.program_logic Require Import ownership.
 From iris.proofmode Require Import pviewshifts.
-Local Hint Extern 10 (_ ≤ _) => omega.
-Local Hint Extern 100 (_ ⊥ _) => set_solver.
-Local Hint Extern 10 (✓{_} _) =>
-  repeat match goal with
-  | H : wsat _ _ _ _ |- _ => apply wsat_valid in H; last omega
-  end; solve_validN.
 
 Section lifting.
-Context {Λ : language} {Σ : iFunctor}.
+Context `{irisG Λ Σ}.
 Implicit Types v : val Λ.
 Implicit Types e : expr Λ.
 Implicit Types σ : state Λ.
-Implicit Types P Q : iProp Λ Σ.
-Implicit Types Φ : val Λ → iProp Λ Σ.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val Λ → iProp Σ.
 
-Notation wp_fork ef := (default True ef (flip (wp ⊤) (λ _, True)))%I.
-
-Lemma wp_lift_step E1 E2 Φ e1 :
-  E2 ⊆ E1 → to_val e1 = None →
-  (|={E1,E2}=> ∃ σ1, ■ reducible e1 σ1 ∧ ▷ ownP σ1 ★
-       ▷ ∀ e2 σ2 ef, (■ prim_step e1 σ1 e2 σ2 ef ∧ ownP σ2)
-                     ={E2,E1}=★ WP e2 @ E1 {{ Φ }} ★ wp_fork ef)
-  ⊢ WP e1 @ E1 {{ Φ }}.
+Lemma wp_lift_step E Φ e1 :
+  to_val e1 = None →
+  (|={E,∅}=> ∃ σ1, ■ reducible e1 σ1 ★ ▷ ownP σ1 ★
+     ▷ ∀ e2 σ2 ef, ■ prim_step e1 σ1 e2 σ2 ef ★ ownP σ2
+          ={∅,E}=★ WP e2 @ E {{ Φ }} ★ wp_fork ef)
+  ⊢ WP e1 @ E {{ Φ }}.
 Proof.
-  intros ? He. rewrite pvs_eq wp_eq.
-  uPred.unseal; split=> n r ? Hvs; constructor; auto. intros k Ef σ1' rf ???.
-  destruct (Hvs (S k) Ef σ1' rf) as (r'&(σ1&Hsafe&r1&r2&?&?&Hwp)&Hws);
-    auto; clear Hvs; cofe_subst r'.
-  destruct (wsat_update_pst k (E2 ∪ Ef) σ1 σ1' r1 (r2 ⋅ rf)) as [-> Hws'].
-  { apply equiv_dist. rewrite -(ownP_spec k); auto. }
-  { by rewrite assoc. }
-  constructor; [done|intros e2 σ2 ef ?; specialize (Hws' σ2)].
-  destruct (λ H1 H2 H3, Hwp e2 σ2 ef k (update_pst σ2 r1) H1 H2 H3 k Ef σ2 rf)
-    as (r'&(r1'&r2'&?&?&?)&?); auto; cofe_subst r'.
-  { split. done. apply ownP_spec; auto. }
-  { rewrite (comm _ r2) -assoc; eauto using wsat_le. }
-  exists r1', r2'; split_and?; try done. by uPred.unseal; intros ? ->.
+  iIntros (?) "H". rewrite wp_unfold /wp_pre; iRight; iSplit; auto.
+  iIntros (σ1) "Hσ". iVs "H" as (σ1') "(% & Hσf & H)". iTimeless "Hσf".
+  iDestruct (ownP_agree σ1 σ1' with "[#]") as %<-; first by iFrame.
+  iVsIntro; iSplit; [done|]; iNext; iIntros (e2 σ2 ef Hstep).
+  iVs (ownP_update σ1 σ2 with "[-H]") as "[$ ?]"; first by iFrame.
+  iApply "H"; eauto.
 Qed.
 
 Lemma wp_lift_pure_step E Φ e1 :
@@ -47,19 +32,14 @@ Lemma wp_lift_pure_step E Φ e1 :
   (▷ ∀ e2 ef σ, ■ prim_step e1 σ e2 σ ef → WP e2 @ E {{ Φ }} ★ wp_fork ef)
   ⊢ WP e1 @ E {{ Φ }}.
 Proof.
-  intros He Hsafe Hstep; rewrite wp_eq; uPred.unseal.
-  split=> n r ? Hwp; constructor; auto.
-  intros k Ef σ1 rf ???; split; [done|]. destruct n as [|n]; first lia.
-  intros e2 σ2 ef ?; destruct (Hstep σ1 e2 σ2 ef); auto; subst.
-  destruct (Hwp e2 ef σ1 k r) as (r1&r2&Hr&?&?); auto.
-  exists r1,r2; split_and?; try done.
-  - rewrite -Hr; eauto using wsat_le.
-  - uPred.unseal; by intros ? ->.
+  iIntros (He Hsafe Hstep) "H". rewrite wp_unfold /wp_pre; iRight; iSplit; auto.
+  iIntros (σ1) "Hσ". iApply pvs_intro'; [set_solver|iIntros "Hclose"].
+  iSplit; [done|]; iNext; iIntros (e2 σ2 ef ?).
+  destruct (Hstep σ1 e2 σ2 ef); auto; subst.
+  iVs "Hclose"; iVsIntro. iFrame "Hσ". iApply "H"; auto.
 Qed.
 
 (** Derived lifting lemmas. *)
-Import uPred.
-
 Lemma wp_lift_atomic_step {E Φ} e1 σ1 :
   atomic e1 →
   reducible e1 σ1 →
@@ -67,12 +47,14 @@ Lemma wp_lift_atomic_step {E Φ} e1 σ1 :
     ■ prim_step e1 σ1 (of_val v2) σ2 ef ∧ ownP σ2 -★ (|={E}=> Φ v2) ★ wp_fork ef)
   ⊢ WP e1 @ E {{ Φ }}.
 Proof.
-  iIntros (Hatom ?) "[Hσ1 Hwp]". iApply (wp_lift_step E E _ e1); eauto using reducible_not_val.
-  iPvsIntro. iExists σ1. repeat iSplit; eauto 10.
-  iFrame. iNext. iIntros (e2 σ2 ef) "[% Hσ2]".
-  edestruct Hatom as [v2 Hv%of_to_val]; eauto. subst e2.
-  iDestruct ("Hwp" $! v2 σ2 ef with "[Hσ2]") as "[HΦ ?]". by eauto.
-  iFrame. iPvs "HΦ". iPvsIntro. iApply wp_value; auto using to_of_val.
+  iIntros (Hatomic ?) "[Hσ H]".
+  iApply (wp_lift_step E _ e1); eauto using reducible_not_val.
+  iApply pvs_intro'; [set_solver|iIntros "Hclose"].
+  iExists σ1. iFrame "Hσ"; iSplit; eauto.
+  iNext; iIntros (e2 σ2 ef) "[% Hσ]".
+  edestruct (Hatomic σ1 e2 σ2 ef) as [v2 <-%of_to_val]; eauto.
+  iDestruct ("H" $! v2 σ2 ef with "[Hσ]") as "[HΦ $]"; first by eauto.
+  iVs "Hclose". iVs "HΦ". iApply wp_value; auto using to_of_val.
 Qed.
 
 Lemma wp_lift_atomic_det_step {E Φ e1} σ1 v2 σ2 ef :
diff --git a/program_logic/model.v b/program_logic/model.v
index 55cdb0ac6f870c12c7e3614ab91ac8afc1e19c41..6603df4d70964ea1556376fb778f77ef81365186 100644
--- a/program_logic/model.v
+++ b/program_logic/model.v
@@ -1,61 +1,68 @@
 From iris.algebra Require Export upred.
-From iris.program_logic Require Export resources.
+From iris.algebra Require Import iprod gmap.
 From iris.algebra Require cofe_solver.
 
-(* The Iris program logic is parametrized by a locally contractive functor
-from the category of COFEs to the category of CMRAs, which is instantiated
-with [iProp]. The [iProp] can be used to construct impredicate CMRAs, such as
-the stored propositions using the agreement CMRA. *)
-Structure iFunctor := IFunctor {
-  iFunctor_F :> urFunctor;
-  iFunctor_contractive : urFunctorContractive iFunctor_F
+(* The Iris program logic is parametrized by a dependent product of a bunch of
+[gFunctor]s, which are locally contractive functor from the category of COFEs to
+the category of CMRAs. These functors are instantiated with [iProp], the type
+of Iris propositions, which allows one to construct impredicate CMRAs, such as
+invariants and stored propositions using the agreement CMRA. *)
+Structure gFunctor := GFunctor {
+  gFunctor_F :> rFunctor;
+  gFunctor_contractive : rFunctorContractive gFunctor_F;
 }.
-Arguments IFunctor _ {_}.
-Existing Instances iFunctor_contractive.
+Arguments GFunctor _ {_}.
+Existing Instance gFunctor_contractive.
 
+(** The global CMRA: Indexed product over a gid i to (gname --fin--> Σ i) *)
+Definition gFunctors := { n : nat & fin n → gFunctor }.
+Definition gid (Σ : gFunctors) := fin (projT1 Σ).
+
+(** Name of one instance of a particular CMRA in the ghost state. *)
+Definition gname := positive.
+
+(** Solution of the recursive domain equation *)
 Module Type iProp_solution_sig.
-Parameter iPreProp : language → iFunctor → cofeT.
-Definition iGst (Λ : language) (Σ : iFunctor) : ucmraT := Σ (iPreProp Λ Σ).
-Definition iRes Λ Σ := res Λ (laterC (iPreProp Λ Σ)) (iGst Λ Σ).
-Definition iResUR Λ Σ := resUR Λ (laterC (iPreProp Λ Σ)) (iGst Λ Σ).
-Definition iWld Λ Σ := gmap positive (agree (laterC (iPreProp Λ Σ))).
-Definition iPst Λ := option (excl (state Λ)).
-Definition iProp (Λ : language) (Σ : iFunctor) : cofeT := uPredC (iResUR Λ Σ).
-
-Parameter iProp_unfold: ∀ {Λ Σ}, iProp Λ Σ -n> iPreProp Λ Σ.
-Parameter iProp_fold: ∀ {Λ Σ}, iPreProp Λ Σ -n> iProp Λ Σ.
-Parameter iProp_fold_unfold: ∀ {Λ Σ} (P : iProp Λ Σ),
+Parameter iPreProp : gFunctors → cofeT.
+Definition iResUR (Σ : gFunctors) : ucmraT :=
+  iprodUR (λ i, gmapUR gname (projT2 Σ i (iPreProp Σ))).
+Definition iProp (Σ : gFunctors) : cofeT := uPredC (iResUR Σ).
+
+Parameter iProp_unfold: ∀ {Σ}, iProp Σ -n> iPreProp Σ.
+Parameter iProp_fold: ∀ {Σ}, iPreProp Σ -n> iProp Σ.
+Parameter iProp_fold_unfold: ∀ {Σ} (P : iProp Σ),
   iProp_fold (iProp_unfold P) ≡ P.
-Parameter iProp_unfold_fold: ∀ {Λ Σ} (P : iPreProp Λ Σ),
+Parameter iProp_unfold_fold: ∀ {Σ} (P : iPreProp Σ),
   iProp_unfold (iProp_fold P) ≡ P.
 End iProp_solution_sig.
 
 Module Export iProp_solution : iProp_solution_sig.
 Import cofe_solver.
-Definition iProp_result (Λ : language) (Σ : iFunctor) :
-  solution (uPredCF (resURF Λ (▶ ∙) Σ)) := solver.result _.
-
-Definition iPreProp (Λ : language) (Σ : iFunctor) : cofeT := iProp_result Λ Σ.
-Definition iGst (Λ : language) (Σ : iFunctor) : ucmraT := Σ (iPreProp Λ Σ).
-Definition iRes Λ Σ := res Λ (laterC (iPreProp Λ Σ)) (iGst Λ Σ).
-Definition iResUR Λ Σ := resUR Λ (laterC (iPreProp Λ Σ)) (iGst Λ Σ).
-Definition iWld Λ Σ := gmap positive (agree (laterC (iPreProp Λ Σ))).
-Definition iPst Λ := option (excl (state Λ)).
-
-Definition iProp (Λ : language) (Σ : iFunctor) : cofeT := uPredC (iResUR Λ Σ).
-Definition iProp_unfold {Λ Σ} : iProp Λ Σ -n> iPreProp Λ Σ :=
-  solution_fold (iProp_result Λ Σ).
-Definition iProp_fold {Λ Σ} : iPreProp Λ Σ -n> iProp Λ Σ := solution_unfold _.
-Lemma iProp_fold_unfold {Λ Σ} (P : iProp Λ Σ) : iProp_fold (iProp_unfold P) ≡ P.
+Definition iResF (Σ : gFunctors) : urFunctor :=
+  iprodURF (λ i, gmapURF gname (projT2 Σ i)).
+Definition iProp_result (Σ : gFunctors) :
+  solution (uPredCF (iResF Σ)) := solver.result _.
+
+Definition iPreProp (Σ : gFunctors) : cofeT := iProp_result Σ.
+Definition iResUR (Σ : gFunctors) : ucmraT :=
+  iprodUR (λ i, gmapUR gname (projT2 Σ i (iPreProp Σ))).
+Definition iProp (Σ : gFunctors) : cofeT := uPredC (iResUR Σ).
+
+Definition iProp_unfold {Σ} : iProp Σ -n> iPreProp Σ :=
+  solution_fold (iProp_result Σ).
+Definition iProp_fold {Σ} : iPreProp Σ -n> iProp Σ := solution_unfold _.
+Lemma iProp_fold_unfold {Σ} (P : iProp Σ) : iProp_fold (iProp_unfold P) ≡ P.
 Proof. apply solution_unfold_fold. Qed.
-Lemma iProp_unfold_fold {Λ Σ} (P : iPreProp Λ Σ) :
-  iProp_unfold (iProp_fold P) ≡ P.
+Lemma iProp_unfold_fold {Σ} (P : iPreProp Σ) : iProp_unfold (iProp_fold P) ≡ P.
 Proof. apply solution_fold_unfold. Qed.
 End iProp_solution.
 
 Bind Scope uPred_scope with iProp.
 
-Instance iProp_fold_inj n Λ Σ : Inj (dist n) (dist n) (@iProp_fold Λ Σ).
-Proof. by intros X Y H; rewrite -(iProp_unfold_fold X) H iProp_unfold_fold. Qed.
-Instance iProp_unfold_inj n Λ Σ : Inj (dist n) (dist n) (@iProp_unfold Λ Σ).
-Proof. by intros X Y H; rewrite -(iProp_fold_unfold X) H iProp_fold_unfold. Qed.
+Lemma iProp_unfold_equivI {Σ} (P Q : iProp Σ) :
+  iProp_unfold P ≡ iProp_unfold Q ⊢ (P ≡ Q : iProp Σ).
+Proof.
+  rewrite -{2}(iProp_fold_unfold P) -{2}(iProp_fold_unfold Q).
+  eapply (uPred.eq_rewrite _ _ (λ z,
+    iProp_fold (iProp_unfold P) ≡ iProp_fold z))%I; auto with I; solve_proper.
+Qed.
diff --git a/program_logic/ownership.v b/program_logic/ownership.v
index f7dfc25aad62c7a9991dbb34c527da90dde88a9a..ad1ae6176d518c8971b2e7bc352dffb408dc0386 100644
--- a/program_logic/ownership.v
+++ b/program_logic/ownership.v
@@ -1,83 +1,173 @@
-From iris.program_logic Require Export model.
+From iris.program_logic Require Export iris.
+From iris.algebra Require Import gmap auth agree gset coPset upred_big_op.
+From iris.proofmode Require Import ghost_ownership tactics.
 
-Definition ownI {Λ Σ} (i : positive) (P : iProp Λ Σ) : iProp Λ Σ :=
-  uPred_ownM (Res {[ i := to_agree (Next (iProp_unfold P)) ]} ∅ ∅).
-Arguments ownI {_ _} _ _%I.
-Definition ownP {Λ Σ} (σ: state Λ) : iProp Λ Σ := uPred_ownM (Res ∅ (Excl' σ) ∅).
-Definition ownG {Λ Σ} (m: iGst Λ Σ) : iProp Λ Σ := uPred_ownM (Res ∅ ∅ m).
-Instance: Params (@ownI) 3.
-Instance: Params (@ownP) 2.
-Instance: Params (@ownG) 2.
+Definition invariant_unfold {Σ} (P : iProp Σ) : agree (later (iPreProp Σ)) :=
+  to_agree (Next (iProp_unfold P)).
+Definition ownI `{irisG Λ Σ} (i : positive) (P : iProp Σ) : iProp Σ :=
+  own invariant_name (â—¯ {[ i := invariant_unfold P ]}).
+Arguments ownI {_ _ _} _ _%I.
+Typeclasses Opaque ownI.
+Instance: Params (@ownI) 4.
 
-Typeclasses Opaque ownI ownG ownP.
+Definition ownP `{irisG Λ Σ} (σ : state Λ) : iProp Σ :=
+  own state_name (◯ (Excl' σ)).
+Typeclasses Opaque ownP.
+Instance: Params (@ownP) 4.
+
+Definition ownP_auth `{irisG Λ Σ} (σ : state Λ) : iProp Σ :=
+  own state_name (● (Excl' σ)).
+Typeclasses Opaque ownP_auth.
+Instance: Params (@ownP_auth) 4.
+
+Definition ownE `{irisG Λ Σ} (E : coPset) : iProp Σ :=
+  own enabled_name (CoPset E).
+Typeclasses Opaque ownE.
+Instance: Params (@ownE) 4.
+
+Definition ownD `{irisG Λ Σ} (E : gset positive) : iProp Σ :=
+  own disabled_name (GSet E).
+Typeclasses Opaque ownD.
+Instance: Params (@ownD) 4.
+
+Definition wsat `{irisG Λ Σ} : iProp Σ :=
+  (∃ I : gmap positive (iProp Σ),
+    own invariant_name (● (invariant_unfold <$> I : gmap _ _)) ★
+    [★ map] i ↦ Q ∈ I, ▷ Q ★ ownD {[i]} ∨ ownE {[i]})%I.
 
 Section ownership.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types r : iRes Λ Σ.
+Context `{irisG Λ Σ}.
 Implicit Types σ : state Λ.
-Implicit Types P : iProp Λ Σ.
-Implicit Types m : iGst Λ Σ.
+Implicit Types P : iProp Σ.
+
+(* Allocation *)
+Lemma iris_alloc `{irisPreG Λ Σ} σ :
+  True =r=> ∃ _ : irisG Λ Σ, wsat ★ ownE ⊤ ★ ownP_auth σ ★ ownP σ.
+Proof.
+  iIntros.
+  iVs (own_alloc (● (Excl' σ) ⋅ ◯ (Excl' σ))) as (γσ) "[Hσ Hσ']"; first done.
+  iVs (own_alloc (● (∅ : gmap _ _))) as (γI) "HI"; first done.
+  iVs (own_alloc (CoPset ⊤)) as (γE) "HE"; first done.
+  iVs (own_alloc (GSet ∅)) as (γD) "HD"; first done.
+  iVsIntro; iExists (IrisG _ _ _ γσ γI γE γD).
+  rewrite /wsat /ownE /ownP_auth /ownP; iFrame.
+  iExists ∅. rewrite fmap_empty big_sepM_empty. by iFrame.
+Qed.
+
+(* Physical state *)
+Lemma ownP_twice σ1 σ2 : ownP σ1 ★ ownP σ2 ⊢ False.
+Proof. rewrite /ownP -own_op own_valid. by iIntros (?). Qed.
+Global Instance ownP_timeless σ : TimelessP (@ownP Λ Σ _ σ).
+Proof. rewrite /ownP; apply _. Qed.
+
+Lemma ownP_agree σ1 σ2 : ownP_auth σ1 ★ ownP σ2 ⊢ σ1 = σ2.
+Proof.
+  rewrite /ownP /ownP_auth -own_op own_valid -auth_both_op.
+  by iIntros ([[[] [=]%leibniz_equiv] _]%auth_valid_discrete).
+Qed.
+Lemma ownP_update σ1 σ2 : ownP_auth σ1 ★ ownP σ1 =r=> ownP_auth σ2 ★ ownP σ2.
+Proof.
+  rewrite /ownP -!own_op. by apply own_update, auth_update_no_frame,
+    option_local_update, exclusive_local_update.
+Qed.
 
 (* Invariants *)
-Global Instance ownI_contractive i : Contractive (@ownI Λ Σ i).
+Global Instance ownI_contractive i : Contractive (@ownI Λ Σ _ i).
 Proof.
-  intros n P Q HPQ. rewrite /ownI.
-  apply uPred.ownM_ne, Res_ne; auto; apply singleton_ne, to_agree_ne.
-  by apply Next_contractive=> j ?; rewrite (HPQ j).
+  intros n P Q HPQ. rewrite /ownI /invariant_unfold. do 4 f_equiv.
+  apply Next_contractive=> j ?; by rewrite (HPQ j).
 Qed.
 Global Instance ownI_persistent i P : PersistentP (ownI i P).
 Proof. rewrite /ownI. apply _. Qed.
 
-(* physical state *)
-Lemma ownP_twice σ1 σ2 : ownP σ1 ★ ownP σ2 ⊢ (False : iProp Λ Σ).
+Lemma ownE_empty : True =r=> ownE ∅.
+Proof. by rewrite (own_empty (A:=coPset_disjUR) enabled_name). Qed.
+Lemma ownE_op E1 E2 : E1 ⊥ E2 → ownE (E1 ∪ E2) ⊣⊢ ownE E1 ★ ownE E2.
+Proof. intros. by rewrite /ownE -own_op coPset_disj_union. Qed.
+Lemma ownE_disjoint E1 E2 : ownE E1 ★ ownE E2 ⊢ E1 ⊥ E2.
+Proof. rewrite /ownE -own_op own_valid. by iIntros (?%coPset_disj_valid_op). Qed.
+Lemma ownE_op' E1 E2 : E1 ⊥ E2 ∧ ownE (E1 ∪ E2) ⊣⊢ ownE E1 ★ ownE E2.
 Proof.
-  rewrite /ownP -uPred.ownM_op Res_op.
-  by apply uPred.ownM_invalid; intros (_&?&_).
+  iSplit; [iIntros "[% ?]"; by iApply ownE_op|].
+  iIntros "HE". iDestruct (ownE_disjoint with "#HE") as %?.
+  iSplit; first done. iApply ownE_op; by try iFrame.
 Qed.
-Global Instance ownP_timeless σ : TimelessP (@ownP Λ Σ σ).
-Proof. rewrite /ownP; apply _. Qed.
+Lemma ownE_singleton_twice i : ownE {[i]} ★ ownE {[i]} ⊢ False.
+Proof. rewrite ownE_disjoint. iIntros (?); set_solver. Qed.
 
-(* ghost state *)
-Global Instance ownG_ne n : Proper (dist n ==> dist n) (@ownG Λ Σ).
-Proof. solve_proper. Qed.
-Global Instance ownG_proper : Proper ((≡) ==> (⊣⊢)) (@ownG Λ Σ) := ne_proper _.
-Lemma ownG_op m1 m2 : ownG (m1 ⋅ m2) ⊣⊢ ownG m1 ★ ownG m2.
-Proof. by rewrite /ownG -uPred.ownM_op Res_op !left_id. Qed.
-Global Instance ownG_mono : Proper (flip (≼) ==> (⊢)) (@ownG Λ Σ).
-Proof. move=>a b [c H]. rewrite H ownG_op. eauto with I. Qed.
-Lemma ownG_valid m : ownG m ⊢ ✓ m.
-Proof. rewrite /ownG uPred.ownM_valid res_validI /=; auto with I. Qed.
-Lemma ownG_valid_r m : ownG m ⊢ ownG m ★ ✓ m.
-Proof. apply (uPred.always_entails_r _ _), ownG_valid. Qed.
-Lemma ownG_empty : True ⊢ (ownG ∅ : iProp Λ Σ).
-Proof. apply: uPred.ownM_empty. Qed.
-Global Instance ownG_timeless m : Timeless m → TimelessP (ownG m).
-Proof. rewrite /ownG; apply _. Qed.
-Global Instance ownG_persistent m : Persistent m → PersistentP (ownG m).
-Proof. rewrite /ownG; apply _. Qed.
-
-(* inversion lemmas *)
-Lemma ownI_spec n r i P :
-  ✓{n} r →
-  (ownI i P) n r ↔ wld r !! i ≡{n}≡ Some (to_agree (Next (iProp_unfold P))).
+Lemma ownD_empty : True =r=> ownD ∅.
+Proof. by rewrite (own_empty (A:=gset_disjUR _) disabled_name). Qed.
+Lemma ownD_op E1 E2 : E1 ⊥ E2 → ownD (E1 ∪ E2) ⊣⊢ ownD E1 ★ ownD E2.
+Proof. intros. by rewrite /ownD -own_op gset_disj_union. Qed.
+Lemma ownD_disjoint E1 E2 : ownD E1 ★ ownD E2 ⊢ E1 ⊥ E2.
+Proof. rewrite /ownD -own_op own_valid. by iIntros (?%gset_disj_valid_op). Qed.
+Lemma ownD_op' E1 E2 : E1 ⊥ E2 ∧ ownD (E1 ∪ E2) ⊣⊢ ownD E1 ★ ownD E2.
 Proof.
-  intros (?&?&?). rewrite /ownI; uPred.unseal.
-  rewrite /uPred_holds/=res_includedN/= singleton_includedN; split.
-  - intros [(P'&Hi&HP) _]; rewrite Hi.
-    constructor; symmetry; apply agree_valid_includedN.
-    + by apply lookup_validN_Some with (wld r) i.
-    + by destruct HP as [?| ->].
-  - intros ?; split_and?; try apply ucmra_unit_leastN; eauto.
+  iSplit; [iIntros "[% ?]"; by iApply ownD_op|].
+  iIntros "HE". iDestruct (ownD_disjoint with "#HE") as %?.
+  iSplit; first done. iApply ownD_op; by try iFrame.
 Qed.
-Lemma ownP_spec n r σ : ✓{n} r → (ownP σ) n r ↔ pst r ≡ Excl' σ.
+Lemma ownD_singleton_twice i : ownD {[i]} ★ ownD {[i]} ⊢ False.
+Proof. rewrite ownD_disjoint. iIntros (?); set_solver. Qed.
+
+Lemma invariant_lookup `{irisG Λ Σ} (I : gmap positive (iProp Σ)) i P :
+  own invariant_name (● (invariant_unfold <$> I : gmap _ _)) ★
+  own invariant_name (◯ {[i := invariant_unfold P]}) ⊢
+  ∃ Q, I !! i = Some Q ★ ▷ (Q ≡ P).
 Proof.
-  intros (?&?&?). rewrite /ownP; uPred.unseal.
-  rewrite /uPred_holds /= res_includedN /= Excl_includedN //.
-  rewrite (timeless_iff n). naive_solver (apply ucmra_unit_leastN).
+  rewrite -own_op own_valid auth_validI /=. iIntros "[#HI #HvI]".
+  iDestruct "HI" as (I') "HI". rewrite gmap_equivI gmap_validI.
+  iSpecialize ("HI" $! i). iSpecialize ("HvI" $! i).
+  rewrite left_id_L lookup_fmap lookup_op lookup_singleton uPred.option_equivI.
+  case: (I !! i)=> [Q|] /=; [|case: (I' !! i)=> [Q'|] /=; by iExFalso].
+  iExists Q; iSplit; first done.
+  iAssert (invariant_unfold Q ≡ invariant_unfold P)%I as "?".
+  { case: (I' !! i)=> [Q'|] //=.
+    iRewrite "HI" in "HvI". rewrite uPred.option_validI agree_validI.
+    iRewrite -"HvI" in "HI". by rewrite agree_idemp. }
+  rewrite /invariant_unfold.
+  by rewrite agree_equivI uPred.later_equivI /= iProp_unfold_equivI.
 Qed.
-Lemma ownG_spec n r m : (ownG m) n r ↔ m ≼{n} gst r.
+
+Lemma ownI_open i P : wsat ★ ownI i P ★ ownE {[i]} ⊢ wsat ★ ▷ P ★ ownD {[i]}.
+Proof.
+  rewrite /ownI. iIntros "(Hw & Hi & HiE)". iDestruct "Hw" as (I) "[? HI]".
+  iDestruct (invariant_lookup I i P with "[#]") as (Q) "[% #HPQ]"; [by iFrame|].
+  iDestruct (big_sepM_delete _ _ i with "HI") as "[[[HQ $]|?] HI]"; eauto.
+  - iSplitR "HQ"; last by iNext; iRewrite -"HPQ".
+    iExists I. iFrame "Hw". iApply (big_sepM_delete _ _ i); eauto.
+    iFrame "HI"; eauto.
+  - iDestruct (ownE_singleton_twice with "[#]") as %[]. by iFrame.
+Qed.
+Lemma ownI_close i P : wsat ★ ownI i P ★ ▷ P ★ ownD {[i]} ⊢ wsat ★ ownE {[i]}.
+Proof.
+  rewrite /ownI. iIntros "(Hw & Hi & HP & HiD)". iDestruct "Hw" as (I) "[? HI]".
+  iDestruct (invariant_lookup with "[#]") as (Q) "[% #HPQ]"; first by iFrame.
+  iDestruct (big_sepM_delete _ _ i with "HI") as "[[[HQ ?]|$] HI]"; eauto.
+  - iDestruct (ownD_singleton_twice with "[#]") as %[]. by iFrame.
+  - iExists I. iFrame "Hw". iApply (big_sepM_delete _ _ i); eauto.
+    iFrame "HI". iLeft. iFrame "HiD". by iNext; iRewrite "HPQ".
+Qed.
+
+Lemma ownI_alloc φ P :
+  (∀ E : gset positive, ∃ i, i ∉ E ∧ φ i) →
+  wsat ★ ▷ P =r=> ∃ i, ■ (φ i) ★ wsat ★ ownI i P.
 Proof.
-  rewrite /ownG; uPred.unseal.
-  rewrite /uPred_holds /= res_includedN; naive_solver (apply ucmra_unit_leastN).
+  iIntros (Hfresh) "[Hw HP]". iDestruct "Hw" as (I) "[? HI]".
+  iVs (own_empty (A:=gset_disjUR positive) disabled_name) as "HE".
+  iVs (own_updateP with "HE") as "HE".
+  { apply (gset_alloc_empty_updateP_strong' (λ i, I !! i = None ∧ φ i)).
+    intros E. destruct (Hfresh (E ∪ dom _ I))
+      as (i & [? HIi%not_elem_of_dom]%not_elem_of_union & ?); eauto. }
+  iDestruct "HE" as (X) "[Hi HE]"; iDestruct "Hi" as %(i & -> & HIi & ?).
+  iVs (own_update with "Hw") as "[Hw HiP]".
+  { apply (auth_update_no_frag _ {[ i := invariant_unfold P ]}).
+    apply alloc_unit_singleton_local_update; last done.
+    by rewrite /= lookup_fmap HIi. }
+  iVsIntro; iExists i;  iSplit; [done|]. rewrite /ownI; iFrame "HiP".
+  iExists (<[i:=P]>I); iSplitL "Hw".
+  { by rewrite fmap_insert insert_singleton_op ?lookup_fmap ?HIi. }
+  iApply (big_sepM_insert _ I); first done.
+  iFrame "HI". iLeft. by rewrite /ownD; iFrame.
 Qed.
 End ownership.
diff --git a/program_logic/pviewshifts.v b/program_logic/pviewshifts.v
index 9cbab38f8cd2baad9df3dca216fde285eec8567c..9c14f02f75af92cba7876558a410cdd4d37013a0 100644
--- a/program_logic/pviewshifts.v
+++ b/program_logic/pviewshifts.v
@@ -1,290 +1,137 @@
-From iris.prelude Require Export coPset.
-From iris.algebra Require Export upred_big_op updates.
-From iris.program_logic Require Export model.
-From iris.program_logic Require Import ownership wsat.
-Local Hint Extern 10 (_ ≤ _) => omega.
-Local Hint Extern 100 (_ ⊥ _) => set_solver.
-Local Hint Extern 100 (_ ∉ _) => set_solver.
-Local Hint Extern 10 (✓{_} _) =>
-  repeat match goal with
-  | H : wsat _ _ _ _ |- _ => apply wsat_valid in H; last omega
-  end; solve_validN.
-
-Program Definition pvs_def {Λ Σ} (E1 E2 : coPset) (P : iProp Λ Σ) : iProp Λ Σ :=
-  {| uPred_holds n r1 := ∀ k Ef σ rf,
-       0 < k ≤ n → E1 ∪ E2 ⊥ Ef →
-       wsat k (E1 ∪ Ef) σ (r1 ⋅ rf) →
-       ∃ r2, P k r2 ∧ wsat k (E2 ∪ Ef) σ (r2 ⋅ rf) |}.
-Next Obligation.
-  intros Λ Σ E1 E2 P n r1 r2 HP [r3 Hr2] k Ef σ rf ??.
-  rewrite (dist_le _ _ _ _ Hr2); last lia. intros Hws.
-  destruct (HP k Ef σ (r3 ⋅ rf)) as (r'&?&Hws'); rewrite ?(assoc op); auto.
-  exists (r' â‹… r3); rewrite -assoc; split; last done.
-  apply uPred_mono with r'; eauto using cmra_includedN_l.
-Qed.
-Next Obligation. naive_solver. Qed.
+From iris.program_logic Require Export iris.
+From iris.program_logic Require Import ownership.
+From iris.algebra Require Import upred_big_op gmap.
+From iris.proofmode Require Import tactics.
+Import uPred.
 
+Program Definition pvs_def `{irisG Λ Σ}
+    (E1 E2 : coPset) (P : iProp Σ) : iProp Σ :=
+  (wsat ★ ownE E1 =r=★ ◇ (wsat ★ ownE E2 ★ P))%I.
 Definition pvs_aux : { x | x = @pvs_def }. by eexists. Qed.
 Definition pvs := proj1_sig pvs_aux.
 Definition pvs_eq : @pvs = @pvs_def := proj2_sig pvs_aux.
-
-Arguments pvs {_ _} _ _ _%I.
-Instance: Params (@pvs) 4.
+Arguments pvs {_ _ _} _ _ _%I.
+Instance: Params (@pvs) 5.
 
 Notation "|={ E1 , E2 }=> Q" := (pvs E1 E2 Q%I)
   (at level 99, E1, E2 at level 50, Q at level 200,
    format "|={ E1 , E2 }=>  Q") : uPred_scope.
-Notation "|={ E }=> Q" := (pvs E E Q%I)
-  (at level 99, E at level 50, Q at level 200,
-   format "|={ E }=>  Q") : uPred_scope.
-Notation "|==> Q" := (pvs ⊤ ⊤ Q%I)
-  (at level 99, Q at level 200, format "|==>  Q") : uPred_scope.
-
-Notation "P ={ E1 , E2 }=> Q" := (P ⊢ |={E1,E2}=> Q)
-  (at level 99, E1, E2 at level 50, Q at level 200, only parsing) : C_scope.
-Notation "P ={ E }=> Q" := (P ⊢ |={E}=> Q)
-  (at level 99, E at level 50, Q at level 200, only parsing) : C_scope.
-
 Notation "P ={ E1 , E2 }=★ Q" := (P -★ |={E1,E2}=> Q)%I
   (at level 99, E1,E2 at level 50, Q at level 200,
    format "P  ={ E1 , E2 }=★  Q") : uPred_scope.
-Notation "P ={ E }=★ Q" := (P ={E,E}=★ Q)%I
+Notation "P ={ E1 , E2 }=> Q" := (P ⊢ |={E1,E2}=> Q)
+  (at level 99, E1, E2 at level 50, Q at level 200, only parsing) : C_scope.
+
+Notation "|={ E }=> Q" := (pvs E E Q%I)
+  (at level 99, E at level 50, Q at level 200,
+   format "|={ E }=>  Q") : uPred_scope.
+Notation "P ={ E }=★ Q" := (P -★ |={E}=> Q)%I
   (at level 99, E at level 50, Q at level 200,
    format "P  ={ E }=★  Q") : uPred_scope.
+Notation "P ={ E }=> Q" := (P ⊢ |={E}=> Q)
+  (at level 99, E at level 50, Q at level 200, only parsing) : C_scope.
 
 Section pvs.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P Q : iProp Λ Σ.
-Implicit Types m : iGst Λ Σ.
-
-Lemma pvs_zero E1 E2 P r : pvs_def E1 E2 P 0 r.
-Proof. intros ?????. exfalso. omega. Qed.
+Context `{irisG Λ Σ}.
+Implicit Types P Q : iProp Σ.
 
-Global Instance pvs_ne E1 E2 n : Proper (dist n ==> dist n) (@pvs Λ Σ E1 E2).
-Proof.
-  rewrite pvs_eq.
-  intros P Q HPQ; split=> n' r1 ??; simpl; split; intros HP k Ef σ rf ???;
-    destruct (HP k Ef σ rf) as (r2&?&?); auto;
-    exists r2; split_and?; auto; apply HPQ; eauto.
-Qed.
-Global Instance pvs_proper E1 E2 : Proper ((≡) ==> (≡)) (@pvs Λ Σ E1 E2).
+Global Instance pvs_ne E1 E2 n : Proper (dist n ==> dist n) (@pvs Λ Σ _ E1 E2).
+Proof. rewrite pvs_eq. solve_proper. Qed.
+Global Instance pvs_proper E1 E2 : Proper ((≡) ==> (≡)) (@pvs Λ Σ _ E1 E2).
 Proof. apply ne_proper, _. Qed.
 
-Lemma pvs_intro E P : P ={E}=> P.
-Proof.
-  rewrite pvs_eq. split=> n r ? HP k Ef σ rf ???; exists r; split; last done.
-  apply uPred_closed with n; eauto.
-Qed.
-Lemma pvs_mono E1 E2 P Q : (P ⊢ Q) → (|={E1,E2}=> P) ={E1,E2}=> Q.
+Lemma pvs_intro' E1 E2 P : E2 ⊆ E1 → ((|={E2,E1}=> True) -★ P) ={E1,E2}=> P.
 Proof.
-  rewrite pvs_eq. intros HPQ; split=> n r ? HP k Ef σ rf ???.
-  destruct (HP k Ef σ rf) as (r2&?&?); eauto.
-  exists r2; eauto using uPred_in_entails.
+  intros (E1''&->&?)%subseteq_disjoint_union_L.
+  rewrite pvs_eq /pvs_def ownE_op //; iIntros "H ($ & $ & HE)".
+  iVsIntro. iApply now_True_intro. iApply "H".
+  iIntros "[$ $]". iVsIntro; iRight; eauto.
 Qed.
-Lemma pvs_timeless E P : TimelessP P → ▷ P ={E}=> P.
+Lemma now_True_pvs E1 E2 P : â—‡ (|={E1,E2}=> P) ={E1,E2}=> P.
 Proof.
-  rewrite pvs_eq uPred.timelessP_spec=> HP.
-  uPred.unseal; split=>-[|n] r ? HP' k Ef σ rf ???; first lia.
-  exists r; split; last done.
-  apply HP, uPred_closed with n; eauto using cmra_validN_le.
+  rewrite pvs_eq. iIntros "[?|H] [Hw HE]".
+  - rewrite /uPred_now_True; eauto.
+  - iApply "H"; by iFrame.
 Qed.
-Lemma pvs_trans E1 E2 E3 P :
-  E2 ⊆ E1 ∪ E3 → (|={E1,E2}=> |={E2,E3}=> P) ={E1,E3}=> P.
+Lemma rvs_pvs E P : (|=r=> P) ={E}=> P.
 Proof.
-  rewrite pvs_eq. intros ?; split=> n r1 ? HP1 k Ef σ rf ???.
-  destruct (HP1 k Ef σ rf) as (r2&HP2&?); auto.
+  rewrite pvs_eq /pvs_def. iIntros "H [$ $]". by rewrite -uPred.now_True_intro.
 Qed.
-Lemma pvs_mask_frame E1 E2 Ef P :
-  Ef ⊥ E1 ∪ E2 → (|={E1,E2}=> P) ={E1 ∪ Ef,E2 ∪ Ef}=> P.
-Proof.
-  rewrite pvs_eq. intros ?; split=> n r ? HP k Ef' σ rf ???.
-  destruct (HP k (Ef∪Ef') σ rf) as (r'&?&?); rewrite ?(assoc_L _); eauto.
-  by exists r'; rewrite -(assoc_L _).
-Qed.
-Lemma pvs_frame_r E1 E2 P Q : (|={E1,E2}=> P) ★ Q ={E1,E2}=> P ★ Q.
-Proof.
-  rewrite pvs_eq. uPred.unseal; split; intros n r ? (r1&r2&Hr&HP&?) k Ef σ rf ???.
-  destruct (HP k Ef σ (r2 ⋅ rf)) as (r'&?&?); eauto.
-  { by rewrite assoc -(dist_le _ _ _ _ Hr); last lia. }
-  exists (r' â‹… r2); split; last by rewrite -assoc.
-  exists r', r2; split_and?; auto. apply uPred_closed with n; auto.
-Qed.
-Lemma pvs_openI i P : ownI i P ={{[i]},∅}=> ▷ P.
-Proof.
-  rewrite pvs_eq. uPred.unseal; split=> -[|n] r ? Hinv [|k] Ef σ rf ???; try lia.
-  apply ownI_spec in Hinv; last auto.
-  destruct (wsat_open k Ef σ (r ⋅ rf) i P) as (rP&?&?); auto.
-  { rewrite lookup_wld_op_l ?Hinv; eauto; apply dist_le with (S n); eauto. }
-  exists (rP â‹… r); split; last by rewrite (left_id_L _ _) -assoc.
-  eapply uPred_mono with rP; eauto using cmra_includedN_l.
-Qed.
-Lemma pvs_closeI i P : ownI i P ∧ ▷ P ={∅,{[i]}}=> True.
+Lemma pvs_mono E1 E2 P Q : (P ⊢ Q) → (|={E1,E2}=> P) ={E1,E2}=> Q.
 Proof.
-  rewrite pvs_eq.
-  uPred.unseal; split=> -[|n] r ? [? HP] [|k] Ef σ rf ? HE ?; try lia.
-  exists ∅; split; [done|].
-  rewrite left_id; apply wsat_close with P r.
-  - apply ownI_spec, uPred_closed with (S n); auto.
-  - set_solver +HE.
-  - by rewrite -(left_id_L ∅ (∪) Ef).
-  - apply uPred_closed with n; auto.
+  rewrite pvs_eq /pvs_def. iIntros (HPQ) "HP HwE".
+  rewrite -HPQ. by iApply "HP".
 Qed.
-Lemma pvs_ownG_updateP E m (P : iGst Λ Σ → Prop) :
-  m ~~>: P → ownG m ={E}=> ∃ m', ■ P m' ∧ ownG m'.
+Lemma pvs_trans E1 E2 E3 P : (|={E1,E2}=> |={E2,E3}=> P) ={E1,E3}=> P.
 Proof.
-  rewrite pvs_eq. intros Hup.
-  uPred.unseal; split=> -[|n] r ? /ownG_spec Hinv [|k] Ef σ rf ???; try lia.
-  destruct (wsat_update_gst k (E ∪ Ef) σ r rf m P) as (m'&?&?); eauto.
-  { apply cmra_includedN_le with (S n); auto. }
-  by exists (update_gst m' r); split; [exists m'; split; [|apply ownG_spec]|].
+  rewrite pvs_eq /pvs_def. iIntros "HP HwE".
+  iVs ("HP" with "HwE") as "[?|(Hw & HE & HP)]".
+  - rewrite /uPred_now_True; eauto.
+  - iApply "HP"; by iFrame. 
 Qed.
-Lemma pvs_allocI E P : ¬set_finite E → ▷ P ={E}=> ∃ i, ■ (i ∈ E) ∧ ownI i P.
+Lemma pvs_mask_frame_r' E1 E2 Ef P :
+  E1 ⊥ Ef → (|={E1,E2}=> E2 ⊥ Ef → P) ={E1 ∪ Ef,E2 ∪ Ef}=> P.
 Proof.
-  rewrite pvs_eq. intros ?; rewrite /ownI; uPred.unseal.
-  split=> -[|n] r ? HP [|k] Ef σ rf ???; try lia.
-  destruct (wsat_alloc k E Ef σ rf P r) as (i&?&?&?); auto.
-  { apply uPred_closed with n; eauto. }
-  exists (Res {[ i := to_agree (Next (iProp_unfold P)) ]} ∅ ∅).
-  split; [|done]. by exists i; split; rewrite /uPred_holds /=.
+  intros. rewrite pvs_eq /pvs_def ownE_op //. iIntros "Hvs (Hw & HE1 &HEf)".
+  iVs ("Hvs" with "[Hw HE1]") as "[?|($ & HE2 & HP)]"; first by iFrame.
+  - rewrite /uPred_now_True; eauto.
+  - iDestruct (ownE_op' with "[HE2 HEf]") as "[? $]"; first by iFrame.
+    iVsIntro; iRight; by iApply "HP".
 Qed.
+Lemma pvs_frame_r E1 E2 P Q : (|={E1,E2}=> P) ★ Q ={E1,E2}=> P ★ Q.
+Proof. rewrite pvs_eq /pvs_def. by iIntros "[HwP $]". Qed.
 
 (** * Derived rules *)
-Import uPred.
-Global Instance pvs_mono' E1 E2 : Proper ((⊢) ==> (⊢)) (@pvs Λ Σ E1 E2).
+Global Instance pvs_mono' E1 E2 : Proper ((⊢) ==> (⊢)) (@pvs Λ Σ _ E1 E2).
 Proof. intros P Q; apply pvs_mono. Qed.
 Global Instance pvs_flip_mono' E1 E2 :
-  Proper (flip (⊢) ==> flip (⊢)) (@pvs Λ Σ E1 E2).
+  Proper (flip (⊢) ==> flip (⊢)) (@pvs Λ Σ _ E1 E2).
 Proof. intros P Q; apply pvs_mono. Qed.
-Lemma pvs_trans' E P : (|={E}=> |={E}=> P) ={E}=> P.
-Proof. apply pvs_trans; set_solver. Qed.
-Lemma pvs_strip_pvs E P Q : (P ={E}=> Q) → (|={E}=> P) ={E}=> Q.
-Proof. move=>->. by rewrite pvs_trans'. Qed.
+
+Lemma pvs_intro E P : P ={E}=> P.
+Proof. iIntros "HP". by iApply rvs_pvs. Qed.
+Lemma pvs_now_True E1 E2 P : (|={E1,E2}=> â—‡ P) ={E1,E2}=> P.
+Proof. by rewrite {1}(pvs_intro E2 P) now_True_pvs pvs_trans. Qed.
+
 Lemma pvs_frame_l E1 E2 P Q : (P ★ |={E1,E2}=> Q) ={E1,E2}=> P ★ Q.
 Proof. rewrite !(comm _ P); apply pvs_frame_r. Qed.
-Lemma pvs_always_l E1 E2 P Q `{!PersistentP P} :
-  P ∧ (|={E1,E2}=> Q) ={E1,E2}=> P ∧ Q.
-Proof. by rewrite !always_and_sep_l pvs_frame_l. Qed.
-Lemma pvs_always_r E1 E2 P Q `{!PersistentP Q} :
-  (|={E1,E2}=> P) ∧ Q ={E1,E2}=> P ∧ Q.
-Proof. by rewrite !always_and_sep_r pvs_frame_r. Qed.
-Lemma pvs_impl_l E1 E2 P Q : □ (P → Q) ∧ (|={E1,E2}=> P) ={E1,E2}=> Q.
-Proof. by rewrite pvs_always_l always_elim impl_elim_l. Qed.
-Lemma pvs_impl_r E1 E2 P Q : (|={E1,E2}=> P) ∧ □ (P → Q) ={E1,E2}=> Q.
-Proof. by rewrite comm pvs_impl_l. Qed.
 Lemma pvs_wand_l E1 E2 P Q : (P -★ Q) ★ (|={E1,E2}=> P) ={E1,E2}=> Q.
 Proof. by rewrite pvs_frame_l wand_elim_l. Qed.
 Lemma pvs_wand_r E1 E2 P Q : (|={E1,E2}=> P) ★ (P -★ Q) ={E1,E2}=> Q.
 Proof. by rewrite pvs_frame_r wand_elim_r. Qed.
+
+Lemma pvs_trans_frame E1 E2 E3 P Q :
+  ((Q ={E2,E3}=★ True) ★ |={E1,E2}=> (Q ★ P)) ={E1,E3}=> P.
+Proof.
+  rewrite pvs_frame_l assoc -(comm _ Q) wand_elim_r.
+  by rewrite pvs_frame_r left_id pvs_trans.
+Qed.
+
+Lemma pvs_mask_frame_r E1 E2 Ef P :
+  E1 ⊥ Ef → (|={E1,E2}=> P) ={E1 ∪ Ef,E2 ∪ Ef}=> P.
+Proof.
+  iIntros (?) "H". iApply pvs_mask_frame_r'; auto.
+  iApply pvs_wand_r; iFrame "H"; eauto.
+Qed.
+Lemma pvs_mask_mono E1 E2 P : E1 ⊆ E2 → (|={E1}=> P) ={E2}=> P.
+Proof.
+  intros (Ef&->&?)%subseteq_disjoint_union_L. by apply pvs_mask_frame_r.
+Qed.
+
 Lemma pvs_sep E P Q : (|={E}=> P) ★ (|={E}=> Q) ={E}=> P ★ Q.
-Proof. rewrite pvs_frame_r pvs_frame_l pvs_trans //. set_solver. Qed.
-Lemma pvs_big_sepM `{Countable K} {A} E (Φ : K → A → iProp Λ Σ) (m : gmap K A) :
+Proof. by rewrite pvs_frame_r pvs_frame_l pvs_trans. Qed.
+Lemma pvs_big_sepM `{Countable K} {A} E (Φ : K → A → iProp Σ) (m : gmap K A) :
   ([★ map] k↦x ∈ m, |={E}=> Φ k x) ={E}=> [★ map] k↦x ∈ m, Φ k x.
 Proof.
   rewrite /uPred_big_sepM.
   induction (map_to_list m) as [|[i x] l IH]; csimpl; auto using pvs_intro.
   by rewrite IH pvs_sep.
 Qed.
-Lemma pvs_big_sepS `{Countable A} E (Φ : A → iProp Λ Σ) X :
+Lemma pvs_big_sepS `{Countable A} E (Φ : A → iProp Σ) X :
   ([★ set] x ∈ X, |={E}=> Φ x) ={E}=> [★ set] x ∈ X, Φ x.
 Proof.
   rewrite /uPred_big_sepS.
   induction (elements X) as [|x l IH]; csimpl; csimpl; auto using pvs_intro.
   by rewrite IH pvs_sep.
 Qed.
-
-Lemma pvs_mask_frame' E1 E1' E2 E2' P :
-  E1' ⊆ E1 → E2' ⊆ E2 → E1 ∖ E1' = E2 ∖ E2' → (|={E1',E2'}=> P) ={E1,E2}=> P.
-Proof.
-  intros HE1 HE2 HEE.
-  rewrite (pvs_mask_frame _ _ (E1 ∖ E1')); last set_solver.
-  by rewrite {2}HEE -!union_difference_L.
-Qed.
-Lemma pvs_openI' E i P : i ∈ E → ownI i P ={E, E ∖ {[i]}}=> ▷ P.
-Proof. intros. etrans. apply pvs_openI. apply pvs_mask_frame'; set_solver. Qed.
-Lemma pvs_closeI' E i P : i ∈ E → (ownI i P ∧ ▷ P) ={E ∖ {[i]}, E}=> True.
-Proof. intros. etrans. apply pvs_closeI. apply pvs_mask_frame'; set_solver. Qed.
-
-Lemma pvs_mask_frame_mono E1 E1' E2 E2' P Q :
-  E1' ⊆ E1 → E2' ⊆ E2 → E1 ∖ E1' = E2 ∖ E2' →
-  (P ⊢ Q) → (|={E1',E2'}=> P) ={E1,E2}=> Q.
-Proof. intros HE1 HE2 HEE ->. by apply pvs_mask_frame'. Qed.
-
-(** It should be possible to give a stronger version of this rule
-   that does not force the conclusion view shift to have twice the
-   same mask. However, even expressing the side-conditions on the
-   mask becomes really ugly then, and we have not found an instance
-   where that would be useful. *)
-Lemma pvs_trans3 E1 E2 Q :
-  E2 ⊆ E1 → (|={E1,E2}=> |={E2}=> |={E2,E1}=> Q) ={E1}=> Q.
-Proof. intros HE. rewrite !pvs_trans; set_solver. Qed.
-
-Lemma pvs_mask_weaken E1 E2 P : E1 ⊆ E2 → (|={E1}=> P) ={E2}=> P.
-Proof. auto using pvs_mask_frame'. Qed.
-
-Lemma pvs_ownG_update E m m' : m ~~> m' → ownG m ={E}=> ownG m'.
-Proof.
-  intros; rewrite (pvs_ownG_updateP E _ (m' =)); last by apply cmra_update_updateP.
-  by apply pvs_mono, uPred.exist_elim=> m''; apply uPred.pure_elim_l=> ->.
-Qed.
 End pvs.
-
-(** * Frame Shift Assertions. *)
-(* Yes, the name is horrible...
-   Frame Shift Assertions take a mask and a predicate over some type (that's
-   their "postcondition"). They support weakening the mask, framing resources
-   into the postcondition, and composition witn mask-changing view shifts. *)
-Notation FSA Λ Σ A := (coPset → (A → iProp Λ Σ) → iProp Λ Σ).
-Class FrameShiftAssertion {Λ Σ A} (fsaV : Prop) (fsa : FSA Λ Σ A) := {
-  fsa_mask_frame_mono E1 E2 Φ Ψ :
-    E1 ⊆ E2 → (∀ a, Φ a ⊢ Ψ a) → fsa E1 Φ ⊢ fsa E2 Ψ;
-  fsa_trans3 E Φ : (|={E}=> fsa E (λ a, |={E}=> Φ a)) ⊢ fsa E Φ;
-  fsa_open_close E1 E2 Φ :
-    fsaV → E2 ⊆ E1 → (|={E1,E2}=> fsa E2 (λ a, |={E2,E1}=> Φ a)) ⊢ fsa E1 Φ;
-  fsa_frame_r E P Φ : (fsa E Φ ★ P) ⊢ fsa E (λ a, Φ a ★ P)
-}.
-
-(* Used to solve side-conditions related to [fsaV] *)
-Create HintDb fsaV.
-
-Section fsa.
-Context {Λ Σ A} (fsa : FSA Λ Σ A) `{!FrameShiftAssertion fsaV fsa}.
-Implicit Types Φ Ψ : A → iProp Λ Σ.
-Import uPred.
-
-Lemma fsa_mono E Φ Ψ : (∀ a, Φ a ⊢ Ψ a) → fsa E Φ ⊢ fsa E Ψ.
-Proof. apply fsa_mask_frame_mono; auto. Qed.
-Lemma fsa_mask_weaken E1 E2 Φ : E1 ⊆ E2 → fsa E1 Φ ⊢ fsa E2 Φ.
-Proof. intros. apply fsa_mask_frame_mono; auto. Qed.
-Lemma fsa_frame_l E P Φ : P ★ fsa E Φ ⊢ fsa E (λ a, P ★ Φ a).
-Proof. rewrite comm fsa_frame_r. apply fsa_mono=>a. by rewrite comm. Qed.
-Lemma fsa_strip_pvs E P Φ : (P ⊢ fsa E Φ) → (|={E}=> P) ⊢ fsa E Φ.
-Proof.
-  move=>->. rewrite -{2}fsa_trans3.
-  apply pvs_mono, fsa_mono=>a; apply pvs_intro.
-Qed.
-Lemma fsa_pvs_fsa E Φ : (|={E}=> fsa E Φ) ⊣⊢ fsa E Φ.
-Proof. apply (anti_symm (⊢)); [by apply fsa_strip_pvs|apply pvs_intro]. Qed.
-Lemma pvs_fsa_fsa E Φ : fsa E (λ a, |={E}=> Φ a) ⊣⊢ fsa E Φ.
-Proof.
-  apply (anti_symm (⊢)); [|apply fsa_mono=> a; apply pvs_intro].
-  by rewrite (pvs_intro E (fsa E _)) fsa_trans3.
-Qed.
-Lemma fsa_mono_pvs E Φ Ψ : (∀ a, Φ a ={E}=> Ψ a) → fsa E Φ ⊢ fsa E Ψ.
-Proof. intros. rewrite -[fsa E Ψ]fsa_trans3 -pvs_intro. by apply fsa_mono. Qed.
-Lemma fsa_wand_l E Φ Ψ : (∀ a, Φ a -★ Ψ a) ★ fsa E Φ ⊢ fsa E Ψ.
-Proof.
-  rewrite fsa_frame_l. apply fsa_mono=> a.
-  by rewrite (forall_elim a) wand_elim_l.
-Qed.
-Lemma fsa_wand_r E Φ Ψ : fsa E Φ ★ (∀ a, Φ a -★ Ψ a) ⊢ fsa E Ψ.
-Proof. by rewrite (comm _ (fsa _ _)) fsa_wand_l. Qed.
-End fsa.
-
-Definition pvs_fsa {Λ Σ} : FSA Λ Σ () := λ E Φ, (|={E}=> Φ ())%I.
-Arguments pvs_fsa _ _ _ _/.
-
-Instance pvs_fsa_prf {Λ Σ} : FrameShiftAssertion True (@pvs_fsa Λ Σ).
-Proof.
-  rewrite /pvs_fsa.
-  split; auto using pvs_mask_frame_mono, pvs_trans3, pvs_frame_r.
-Qed.
diff --git a/program_logic/resources.v b/program_logic/resources.v
deleted file mode 100644
index 1590ee7374bdb68b0b5ddd2fb16d1b7f7b28f680..0000000000000000000000000000000000000000
--- a/program_logic/resources.v
+++ /dev/null
@@ -1,253 +0,0 @@
-From iris.algebra Require Export gmap agree excl.
-From iris.algebra Require Import upred.
-From iris.program_logic Require Export language.
-
-Record res (Λ : language) (A : cofeT) (M : cmraT) := Res {
-  wld : gmapR positive (agreeR A);
-  pst : optionR (exclR (stateC Λ));
-  gst : M;
-}.
-Add Printing Constructor res.
-Arguments Res {_ _ _} _ _ _.
-Arguments wld {_ _ _} _.
-Arguments pst {_ _ _} _.
-Arguments gst {_ _ _} _.
-Instance: Params (@Res) 3.
-Instance: Params (@wld) 3.
-Instance: Params (@pst) 3.
-Instance: Params (@gst) 3.
-
-Section res.
-Context {Λ : language} {A : cofeT} {M : ucmraT}.
-Implicit Types r : res Λ A M.
-
-Inductive res_equiv' (r1 r2 : res Λ A M) := Res_equiv :
-  wld r1 ≡ wld r2 → pst r1 ≡ pst r2 → gst r1 ≡ gst r2 → res_equiv' r1 r2.
-Instance res_equiv : Equiv (res Λ A M) := res_equiv'.
-Inductive res_dist' (n : nat) (r1 r2 : res Λ A M) := Res_dist :
-  wld r1 ≡{n}≡ wld r2 → pst r1 ≡{n}≡ pst r2 → gst r1 ≡{n}≡ gst r2 →
-  res_dist' n r1 r2.
-Instance res_dist : Dist (res Λ A M) := res_dist'.
-Global Instance Res_ne n :
-  Proper (dist n ==> dist n ==> dist n ==> dist n) (@Res Λ A M).
-Proof. done. Qed.
-Global Instance Res_proper : Proper ((≡) ==> (≡) ==> (≡) ==> (≡)) (@Res Λ A M).
-Proof. done. Qed.
-Global Instance wld_ne n : Proper (dist n ==> dist n) (@wld Λ A M).
-Proof. by destruct 1. Qed.
-Global Instance wld_proper : Proper ((≡) ==> (≡)) (@wld Λ A M).
-Proof. by destruct 1. Qed.
-Global Instance pst_ne n : Proper (dist n ==> dist n) (@pst Λ A M).
-Proof. by destruct 1. Qed.
-Global Instance pst_ne' n : Proper (dist n ==> (≡)) (@pst Λ A M).
-Proof. destruct 1; apply: timeless; apply dist_le with n; auto with lia. Qed.
-Global Instance pst_proper : Proper ((≡) ==> (=)) (@pst Λ A M).
-Proof. by destruct 1; unfold_leibniz. Qed.
-Global Instance gst_ne n : Proper (dist n ==> dist n) (@gst Λ A M).
-Proof. by destruct 1. Qed.
-Global Instance gst_proper : Proper ((≡) ==> (≡)) (@gst Λ A M).
-Proof. by destruct 1. Qed.
-Instance res_compl : Compl (res Λ A M) := λ c,
-  Res (compl (chain_map wld c))
-      (compl (chain_map pst c)) (compl (chain_map gst c)).
-Definition res_cofe_mixin : CofeMixin (res Λ A M).
-Proof.
-  split.
-  - intros w1 w2; split.
-    + by destruct 1; constructor; apply equiv_dist.
-    + by intros Hw; constructor; apply equiv_dist=>n; destruct (Hw n).
-  - intros n; split.
-    + done.
-    + by destruct 1; constructor.
-    + do 2 destruct 1; constructor; etrans; eauto.
-  - by destruct 1; constructor; apply dist_S.
-  - intros n c; constructor.
-    + apply (conv_compl n (chain_map wld c)).
-    + apply (conv_compl n (chain_map pst c)).
-    + apply (conv_compl n (chain_map gst c)).
-Qed.
-Canonical Structure resC : cofeT := CofeT (res Λ A M) res_cofe_mixin.
-Global Instance res_timeless r :
-  Timeless (wld r) → Timeless (gst r) → Timeless r.
-Proof. destruct 3; constructor; by apply (timeless _). Qed.
-
-Instance res_op : Op (res Λ A M) := λ r1 r2,
-  Res (wld r1 â‹… wld r2) (pst r1 â‹… pst r2) (gst r1 â‹… gst r2).
-Instance res_pcore : PCore (res Λ A M) := λ r,
-  Some $ Res (core (wld r)) (core (pst r)) (core (gst r)).
-Instance res_valid : Valid (res Λ A M) := λ r, ✓ wld r ∧ ✓ pst r ∧ ✓ gst r.
-Instance res_validN : ValidN (res Λ A M) := λ n r,
-  ✓{n} wld r ∧ ✓{n} pst r ∧ ✓{n} gst r.
-
-Lemma res_included (r1 r2 : res Λ A M) :
-  r1 ≼ r2 ↔ wld r1 ≼ wld r2 ∧ pst r1 ≼ pst r2 ∧ gst r1 ≼ gst r2.
-Proof.
-  split; [|by intros ([w ?]&[σ ?]&[m ?]); exists (Res w σ m)].
-  intros [r Hr]; split_and?;
-    [exists (wld r)|exists (pst r)|exists (gst r)]; apply Hr.
-Qed.
-Lemma res_includedN (r1 r2 : res Λ A M) n :
-  r1 ≼{n} r2 ↔ wld r1 ≼{n} wld r2 ∧ pst r1 ≼{n} pst r2 ∧ gst r1 ≼{n} gst r2.
-Proof.
-  split; [|by intros ([w ?]&[σ ?]&[m ?]); exists (Res w σ m)].
-  intros [r Hr]; split_and?;
-    [exists (wld r)|exists (pst r)|exists (gst r)]; apply Hr.
-Qed.
-Definition res_cmra_mixin : CMRAMixin (res Λ A M).
-Proof.
-  apply cmra_total_mixin.
-  - eauto.
-  - by intros n x [???] ? [???]; constructor; cofe_subst.
-  - by intros n [???] ? [???]; constructor; cofe_subst.
-  - by intros n [???] ? [???] (?&?&?); split_and!; cofe_subst.
-  - intros r; split.
-    + intros (?&?&?) n; split_and!; by apply cmra_valid_validN.
-    + intros Hr; split_and!; apply cmra_valid_validN=> n; apply Hr.
-  - by intros n ? (?&?&?); split_and!; apply cmra_validN_S.
-  - by intros ???; constructor; rewrite /= assoc.
-  - by intros ??; constructor; rewrite /= comm.
-  - by intros ?; constructor; rewrite /= cmra_core_l.
-  - by intros ?; constructor; rewrite /= cmra_core_idemp.
-  - intros r1 r2; rewrite !res_included.
-    by intros (?&?&?); split_and!; apply cmra_core_mono.
-  - intros n r1 r2 (?&?&?);
-      split_and!; simpl in *; eapply cmra_validN_op_l; eauto.
-  - intros n r r1 r2 (?&?&?) [???]; simpl in *.
-    destruct (cmra_extend n (wld r) (wld r1) (wld r2)) as ([w w']&?&?&?),
-      (cmra_extend n (pst r) (pst r1) (pst r2)) as ([σ σ']&?&?&?),
-      (cmra_extend n (gst r) (gst r1) (gst r2)) as ([m m']&?&?&?); auto.
-    by exists (Res w σ m, Res w' σ' m').
-Qed.
-Canonical Structure resR := CMRAT (res Λ A M) res_cofe_mixin res_cmra_mixin.
-
-Instance res_empty : Empty (res Λ A M) := Res ∅ ∅ ∅.
-Definition res_ucmra_mixin : UCMRAMixin (res Λ A M).
-Proof.
-  split.
-  - split_and!; apply ucmra_unit_valid.
-  - by split; rewrite /= left_id.
-  - apply _.
-  - do 2 constructor; simpl; apply (persistent_core _).
-Qed.
-Canonical Structure resUR :=
-  UCMRAT (res Λ A M) res_cofe_mixin res_cmra_mixin res_ucmra_mixin.
-
-Definition update_pst (σ : state Λ) (r : res Λ A M) : res Λ A M :=
-  Res (wld r) (Excl' σ) (gst r).
-Definition update_gst (m : M) (r : res Λ A M) : res Λ A M :=
-  Res (wld r) (pst r) m.
-
-Lemma wld_validN n r : ✓{n} r → ✓{n} wld r.
-Proof. by intros (?&?&?). Qed.
-Lemma gst_validN n r : ✓{n} r → ✓{n} gst r.
-Proof. by intros (?&?&?). Qed.
-Lemma Res_op w1 w2 σ1 σ2 m1 m2 :
-  Res w1 σ1 m1 ⋅ Res w2 σ2 m2 = Res (w1 ⋅ w2) (σ1 ⋅ σ2) (m1 ⋅ m2).
-Proof. done. Qed.
-Lemma Res_core w σ m : core (Res w σ m) = Res (core w) (core σ) (core m).
-Proof. done. Qed.
-Lemma lookup_wld_op_l n r1 r2 i P :
-  ✓{n} (r1⋅r2) → wld r1 !! i ≡{n}≡ Some P → (wld r1 ⋅ wld r2) !! i ≡{n}≡ Some P.
-Proof.
-  move=>/wld_validN /(_ i) Hval Hi1P; move: Hi1P Hval; rewrite lookup_op.
-  destruct (wld r2 !! i) as [P'|] eqn:Hi; rewrite !Hi ?right_id // =>-> ?.
-  by constructor; rewrite (agree_op_inv _ P P') // agree_idemp.
-Qed.
-Lemma lookup_wld_op_r n r1 r2 i P :
-  ✓{n} (r1⋅r2) → wld r2 !! i ≡{n}≡ Some P → (wld r1 ⋅ wld r2) !! i ≡{n}≡ Some P.
-Proof. rewrite (comm _ r1) (comm _ (wld r1)); apply lookup_wld_op_l. Qed.
-Global Instance Res_timeless eσ m : Timeless m → Timeless (@Res Λ A M ∅ eσ m).
-Proof. by intros ? ? [???]; constructor; apply (timeless _). Qed.
-Global Instance Res_persistent w m: Persistent m → Persistent (@Res Λ A M w ∅ m).
-Proof. do 2 constructor; apply (persistent_core _). Qed.
-
-(** Internalized properties *)
-Lemma res_equivI {M'} r1 r2 :
-  r1 ≡ r2 ⊣⊢ (wld r1 ≡ wld r2 ∧ pst r1 ≡ pst r2 ∧ gst r1 ≡ gst r2 : uPred M').
-Proof.
-  uPred.unseal. do 2 split. by destruct 1. by intros (?&?&?); by constructor.
-Qed.
-Lemma res_validI {M'} r : ✓ r ⊣⊢ (✓ wld r ∧ ✓ pst r ∧ ✓ gst r : uPred M').
-Proof. by uPred.unseal. Qed.
-End res.
-
-Arguments resC : clear implicits.
-Arguments resR : clear implicits.
-Arguments resUR : clear implicits.
-
-(* Functor *)
-Definition res_map {Λ} {A A' : cofeT} {M M' : cmraT}
-    (f : A → A') (g : M → M') (r : res Λ A M) : res Λ A' M' :=
-  Res (agree_map f <$> wld r) (pst r) (g $ gst r).
-Instance res_map_ne {Λ} {A A': cofeT} {M M' : ucmraT} (f : A → A') (g : M → M'):
-  (∀ n, Proper (dist n ==> dist n) f) → (∀ n, Proper (dist n ==> dist n) g) →
-  ∀ n, Proper (dist n ==> dist n) (@res_map Λ _ _ _ _ f g).
-Proof. intros Hf n [] ? [???]; constructor; by cofe_subst. Qed.
-Lemma res_map_id {Λ A} {M : ucmraT} (r : res Λ A M) : res_map id id r ≡ r.
-Proof.
-  constructor; rewrite /res_map /=; f_equal.
-  rewrite -{2}(map_fmap_id (wld r)). apply map_fmap_setoid_ext=> i y ? /=.
-  by rewrite -{2}(agree_map_id y).
-Qed.
-Lemma res_map_compose {Λ} {A1 A2 A3 : cofeT} {M1 M2 M3 : ucmraT}
-   (f : A1 → A2) (f' : A2 → A3) (g : M1 → M2) (g' : M2 → M3) (r : res Λ A1 M1) :
-  res_map (f' ∘ f) (g' ∘ g) r ≡ res_map f' g' (res_map f g r).
-Proof.
-  constructor; rewrite /res_map /=; f_equal.
-  rewrite -map_fmap_compose; apply map_fmap_setoid_ext=> i y _ /=.
-  by rewrite -agree_map_compose.
-Qed.
-Lemma res_map_ext {Λ} {A A' : cofeT} {M M' : ucmraT}
-    (f f' : A → A') (g g' : M → M') (r : res Λ A M) :
-  (∀ x, f x ≡ f' x) → (∀ m, g m ≡ g' m) → res_map f g r ≡ res_map f' g' r.
-Proof.
-  intros Hf Hg; split; simpl; auto.
-  by apply map_fmap_setoid_ext=>i x ?; apply agree_map_ext.
-Qed.
-Instance res_map_cmra_monotone {Λ}
-    {A A' : cofeT} {M M': ucmraT} (f: A → A') (g: M → M') :
-  (∀ n, Proper (dist n ==> dist n) f) → CMRAMonotone g →
-  CMRAMonotone (@res_map Λ _ _ _ _ f g).
-Proof.
-  split; first apply _.
-  - intros n r (?&?&?); split_and!; simpl; by try apply: validN_preserving.
-  - by intros r1 r2; rewrite !res_included;
-      intros (?&?&?); split_and!; simpl; try apply: cmra_monotone.
-Qed.
-Definition resC_map {Λ} {A A' : cofeT} {M M' : ucmraT}
-    (f : A -n> A') (g : M -n> M') : resC Λ A M -n> resC Λ A' M' :=
-  CofeMor (res_map f g : resC Λ A M → resC Λ A' M').
-Instance resC_map_ne {Λ A A' M M'} n :
-  Proper (dist n ==> dist n ==> dist n) (@resC_map Λ A A' M M').
-Proof.
-  intros f g Hfg r; split; simpl; auto.
-  by apply (gmapC_map_ne _ (agreeC_map f) (agreeC_map g)), agreeC_map_ne.
-Qed.
-
-Program Definition resURF (Λ : language)
-    (F1 : cFunctor) (F2 : urFunctor) : urFunctor := {|
-  urFunctor_car A B := resUR Λ (cFunctor_car F1 A B) (urFunctor_car F2 A B);
-  urFunctor_map A1 A2 B1 B2 fg :=resC_map (cFunctor_map F1 fg) (urFunctor_map F2 fg)
-|}.
-Next Obligation.
-  intros Λ F1 F2 A1 A2 B1 B2 n f g Hfg; apply resC_map_ne.
-  - by apply cFunctor_ne.
-  - by apply urFunctor_ne.
-Qed.
-Next Obligation.
-  intros Λ F Σ A B x. rewrite /= -{2}(res_map_id x).
-  apply res_map_ext=>y. apply cFunctor_id. apply urFunctor_id.
-Qed.
-Next Obligation.
-  intros Λ F Σ A1 A2 A3 B1 B2 B3 f g f' g' x. rewrite /= -res_map_compose.
-  apply res_map_ext=>y. apply cFunctor_compose. apply urFunctor_compose.
-Qed.
-
-Instance resRF_contractive Λ F1 F2 :
-  cFunctorContractive F1 → urFunctorContractive F2 →
-  urFunctorContractive (resURF Λ F1 F2).
-Proof.
-  intros ?? A1 A2 B1 B2 n f g Hfg; apply resC_map_ne.
-  - by apply cFunctor_contractive.
-  - by apply urFunctor_contractive.
-Qed.
diff --git a/program_logic/saved_prop.v b/program_logic/saved_prop.v
index b369d833282c957d2190011fef908d0cd6c6d0a3..c27e0c886281c94f6507b7c09bad80c860de15d7 100644
--- a/program_logic/saved_prop.v
+++ b/program_logic/saved_prop.v
@@ -1,42 +1,42 @@
-From iris.algebra Require Export agree.
 From iris.program_logic Require Export ghost_ownership.
+From iris.algebra Require Import agree.
+From iris.prelude Require Import gmap.
 Import uPred.
 
-Class savedPropG (Λ : language) (Σ : gFunctors) (F : cFunctor) :=
-  saved_prop_inG :> inG Λ Σ (agreeR (laterC (F (iPrePropG Λ Σ)))).
+Class savedPropG (Σ : gFunctors) (F : cFunctor) :=
+  saved_prop_inG :> inG Σ (agreeR (laterC (F (iPreProp Σ)))).
 Definition savedPropGF (F : cFunctor) : gFunctor :=
   GFunctor (agreeRF (â–¶ F)).
-Instance inGF_savedPropG  `{inGF Λ Σ (savedPropGF F)} : savedPropG Λ Σ F.
+Instance inGF_savedPropG  `{inGF Σ (savedPropGF F)} : savedPropG Σ F.
 Proof. apply: inGF_inG. Qed.
 
-Definition saved_prop_own `{savedPropG Λ Σ F}
-    (γ : gname) (x : F (iPropG Λ Σ)) : iPropG Λ Σ :=
+Definition saved_prop_own `{savedPropG Σ F}
+    (γ : gname) (x : F (iProp Σ)) : iProp Σ :=
   own γ (to_agree $ Next (cFunctor_map F (iProp_fold, iProp_unfold) x)).
 Typeclasses Opaque saved_prop_own.
-Instance: Params (@saved_prop_own) 4.
+Instance: Params (@saved_prop_own) 3.
 
 Section saved_prop.
-  Context `{savedPropG Λ Σ F}.
-  Implicit Types x y : F (iPropG Λ Σ).
+  Context `{savedPropG Σ F}.
+  Implicit Types x y : F (iProp Σ).
   Implicit Types γ : gname.
 
   Global Instance saved_prop_persistent γ x : PersistentP (saved_prop_own γ x).
   Proof. rewrite /saved_prop_own; apply _. Qed.
 
-  Lemma saved_prop_alloc_strong E x (G : gset gname) :
-    True ={E}=> ∃ γ, ■ (γ ∉ G) ∧ saved_prop_own γ x.
+  Lemma saved_prop_alloc_strong x (G : gset gname) :
+    True =r=> ∃ γ, ■ (γ ∉ G) ∧ saved_prop_own γ x.
   Proof. by apply own_alloc_strong. Qed.
 
-  Lemma saved_prop_alloc E x : True ={E}=> ∃ γ, saved_prop_own γ x.
+  Lemma saved_prop_alloc x : True =r=> ∃ γ, saved_prop_own γ x.
   Proof. by apply own_alloc. Qed.
 
   Lemma saved_prop_agree γ x y :
     saved_prop_own γ x ★ saved_prop_own γ y ⊢ ▷ (x ≡ y).
   Proof.
-    rewrite -own_op own_valid agree_validI agree_equivI later_equivI.
+    rewrite -own_op own_valid agree_validI agree_equivI later_equivI /=.
     set (G1 := cFunctor_map F (iProp_fold, iProp_unfold)).
-    set (G2 := cFunctor_map F (@iProp_unfold Λ (globalF Σ),
-                               @iProp_fold Λ (globalF Σ))).
+    set (G2 := cFunctor_map F (@iProp_unfold Σ, @iProp_fold Σ)).
     assert (∀ z, G2 (G1 z) ≡ z) as help.
     { intros z. rewrite /G1 /G2 -cFunctor_compose -{2}[z]cFunctor_id.
       apply (ne_proper (cFunctor_map F)); split=>?; apply iProp_fold_unfold. }
diff --git a/program_logic/sts.v b/program_logic/sts.v
index a8180d0e93bc936251dfef5b50a3158fc78a2af6..4f599a8b1c1edd1a4922b0512d2dda818e2af1f7 100644
--- a/program_logic/sts.v
+++ b/program_logic/sts.v
@@ -1,30 +1,31 @@
 From iris.algebra Require Export sts upred_tactics.
-From iris.program_logic Require Export invariants ghost_ownership.
-From iris.proofmode Require Import invariants ghost_ownership.
+From iris.program_logic Require Export invariants.
+From iris.proofmode Require Import invariants.
 Import uPred.
 
 (** The CMRA we need. *)
-Class stsG Λ Σ (sts : stsT) := StsG {
-  sts_inG :> inG Λ Σ (stsR sts);
+Class stsG Σ (sts : stsT) := StsG {
+  sts_inG :> inG Σ (stsR sts);
   sts_inhabited :> Inhabited (sts.state sts);
 }.
 Coercion sts_inG : stsG >-> inG.
 (** The Functor we need. *)
 Definition stsGF (sts : stsT) : gFunctor := GFunctor (constRF (stsR sts)).
 (* Show and register that they match. *)
-Instance inGF_stsG sts `{inGF Λ Σ (stsGF sts)}
-  `{Inhabited (sts.state sts)} : stsG Λ Σ sts.
+Instance inGF_stsG sts `{inGF Λ (stsGF sts)}
+  `{Inhabited (sts.state sts)} : stsG Λ sts.
 Proof. split; try apply _. apply: inGF_inG. Qed.
 
 Section definitions.
-  Context `{i : stsG Λ Σ sts} (γ : gname).
-  Definition sts_ownS (S : sts.states sts) (T : sts.tokens sts) : iPropG Λ Σ:=
+  Context `{irisG Λ Σ, stsG Σ sts} (γ : gname).
+
+  Definition sts_ownS (S : sts.states sts) (T : sts.tokens sts) : iProp Σ :=
     own γ (sts_frag S T).
-  Definition sts_own (s : sts.state sts) (T : sts.tokens sts) : iPropG Λ Σ :=
+  Definition sts_own (s : sts.state sts) (T : sts.tokens sts) : iProp Σ :=
     own γ (sts_frag_up s T).
-  Definition sts_inv (φ : sts.state sts → iPropG Λ Σ) : iPropG Λ Σ :=
+  Definition sts_inv (φ : sts.state sts → iProp Σ) : iProp Σ :=
     (∃ s, own γ (sts_auth s ∅) ★ φ s)%I.
-  Definition sts_ctx (N : namespace) (φ: sts.state sts → iPropG Λ Σ) : iPropG Λ Σ :=
+  Definition sts_ctx (N : namespace) (φ: sts.state sts → iProp Σ) : iProp Σ :=
     inv N (sts_inv φ).
 
   Global Instance sts_inv_ne n :
@@ -54,23 +55,23 @@ Instance: Params (@sts_own) 6.
 Instance: Params (@sts_ctx) 6.
 
 Section sts.
-  Context `{stsG Λ Σ sts} (φ : sts.state sts → iPropG Λ Σ).
+  Context `{irisG Λ Σ, stsG Σ sts} (φ : sts.state sts → iProp Σ).
   Implicit Types N : namespace.
-  Implicit Types P Q R : iPropG Λ Σ.
+  Implicit Types P Q R : iProp Σ.
   Implicit Types γ : gname.
   Implicit Types S : sts.states sts.
   Implicit Types T : sts.tokens sts.
 
   (* The same rule as implication does *not* hold, as could be shown using
      sts_frag_included. *)
-  Lemma sts_ownS_weaken E γ S1 S2 T1 T2 :
+  Lemma sts_ownS_weaken γ S1 S2 T1 T2 :
     T2 ⊆ T1 → S1 ⊆ S2 → sts.closed S2 T2 →
-    sts_ownS γ S1 T1 ={E}=> sts_ownS γ S2 T2.
+    sts_ownS γ S1 T1 =r=> sts_ownS γ S2 T2.
   Proof. intros ???. by apply own_update, sts_update_frag. Qed.
 
-  Lemma sts_own_weaken E γ s S T1 T2 :
+  Lemma sts_own_weaken γ s S T1 T2 :
     T2 ⊆ T1 → s ∈ S → sts.closed S T2 →
-    sts_own γ s T1 ={E}=> sts_ownS γ S T2.
+    sts_own γ s T1 =r=> sts_ownS γ S T2.
   Proof. intros ???. by apply own_update, sts_update_frag_up. Qed.
 
   Lemma sts_ownS_op γ S1 S2 T1 T2 :
@@ -79,49 +80,39 @@ Section sts.
   Proof. intros. by rewrite /sts_ownS -own_op sts_op_frag. Qed.
 
   Lemma sts_alloc E N s :
-    nclose N ⊆ E →
     ▷ φ s ={E}=> ∃ γ, sts_ctx γ N φ ∧ sts_own γ s (⊤ ∖ sts.tok s).
   Proof.
-    iIntros (?) "Hφ". rewrite /sts_ctx /sts_own.
-    iPvs (own_alloc (sts_auth s (⊤ ∖ sts.tok s))) as (γ) "Hγ".
+    iIntros "Hφ". rewrite /sts_ctx /sts_own.
+    iVs (own_alloc (sts_auth s (⊤ ∖ sts.tok s))) as (γ) "Hγ".
     { apply sts_auth_valid; set_solver. }
     iExists γ; iRevert "Hγ"; rewrite -sts_op_auth_frag_up; iIntros "[Hγ $]".
-    iPvs (inv_alloc N _ (sts_inv γ φ) with "[Hφ Hγ]") as "#?"; auto.
+    iVs (inv_alloc N _ (sts_inv γ φ) with "[Hφ Hγ]") as "#?"; auto.
     iNext. iExists s. by iFrame.
   Qed.
 
-  Context {V} (fsa : FSA Λ (globalF Σ) V) `{!FrameShiftAssertion fsaV fsa}.
-
-  Lemma sts_fsaS E N (Ψ : V → iPropG Λ Σ) γ S T :
-    fsaV → nclose N ⊆ E →
-    sts_ctx γ N φ ★ sts_ownS γ S T ★ (∀ s,
-      ■ (s ∈ S) ★ ▷ φ s -★
-      fsa (E ∖ nclose N) (λ x, ∃ s' T',
-        ■ sts.steps (s, T) (s', T') ★ ▷ φ s' ★ (sts_own γ s' T' -★ Ψ x)))
-    ⊢ fsa E Ψ.
+  Lemma sts_openS E N γ S T :
+    nclose N ⊆ E →
+    sts_ctx γ N φ ★ sts_ownS γ S T ={E,E∖N}=> ∃ s,
+      ■ (s ∈ S) ★ ▷ φ s ★ ∀ s' T',
+      ■ sts.steps (s, T) (s', T') ★ ▷ φ s' ={E∖N,E}=★ sts_own γ s' T'.
   Proof.
-    iIntros (??) "(#? & Hγf & HΨ)". rewrite /sts_ctx /sts_ownS /sts_inv /sts_own.
-    iInv N as (s) "[Hγ Hφ]"; iTimeless "Hγ".
+    iIntros (?) "[#? Hγf]". rewrite /sts_ctx /sts_ownS /sts_inv /sts_own.
+    iInv N as (s) "[Hγ Hφ]" "Hclose". iTimeless "Hγ".
     iCombine "Hγ" "Hγf" as "Hγ"; iDestruct (own_valid with "#Hγ") as %Hvalid.
     assert (s ∈ S) by eauto using sts_auth_frag_valid_inv.
     assert (✓ sts_frag S T) as [??] by eauto using cmra_valid_op_r.
-    iRevert "Hγ"; rewrite sts_op_auth_frag //; iIntros "Hγ".
-    iApply pvs_fsa_fsa; iApply fsa_wand_r; iSplitL "HΨ Hφ".
-    { iApply "HΨ"; by iSplit. }
-    iIntros (a); iDestruct 1 as (s' T') "(% & Hφ & HΨ)".
-    iPvs (own_update with "Hγ") as "Hγ"; first eauto using sts_update_auth.
-    iRevert "Hγ"; rewrite -sts_op_auth_frag_up; iIntros "[Hγ Hγf]".
-    iPvsIntro; iSplitL "Hφ Hγ"; last by iApply "HΨ".
-    iNext; iExists s'; by iFrame.
+    rewrite sts_op_auth_frag //.
+    iVsIntro; iExists s; iSplit; [done|]; iFrame "Hφ".
+    iIntros (s' T') "[% Hφ]".
+    iVs (own_update with "Hγ") as "Hγ"; first eauto using sts_update_auth.
+    iRevert "Hγ"; rewrite -sts_op_auth_frag_up; iIntros "[Hγ $]".
+    iApply "Hclose". iNext; iExists s'; by iFrame.
   Qed.
 
-  Lemma sts_fsa E N (Ψ : V → iPropG Λ Σ) γ s0 T :
-    fsaV → nclose N ⊆ E →
-    sts_ctx γ N φ ★ sts_own γ s0 T ★ (∀ s,
-      ■ (s ∈ sts.up s0 T) ★ ▷ φ s -★
-      fsa (E ∖ nclose N) (λ x, ∃ s' T',
-        ■ (sts.steps (s, T) (s', T')) ★ ▷ φ s' ★
-        (sts_own γ s' T' -★ Ψ x)))
-    ⊢ fsa E Ψ.
-  Proof. by apply sts_fsaS. Qed.
+  Lemma sts_open E N γ s0 T :
+    nclose N ⊆ E →
+    sts_ctx γ N φ ★ sts_own γ s0 T ={E,E∖N}=> ∃ s,
+      ■ (s ∈ sts.up s0 T) ★ ▷ φ s ★ ∀ s' T',
+      ■ (sts.steps (s, T) (s', T')) ★ ▷ φ s' ={E∖N,E}=★ sts_own γ s' T'.
+  Proof. by apply sts_openS. Qed.
 End sts.
diff --git a/program_logic/viewshifts.v b/program_logic/viewshifts.v
deleted file mode 100644
index b131eed655aada65465a81287a7533e48ce03cd7..0000000000000000000000000000000000000000
--- a/program_logic/viewshifts.v
+++ /dev/null
@@ -1,80 +0,0 @@
-From iris.program_logic Require Import ownership.
-From iris.program_logic Require Export pviewshifts invariants ghost_ownership.
-From iris.proofmode Require Import pviewshifts invariants.
-Import uPred.
-
-Definition vs {Λ Σ} (E1 E2 : coPset) (P Q : iProp Λ Σ) : iProp Λ Σ :=
-  (□ (P → |={E1,E2}=> Q))%I.
-Arguments vs {_ _} _ _ _%I _%I.
-Instance: Params (@vs) 4.
-Notation "P ={ E1 , E2 }=> Q" := (vs E1 E2 P%I Q%I)
-  (at level 99, E1,E2 at level 50, Q at level 200,
-   format "P  ={ E1 , E2 }=>  Q") : uPred_scope.
-Notation "P ={ E }=> Q" := (P ={E,E}=> Q)%I
-  (at level 99, E at level 50, Q at level 200,
-   format "P  ={ E }=>  Q") : uPred_scope.
-
-Section vs.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P Q R : iProp Λ Σ.
-Implicit Types N : namespace.
-
-Global Instance vs_ne E1 E2 n :
-  Proper (dist n ==> dist n ==> dist n) (@vs Λ Σ E1 E2).
-Proof. solve_proper. Qed.
-
-Global Instance vs_proper E1 E2 : Proper ((≡) ==> (≡) ==> (≡)) (@vs Λ Σ E1 E2).
-Proof. apply ne_proper_2, _. Qed.
-
-Lemma vs_mono E1 E2 P P' Q Q' :
-  (P ⊢ P') → (Q' ⊢ Q) → (P' ={E1,E2}=> Q') ⊢ P ={E1,E2}=> Q.
-Proof. by intros HP HQ; rewrite /vs -HP HQ. Qed.
-
-Global Instance vs_mono' E1 E2 :
-  Proper (flip (⊢) ==> (⊢) ==> (⊢)) (@vs Λ Σ E1 E2).
-Proof. solve_proper. Qed.
-
-Lemma vs_false_elim E1 E2 P : False ={E1,E2}=> P.
-Proof. iIntros "[]". Qed.
-Lemma vs_timeless E P : TimelessP P → ▷ P ={E}=> P.
-Proof. by apply pvs_timeless. Qed.
-
-Lemma vs_transitive E1 E2 E3 P Q R :
-  E2 ⊆ E1 ∪ E3 → (P ={E1,E2}=> Q) ∧ (Q ={E2,E3}=> R) ⊢ P ={E1,E3}=> R.
-Proof.
-  iIntros (?) "#[HvsP HvsQ] ! HP".
-  iPvs ("HvsP" with "HP") as "HQ"; first done. by iApply "HvsQ".
-Qed.
-
-Lemma vs_transitive' E P Q R : (P ={E}=> Q) ∧ (Q ={E}=> R) ⊢ (P ={E}=> R).
-Proof. apply vs_transitive; set_solver. Qed.
-Lemma vs_reflexive E P : P ={E}=> P.
-Proof. by iIntros "HP". Qed.
-
-Lemma vs_impl E P Q : □ (P → Q) ⊢ P ={E}=> Q.
-Proof. iIntros "#HPQ ! HP". by iApply "HPQ". Qed.
-
-Lemma vs_frame_l E1 E2 P Q R : (P ={E1,E2}=> Q) ⊢ R ★ P ={E1,E2}=> R ★ Q.
-Proof. iIntros "#Hvs ! [$ HP]". by iApply "Hvs". Qed.
-
-Lemma vs_frame_r E1 E2 P Q R : (P ={E1,E2}=> Q) ⊢ P ★ R ={E1,E2}=> Q ★ R.
-Proof. iIntros "#Hvs ! [HP $]". by iApply "Hvs". Qed.
-
-Lemma vs_mask_frame E1 E2 Ef P Q :
-  Ef ⊥ E1 ∪ E2 → (P ={E1,E2}=> Q) ⊢ P ={E1 ∪ Ef,E2 ∪ Ef}=> Q.
-Proof.
-  iIntros (?) "#Hvs ! HP". iApply pvs_mask_frame; auto. by iApply "Hvs".
-Qed.
-
-Lemma vs_mask_frame' E Ef P Q : Ef ⊥ E → (P ={E}=> Q) ⊢ P ={E ∪ Ef}=> Q.
-Proof. intros; apply vs_mask_frame; set_solver. Qed.
-
-Lemma vs_inv N E P Q R :
-  nclose N ⊆ E → inv N R ★ (▷ R ★ P ={E ∖ nclose N}=> ▷ R ★ Q) ⊢ P ={E}=> Q.
-Proof.
-  iIntros (?) "#[? Hvs] ! HP". iInv N as "HR". iApply "Hvs". by iSplitL "HR".
-Qed.
-
-Lemma vs_alloc N P : â–· P ={N}=> inv N P.
-Proof. iIntros "HP". by iApply inv_alloc. Qed.
-End vs.
diff --git a/program_logic/weakestpre.v b/program_logic/weakestpre.v
index ab2d859b938a2cc26388a2422e7a5036c588c36a..9021ea4510fe0945f7de1aeec8dbb0df99caecec 100644
--- a/program_logic/weakestpre.v
+++ b/program_logic/weakestpre.v
@@ -1,58 +1,39 @@
 From iris.program_logic Require Export pviewshifts.
-From iris.program_logic Require Import wsat.
-Local Hint Extern 10 (_ ≤ _) => omega.
-Local Hint Extern 100 (_ ⊥ _) => set_solver.
-Local Hint Extern 100 (_ ∉_) => set_solver.
-Local Hint Extern 100 (@subseteq coPset _ _ _) => set_solver.
-Local Hint Extern 10 (✓{_} _) =>
-  repeat match goal with
-  | H : wsat _ _ _ _ |- _ => apply wsat_valid in H; last omega
-  end; solve_validN.
+From iris.program_logic Require Import ownership.
+From iris.prelude Require Export coPset.
+From iris.proofmode Require Import tactics pviewshifts.
+Import uPred.
+
+Definition wp_pre `{irisG Λ Σ}
+    (wp : coPset -c> expr Λ -c> (val Λ -c> iProp Σ) -c> iProp Σ) :
+    coPset -c> expr Λ -c> (val Λ -c> iProp Σ) -c> iProp Σ := λ E e1 Φ, (
+  (* value case *)
+  (∃ v, to_val e1 = Some v ∧ |={E}=> Φ v) ∨
+  (* step case *)
+  (to_val e1 = None ∧ ∀ σ1,
+     ownP_auth σ1 ={E,∅}=★ ■ reducible e1 σ1 ★
+     ▷ ∀ e2 σ2 ef, ■ prim_step e1 σ1 e2 σ2 ef ={∅,E}=★
+       ownP_auth σ2 ★ wp E e2 Φ ★
+       from_option (flip (wp ⊤) (λ _, True)) True ef))%I.
 
-Record wp_go {Λ Σ} (E : coPset) (Φ Φfork : expr Λ → nat → iRes Λ Σ → Prop)
-    (k : nat) (σ1 : state Λ) (rf : iRes Λ Σ) (e1 : expr Λ) := {
-  wf_safe : reducible e1 σ1;
-  wp_step e2 σ2 ef :
-    prim_step e1 σ1 e2 σ2 ef →
-    ∃ r2 r2',
-      wsat k E σ2 (r2 ⋅ r2' ⋅ rf) ∧
-      Φ e2 k r2 ∧
-      ∀ e', ef = Some e' → Φfork e' k r2'
-}.
-CoInductive wp_pre {Λ Σ} (E : coPset)
-     (Φ : val Λ → iProp Λ Σ) : expr Λ → nat → iRes Λ Σ → Prop :=
-  | wp_pre_value n r v : (|={E}=> Φ v)%I n r → wp_pre E Φ (of_val v) n r
-  | wp_pre_step n r1 e1 :
-     to_val e1 = None →
-     (∀ k Ef σ1 rf,
-       0 < k < n → E ⊥ Ef →
-       wsat (S k) (E ∪ Ef) σ1 (r1 ⋅ rf) →
-       wp_go (E ∪ Ef) (wp_pre E Φ)
-                      (wp_pre ⊤ (λ _, True%I)) k σ1 rf e1) →
-     wp_pre E Φ e1 n r1.
-Program Definition wp_def {Λ Σ} (E : coPset) (e : expr Λ)
-  (Φ : val Λ → iProp Λ Σ) : iProp Λ Σ := {| uPred_holds := wp_pre E Φ e |}.
-Next Obligation.
-  intros Λ Σ E e Φ n r1 r2; revert Φ E e r1 r2.
-  induction n as [n IH] using lt_wf_ind; intros Φ E e r1 r1'.
-  destruct 1 as [|n r1 e1 ? Hgo].
-  - constructor; eauto using uPred_mono.
-  - intros [rf' Hr]; constructor; [done|intros k Ef σ1 rf ???].
-    destruct (Hgo k Ef σ1 (rf' ⋅ rf)) as [Hsafe Hstep];
-      rewrite ?assoc -?(dist_le _ _ _ _ Hr); auto; constructor; [done|].
-    intros e2 σ2 ef ?; destruct (Hstep e2 σ2 ef) as (r2&r2'&?&?&?); auto.
-    exists r2, (r2' â‹… rf'); split_and?; eauto 10 using (IH k), cmra_includedN_l.
-    by rewrite -!assoc (assoc _ r2).
+Local Instance wp_pre_contractive `{irisG Λ Σ} : Contractive wp_pre.
+Proof.
+  rewrite /wp_pre=> n wp wp' Hwp E e1 Φ.
+  apply or_ne, and_ne, forall_ne; auto=> σ1; apply wand_ne; auto.
+  apply pvs_ne, sep_ne, later_contractive; auto=> i ?.
+  apply forall_ne=> e2; apply forall_ne=> σ2; apply forall_ne=> ef.
+  apply wand_ne, pvs_ne, sep_ne, sep_ne; auto; first by apply Hwp.
+  destruct ef; first apply Hwp; auto.
 Qed.
-Next Obligation. destruct 1; constructor; eauto using uPred_closed. Qed.
 
-(* Perform sealing. *)
+Definition wp_def `{irisG Λ Σ} :
+  coPset → expr Λ → (val Λ → iProp Σ) → iProp Σ := fixpoint wp_pre.
 Definition wp_aux : { x | x = @wp_def }. by eexists. Qed.
 Definition wp := proj1_sig wp_aux.
 Definition wp_eq : @wp = @wp_def := proj2_sig wp_aux.
 
-Arguments wp {_ _} _ _ _.
-Instance: Params (@wp) 4.
+Arguments wp {_ _ _} _ _ _.
+Instance: Params (@wp) 5.
 
 Notation "'WP' e @ E {{ Φ } }" := (wp E e Φ)
   (at level 20, e, Φ at level 200,
@@ -68,198 +49,140 @@ Notation "'WP' e {{ v , Q } }" := (wp ⊤ e (λ v, Q))
   (at level 20, e, Q at level 200,
    format "'WP'  e  {{  v ,  Q  } }") : uPred_scope.
 
+Notation wp_fork ef := (from_option (flip (wp ⊤) (λ _, True)) True ef)%I.
+
 Section wp.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P : iProp Λ Σ.
-Implicit Types Φ : val Λ → iProp Λ Σ.
+Context `{irisG Λ Σ}.
+Implicit Types P : iProp Σ.
+Implicit Types Φ : val Λ → iProp Σ.
 Implicit Types v : val Λ.
 Implicit Types e : expr Λ.
 
+Lemma wp_unfold E e Φ : WP e @ E {{ Φ }} ⊣⊢ wp_pre wp E e Φ.
+Proof. rewrite wp_eq. apply (fixpoint_unfold wp_pre). Qed.
+
 Global Instance wp_ne E e n :
-  Proper (pointwise_relation _ (dist n) ==> dist n) (@wp Λ Σ E e).
+  Proper (pointwise_relation _ (dist n) ==> dist n) (@wp Λ Σ _ E e).
 Proof.
-  cut (∀ Φ Ψ, (∀ v, Φ v ≡{n}≡ Ψ v) →
-    ∀ n' r, n' ≤ n → ✓{n'} r → wp E e Φ n' r → wp E e Ψ n' r).
-  { rewrite wp_eq. intros help Φ Ψ HΦΨ. by do 2 split; apply help. }
-  rewrite wp_eq. intros Φ Ψ HΦΨ n' r; revert e r.
-  induction n' as [n' IH] using lt_wf_ind=> e r.
-  destruct 3 as [n' r v HpvsQ|n' r e1 ? Hgo].
-  { constructor. by eapply pvs_ne, HpvsQ; eauto. }
-  constructor; [done|]=> k Ef σ1 rf ???.
-  destruct (Hgo k Ef σ1 rf) as [Hsafe Hstep]; auto.
-  split; [done|intros e2 σ2 ef ?].
-  destruct (Hstep e2 σ2 ef) as (r2&r2'&?&?&?); auto.
-  exists r2, r2'; split_and?; [|eapply IH|]; eauto.
+  revert e. induction (lt_wf n) as [n _ IH]=> e Φ Ψ HΦ.
+  rewrite !wp_unfold /wp_pre. apply or_ne, and_ne; auto; first solve_proper.
+  apply forall_ne=> σ1.
+  apply wand_ne, pvs_ne, sep_ne, later_contractive; auto=> i ?.
+  apply forall_ne=> e2; apply forall_ne=> σ2; apply forall_ne=> ef.
+  apply wand_ne, pvs_ne, sep_ne, sep_ne; auto.
+  apply IH; [done|]=> v. eapply dist_le; eauto with omega.
 Qed.
 Global Instance wp_proper E e :
-  Proper (pointwise_relation _ (≡) ==> (≡)) (@wp Λ Σ E e).
+  Proper (pointwise_relation _ (≡) ==> (≡)) (@wp Λ Σ _ E e).
 Proof.
   by intros Φ Φ' ?; apply equiv_dist=>n; apply wp_ne=>v; apply equiv_dist.
 Qed.
 
-Lemma wp_mask_frame_mono E1 E2 e Φ Ψ :
-  E1 ⊆ E2 → (∀ v, Φ v ⊢ Ψ v) → WP e @ E1 {{ Φ }} ⊢ WP e @ E2 {{ Ψ }}.
-Proof.
-  rewrite wp_eq. intros HE HΦ; split=> n r.
-  revert e r; induction n as [n IH] using lt_wf_ind=> e r.
-  destruct 2 as [n' r v HpvsQ|n' r e1 ? Hgo].
-  { constructor; eapply pvs_mask_frame_mono, HpvsQ; eauto. }
-  constructor; [done|]=> k Ef σ1 rf ???.
-  assert (E2 ∪ Ef = E1 ∪ (E2 ∖ E1 ∪ Ef)) as HE'.
-  { by rewrite assoc_L -union_difference_L. }
-  destruct (Hgo k ((E2 ∖ E1) ∪ Ef) σ1 rf) as [Hsafe Hstep]; rewrite -?HE'; auto.
-  split; [done|intros e2 σ2 ef ?].
-  destruct (Hstep e2 σ2 ef) as (r2&r2'&?&?&?); auto.
-  exists r2, r2'; split_and?; [rewrite HE'|eapply IH|]; eauto.
-Qed.
-
-Lemma wp_zero E e Φ r : wp_def E e Φ 0 r.
-Proof.
-  case EQ: (to_val e).
-  - rewrite -(of_to_val _ _ EQ). constructor. rewrite pvs_eq.
-    exact: pvs_zero.
-  - constructor; first done. intros ?????. exfalso. omega.
-Qed.
-Lemma wp_value_inv E Φ v n r : wp_def E (of_val v) Φ n r → pvs E E (Φ v) n r.
+Lemma wp_value' E Φ v : Φ v ⊢ WP of_val v @ E {{ Φ }}.
 Proof.
-  by inversion 1 as [|??? He]; [|rewrite ?to_of_val in He]; simplify_eq.
+  iIntros "HΦ". rewrite wp_unfold /wp_pre.
+  iLeft; iExists v; rewrite to_of_val; auto.
 Qed.
-Lemma wp_step_inv E Ef Φ e k n σ r rf :
-  to_val e = None → 0 < k < n → E ⊥ Ef →
-  wp_def E e Φ n r → wsat (S k) (E ∪ Ef) σ (r ⋅ rf) →
-  wp_go (E ∪ Ef) (λ e, wp_def E e Φ) (λ e, wp_def ⊤ e (λ _, True%I)) k σ rf e.
+Lemma wp_value_inv E Φ v : WP of_val v @ E {{ Φ }} ={E}=> Φ v.
 Proof.
-  intros He; destruct 3; [by rewrite ?to_of_val in He|eauto].
+  rewrite wp_unfold /wp_pre to_of_val. iIntros "[H|[% _]]"; [|done].
+  by iDestruct "H" as (v') "[% ?]"; simplify_eq.
 Qed.
 
-Lemma wp_value' E Φ v : Φ v ⊢ WP of_val v @ E {{ Φ }}.
-Proof. rewrite wp_eq. split=> n r; constructor; by apply pvs_intro. Qed.
 Lemma pvs_wp E e Φ : (|={E}=> WP e @ E {{ Φ }}) ⊢ WP e @ E {{ Φ }}.
 Proof.
-  rewrite wp_eq. split=> n r ? Hvs.
-  destruct (to_val e) as [v|] eqn:He; [apply of_to_val in He; subst|].
-  { constructor; eapply pvs_trans', pvs_mono, Hvs; eauto.
-    split=> ???; apply wp_value_inv. }
-  constructor; [done|]=> k Ef σ1 rf ???.
-  rewrite pvs_eq in Hvs. destruct (Hvs (S k) Ef σ1 rf) as (r'&Hwp&?); auto.
-  eapply wp_step_inv with (S k) r'; eauto.
+  rewrite wp_unfold /wp_pre. iIntros "H". destruct (to_val e) as [v|] eqn:?.
+  { iLeft. iExists v; iSplit; first done.
+    by iVs "H" as "[H|[% ?]]"; [iDestruct "H" as (v') "[% ?]"|]; simplify_eq. }
+  iRight; iSplit; [done|]; iIntros (σ1) "Hσ1".
+  iVs "H" as "[H|[% H]]"; last by iApply "H".
+  iDestruct "H" as (v') "[% ?]"; simplify_eq.
 Qed.
 Lemma wp_pvs E e Φ : WP e @ E {{ v, |={E}=> Φ v }} ⊢ WP e @ E {{ Φ }}.
 Proof.
-  rewrite wp_eq. split=> n r; revert e r;
-    induction n as [n IH] using lt_wf_ind=> e r Hr HΦ.
-  destruct (to_val e) as [v|] eqn:He; [apply of_to_val in He; subst|].
-  { constructor; apply pvs_trans', (wp_value_inv _ (pvs E E ∘ Φ)); auto. }
-  constructor; [done|]=> k Ef σ1 rf ???.
-  destruct (wp_step_inv E Ef (pvs E E ∘ Φ) e k n σ1 r rf) as [? Hstep]; auto.
-  split; [done|intros e2 σ2 ef ?].
-  destruct (Hstep e2 σ2 ef) as (r2&r2'&?&Hwp'&?); auto.
-  exists r2, r2'; split_and?; [|apply (IH k)|]; auto.
+  iIntros "H". iLöb (E e Φ) as "IH". rewrite !wp_unfold /wp_pre.
+  iDestruct "H" as "[Hv|[% H]]"; [iLeft|iRight].
+  { iDestruct "Hv" as (v) "[% Hv]". iExists v; iSplit; first done.
+    by iVs "Hv". }
+  iSplit; [done|]; iIntros (σ1) "Hσ1".
+  iVs ("H" $! _ with "Hσ1") as "[$ H]".
+  iVsIntro; iNext; iIntros (e2 σ2 ef Hstep).
+  iVs ("H" $! e2 σ2 ef with "[%]") as "($ & ? & $)"; eauto. by iApply "IH".
 Qed.
+
 Lemma wp_atomic E1 E2 e Φ :
-  E2 ⊆ E1 → atomic e →
+  atomic e →
   (|={E1,E2}=> WP e @ E2 {{ v, |={E2,E1}=> Φ v }}) ⊢ WP e @ E1 {{ Φ }}.
 Proof.
-  rewrite wp_eq pvs_eq. intros ? He; split=> n r ? Hvs.
-  destruct (Some_dec (to_val e)) as [[v <-%of_to_val]|].
-  - eapply wp_pre_value. rewrite pvs_eq.
-    intros k Ef σ rf ???. destruct (Hvs k Ef σ rf) as (r'&Hwp&?); auto.
-    apply wp_value_inv in Hwp. rewrite pvs_eq in Hwp.
-    destruct (Hwp k Ef σ rf) as (r2'&HΦ&?); auto.
-  - apply wp_pre_step. done. intros k Ef σ1 rf ???.
-    destruct (Hvs (S k) Ef σ1 rf) as (r'&Hwp&?); auto.
-    destruct (wp_step_inv E2 Ef (pvs_def E2 E1 ∘ Φ) e k (S k) σ1 r' rf)
-      as [Hsafe Hstep]; auto; [].
-    split; [done|]=> e2 σ2 ef ?.
-    destruct (Hstep e2 σ2 ef) as (r2&r2'&?&Hwp'&?); clear Hsafe Hstep; auto.
-    destruct Hwp' as [k r2 v Hvs'|k r2 e2 Hgo];
-      [|destruct (He σ1 e2 σ2 ef); naive_solver].
-    rewrite -pvs_eq in Hvs'. apply pvs_trans in Hvs';auto. rewrite pvs_eq in Hvs'.
-    destruct (Hvs' k Ef σ2 (r2' ⋅ rf)) as (r3&[]); rewrite ?assoc; auto.
-    exists r3, r2'; split_and?; last done.
-    + by rewrite -assoc.
-    + constructor; apply pvs_intro; auto.
-Qed.
-Lemma wp_frame_r E e Φ R : WP e @ E {{ Φ }} ★ R ⊢ WP e @ E {{ v, Φ v ★ R }}.
-Proof.
-  rewrite wp_eq. uPred.unseal; split; intros n r' Hvalid (r&rR&Hr&Hwp&?).
-  revert Hvalid. rewrite Hr; clear Hr; revert e r Hwp.
-  induction n as [n IH] using lt_wf_ind; intros e r1.
-  destruct 1 as [|n r e ? Hgo]=>?.
-  { constructor. rewrite -uPred_sep_eq; apply pvs_frame_r; auto.
-    uPred.unseal; exists r, rR; eauto. }
-  constructor; [done|]=> k Ef σ1 rf ???.
-  destruct (Hgo k Ef σ1 (rR⋅rf)) as [Hsafe Hstep]; auto.
-  { by rewrite assoc. }
-  split; [done|intros e2 σ2 ef ?].
-  destruct (Hstep e2 σ2 ef) as (r2&r2'&?&?&?); auto.
-  exists (r2 â‹… rR), r2'; split_and?; auto.
-  - by rewrite -(assoc _ r2) (comm _ rR) !assoc -(assoc _ _ rR).
-  - apply IH; eauto using uPred_closed.
+  iIntros (Hatomic) "H". destruct (to_val e) as [v|] eqn:He.
+  { apply of_to_val in He as <-. iApply wp_pvs. iApply wp_value'.
+    iVs "H". by iVs (wp_value_inv with "H"). }
+  setoid_rewrite wp_unfold; rewrite /wp_pre. iRight; iSplit; auto.
+  iIntros (σ1) "Hσ". iVs "H" as "[H|[_ H]]".
+  { iDestruct "H" as (v') "[% ?]"; simplify_eq. }
+  iVs ("H" $! σ1 with "Hσ") as "[$ H]".
+  iVsIntro. iNext. iIntros (e2 σ2 ef Hstep).
+  destruct (Hatomic _ _ _ _ Hstep) as [v <-%of_to_val].
+  iVs ("H" $! _ σ2 ef with "[#]") as "($ & H & $)"; auto.
+  iVs (wp_value_inv with "H") as "H". iVs "H". by iApply wp_value'.
 Qed.
-Lemma wp_frame_step_r E E1 E2 e Φ R :
-  to_val e = None → E ⊥ E1 → E2 ⊆ E1 →
-  WP e @ E {{ Φ }} ★ (|={E1,E2}=> ▷ |={E2,E1}=> R)
-  ⊢ WP e @ E ∪ E1 {{ v, Φ v ★ R }}.
-Proof.
-  rewrite wp_eq pvs_eq=> He ??.
-  uPred.unseal; split; intros n r' Hvalid (r&rR&Hr&Hwp&HR); cofe_subst.
-  constructor; [done|]=> k Ef σ1 rf ?? Hws1.
-  destruct Hwp as [|n r e ? Hgo]; [by rewrite to_of_val in He|].
-  (* "execute" HR *)
-  destruct (HR (S k) (E ∪ Ef) σ1 (r ⋅ rf)) as (s&Hvs&Hws2); auto.
-  { eapply wsat_proper, Hws1; first by set_solver+.
-    by rewrite assoc [rR â‹… _]comm. }
-  clear Hws1 HR.
-  (* Take a step *)
-  destruct (Hgo k (E2 ∪ Ef) σ1 (s ⋅ rf)) as [Hsafe Hstep]; auto.
-  { eapply wsat_proper, Hws2; first by set_solver+.
-    by rewrite !assoc [s â‹… _]comm. }
-  clear Hgo.
-  split; [done|intros e2 σ2 ef ?].
-  destruct (Hstep e2 σ2 ef) as (r2&r2'&Hws3&?&?); auto. clear Hws2.
-  (* Execute 2nd part of the view shift *)
-  destruct (Hvs k (E ∪ Ef) σ2 (r2 ⋅ r2' ⋅ rf)) as (t&HR&Hws4); auto.
-  { eapply wsat_proper, Hws3; first by set_solver+.
-    by rewrite !assoc [_ â‹… s]comm !assoc. }
-  clear Hvs Hws3.
-  (* Execute the rest of e *)
-  exists (r2 â‹… t), r2'; split_and?; auto.
-  - eapply wsat_proper, Hws4; first by set_solver+.
-    by rewrite !assoc [_ â‹… t]comm.
-  - rewrite -uPred_sep_eq. move: wp_frame_r. rewrite wp_eq=>Hframe.
-    apply Hframe; first by auto. uPred.unseal; exists r2, t; split_and?; auto.
-    move: wp_mask_frame_mono. rewrite wp_eq=>Hmask.
-    eapply (Hmask E); by auto.
+
+Lemma wp_strong_mono E1 E2 e Φ Ψ :
+  E1 ⊆ E2 → (∀ v, Φ v -★ Ψ v) ★ WP e @ E1 {{ Φ }} ⊢ WP e @ E2 {{ Ψ }}.
+Proof.
+  iIntros (?) "[HΦ H]". iLöb (e) as "IH". rewrite !wp_unfold /wp_pre.
+  iDestruct "H" as "[Hv|[% H]]"; [iLeft|iRight].
+  { iDestruct "Hv" as (v) "[% Hv]". iExists v; iSplit; first done.
+    iApply (pvs_mask_mono E1 _); first done. by iApply ("HΦ" with "|==>[-]"). }
+  iSplit; [done|]; iIntros (σ1) "Hσ".
+  iApply (pvs_trans _ E1); iApply pvs_intro'; auto. iIntros "Hclose".
+  iVs ("H" $! σ1 with "Hσ") as "[$ H]".
+  iVsIntro. iNext. iIntros (e2 σ2 ef Hstep).
+  iVs ("H" $! _ σ2 ef with "[#]") as "($ & H & $)"; auto.
+  iVs "Hclose" as "_". by iApply ("IH" with "HΦ").
+Qed.
+Lemma wp_frame_step_l E1 E2 e Φ R :
+  to_val e = None → E2 ⊆ E1 →
+  (|={E1,E2}=> ▷ |={E2,E1}=> R) ★ WP e @ E2 {{ Φ }} ⊢ WP e @ E1 {{ v, R ★ Φ v }}.
+Proof.
+  rewrite !wp_unfold /wp_pre. iIntros (??) "[HR [Hv|[_ H]]]".
+  { iDestruct "Hv" as (v) "[% Hv]"; simplify_eq. }
+  iRight; iSplit; [done|]. iIntros (σ1) "Hσ".
+  iVs "HR". iVs ("H" $! _ with "Hσ") as "[$ H]".
+  iVsIntro; iNext; iIntros (e2 σ2 ef Hstep).
+  iVs ("H" $! e2 σ2 ef with "[%]") as "($ & H & $)"; auto.
+  iVs "HR". iVsIntro. iApply (wp_strong_mono E2 _ _ Φ); try iFrame; eauto.
 Qed.
+
 Lemma wp_bind `{LanguageCtx Λ K} E e Φ :
   WP e @ E {{ v, WP K (of_val v) @ E {{ Φ }} }} ⊢ WP K e @ E {{ Φ }}.
 Proof.
-  rewrite wp_eq. split=> n r; revert e r;
-    induction n as [n IH] using lt_wf_ind=> e r ?.
-  destruct 1 as [|n r e ? Hgo].
-  { rewrite -wp_eq. apply pvs_wp; rewrite ?wp_eq; done. }
-  constructor; auto using fill_not_val=> k Ef σ1 rf ???.
-  destruct (Hgo k Ef σ1 rf) as [Hsafe Hstep]; auto.
-  split.
-  { destruct Hsafe as (e2&σ2&ef&?).
-    by exists (K e2), σ2, ef; apply fill_step. }
-  intros e2 σ2 ef ?.
+  iIntros "H". iLöb (E e Φ) as "IH". rewrite wp_unfold /wp_pre.
+  iDestruct "H" as "[Hv|[% H]]".
+  { iDestruct "Hv" as (v) "[Hev Hv]"; iDestruct "Hev" as % <-%of_to_val.
+    by iApply pvs_wp. }
+  rewrite wp_unfold /wp_pre. iRight; iSplit; eauto using fill_not_val.
+  iIntros (σ1) "Hσ". iVs ("H" $! _ with "Hσ") as "[% H]".
+  iVsIntro; iSplit.
+  { iPureIntro. unfold reducible in *. naive_solver eauto using fill_step. }
+  iNext; iIntros (e2 σ2 ef Hstep).
   destruct (fill_step_inv e σ1 e2 σ2 ef) as (e2'&->&?); auto.
-  destruct (Hstep e2' σ2 ef) as (r2&r2'&?&?&?); auto.
-  exists r2, r2'; split_and?; try eapply IH; eauto.
+  iVs ("H" $! e2' σ2 ef with "[%]") as "($ & H & $)"; auto.
+  by iApply "IH".
 Qed.
 
 (** * Derived rules *)
-Import uPred.
 Lemma wp_mono E e Φ Ψ : (∀ v, Φ v ⊢ Ψ v) → WP e @ E {{ Φ }} ⊢ WP e @ E {{ Ψ }}.
-Proof. by apply wp_mask_frame_mono. Qed.
+Proof.
+  iIntros (HΦ) "H"; iApply (wp_strong_mono E E); auto.
+  iFrame. iIntros (v). iApply HΦ.
+Qed.
+Lemma wp_mask_mono E1 E2 e Φ : E1 ⊆ E2 → WP e @ E1 {{ Φ }} ⊢ WP e @ E2 {{ Φ }}.
+Proof. iIntros (?) "H"; iApply (wp_strong_mono E1 E2); auto. iFrame; eauto. Qed.
 Global Instance wp_mono' E e :
-  Proper (pointwise_relation _ (⊢) ==> (⊢)) (@wp Λ Σ E e).
+  Proper (pointwise_relation _ (⊢) ==> (⊢)) (@wp Λ Σ _ E e).
 Proof. by intros Φ Φ' ?; apply wp_mono. Qed.
-Lemma wp_strip_pvs E e P Φ :
-  (P ⊢ WP e @ E {{ Φ }}) → (|={E}=> P) ⊢ WP e @ E {{ Φ }}.
-Proof. move=>->. by rewrite pvs_wp. Qed.
+
 Lemma wp_value E Φ e v : to_val e = Some v → Φ v ⊢ WP e @ E {{ Φ }}.
 Proof. intros; rewrite -(of_to_val e v) //; by apply wp_value'. Qed.
 Lemma wp_value_pvs' E Φ v : (|={E}=> Φ v) ⊢ WP of_val v @ E {{ Φ }}.
@@ -267,54 +190,30 @@ Proof. intros. by rewrite -wp_pvs -wp_value'. Qed.
 Lemma wp_value_pvs E Φ e v :
   to_val e = Some v → (|={E}=> Φ v) ⊢ WP e @ E {{ Φ }}.
 Proof. intros. rewrite -wp_pvs -wp_value //. Qed.
+
 Lemma wp_frame_l E e Φ R : R ★ WP e @ E {{ Φ }} ⊢ WP e @ E {{ v, R ★ Φ v }}.
-Proof. setoid_rewrite (comm _ R); apply wp_frame_r. Qed.
-Lemma wp_frame_step_r' E e Φ R :
-  to_val e = None → WP e @ E {{ Φ }} ★ ▷ R ⊢ WP e @ E {{ v, Φ v ★ R }}.
-Proof.
-  intros. rewrite {2}(_ : E = E ∪ ∅); last by set_solver.
-  rewrite -(wp_frame_step_r E ∅ ∅); [|auto|set_solver..].
-  apply sep_mono_r. rewrite -!pvs_intro. done.
-Qed.
-Lemma wp_frame_step_l E E1 E2 e Φ R :
-  to_val e = None → E ⊥ E1 → E2 ⊆ E1 →
-  (|={E1,E2}=> ▷ |={E2,E1}=> R) ★ WP e @ E {{ Φ }}
-  ⊢ WP e @ (E ∪ E1) {{ v, R ★ Φ v }}.
+Proof. iIntros "[??]". iApply (wp_strong_mono E E _ Φ); try iFrame; eauto. Qed.
+Lemma wp_frame_r E e Φ R : WP e @ E {{ Φ }} ★ R ⊢ WP e @ E {{ v, Φ v ★ R }}.
+Proof. iIntros "[??]". iApply (wp_strong_mono E E _ Φ); try iFrame; eauto. Qed.
+
+Lemma wp_frame_step_r E1 E2 e Φ R :
+  to_val e = None → E2 ⊆ E1 →
+  WP e @ E2 {{ Φ }} ★ (|={E1,E2}=> ▷ |={E2,E1}=> R) ⊢ WP e @ E1 {{ v, Φ v ★ R }}.
 Proof.
-  rewrite [((|={E1,E2}=> _) ★ _)%I]comm; setoid_rewrite (comm _ R).
-  apply wp_frame_step_r.
+  rewrite [(WP _ @ _ {{ _ }} ★ _)%I]comm; setoid_rewrite (comm _ _ R).
+  apply wp_frame_step_l.
 Qed.
 Lemma wp_frame_step_l' E e Φ R :
   to_val e = None → ▷ R ★ WP e @ E {{ Φ }} ⊢ WP e @ E {{ v, R ★ Φ v }}.
-Proof.
-  rewrite (comm _ (â–· R)%I); setoid_rewrite (comm _ R).
-  apply wp_frame_step_r'.
-Qed.
-Lemma wp_always_l E e Φ R `{!PersistentP R} :
-  R ∧ WP e @ E {{ Φ }} ⊢ WP e @ E {{ v, R ∧ Φ v }}.
-Proof. by setoid_rewrite (always_and_sep_l _ _); rewrite wp_frame_l. Qed.
-Lemma wp_always_r E e Φ R `{!PersistentP R} :
-  WP e @ E {{ Φ }} ∧ R ⊢ WP e @ E {{ v, Φ v ∧ R }}.
-Proof. by setoid_rewrite (always_and_sep_r _ _); rewrite wp_frame_r. Qed.
+Proof. iIntros (?) "[??]". iApply (wp_frame_step_l E E); try iFrame; eauto. Qed.
+Lemma wp_frame_step_r' E e Φ R :
+  to_val e = None → WP e @ E {{ Φ }} ★ ▷ R ⊢ WP e @ E {{ v, Φ v ★ R }}.
+Proof. iIntros (?) "[??]". iApply (wp_frame_step_r E E); try iFrame; eauto. Qed.
+
 Lemma wp_wand_l E e Φ Ψ :
   (∀ v, Φ v -★ Ψ v) ★ WP e @ E {{ Φ }} ⊢ WP e @ E {{ Ψ }}.
-Proof.
-  rewrite wp_frame_l. apply wp_mono=> v. by rewrite (forall_elim v) wand_elim_l.
-Qed.
+Proof. by apply wp_strong_mono. Qed.
 Lemma wp_wand_r E e Φ Ψ :
   WP e @ E {{ Φ }} ★ (∀ v, Φ v -★ Ψ v) ⊢ WP e @ E {{ Ψ }}.
 Proof. by rewrite comm wp_wand_l. Qed.
-
-Lemma wp_mask_weaken E1 E2 e Φ :
-  E1 ⊆ E2 → WP e @ E1 {{ Φ }} ⊢ WP e @ E2 {{ Φ }}.
-Proof. auto using wp_mask_frame_mono. Qed.
-
-(** * Weakest-pre is a FSA. *)
-Definition wp_fsa (e : expr Λ) : FSA Λ Σ (val Λ) := λ E, wp E e.
-Global Arguments wp_fsa _ _ / _.
-Global Instance wp_fsa_prf : FrameShiftAssertion (atomic e) (wp_fsa e).
-Proof.
-  rewrite /wp_fsa; split; auto using wp_mask_frame_mono, wp_atomic, wp_frame_r.
-  intros E Φ. by rewrite -(pvs_wp E e Φ) -(wp_pvs E e Φ).
-Qed.
 End wp.
diff --git a/program_logic/weakestpre_fix.v b/program_logic/weakestpre_fix.v
deleted file mode 100644
index 3642b021c1e913cdb7bcd0a7248bc1ce55bf90ea..0000000000000000000000000000000000000000
--- a/program_logic/weakestpre_fix.v
+++ /dev/null
@@ -1,103 +0,0 @@
-From Coq Require Import Wf_nat.
-From iris.program_logic Require Import weakestpre wsat.
-Local Hint Extern 10 (_ ≤ _) => omega.
-Local Hint Extern 10 (_ ⊥ _) => set_solver.
-Local Hint Extern 10 (✓{_} _) =>
-  repeat match goal with
-  | H : wsat _ _ _ _ |- _ => apply wsat_valid in H; last omega
-  end; solve_validN.
-
-(** This files provides an alternative definition of wp in terms of a fixpoint
-of a contractive function, rather than a CoInductive type. This is how we define
-wp on paper.  We show that the two versions are equivalent. *)
-Section def.
-Context {Λ : language} {Σ : iFunctor}.
-Local Notation iProp := (iProp Λ Σ).
-
-Program Definition wp_pre
-    (wp : coPset -c> expr Λ -c> (val Λ -c> iProp) -c> iProp) :
-    coPset -c> expr Λ -c> (val Λ -c> iProp) -c> iProp := λ E e1 Φ,
-  {| uPred_holds n r1 := ∀ k Ef σ1 rf,
-       0 ≤ k < n → E ⊥ Ef →
-       wsat (S k) (E ∪ Ef) σ1 (r1 ⋅ rf) →
-       (∀ v, to_val e1 = Some v →
-         ∃ r2, Φ v (S k) r2 ∧ wsat (S k) (E ∪ Ef) σ1 (r2 ⋅ rf)) ∧
-       (to_val e1 = None → 0 < k →
-         reducible e1 σ1 ∧
-         (∀ e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef →
-           ∃ r2 r2',
-             wsat k (E ∪ Ef) σ2 (r2 ⋅ r2' ⋅ rf) ∧
-             wp E e2 Φ k r2 ∧
-             default True ef (λ ef, wp ⊤ ef (cconst True%I) k r2'))) |}.
-Next Obligation.
-  intros wp E e1 Φ n r1 r2 Hwp [r3 Hr2] k Ef σ1 rf ??.
-  rewrite (dist_le _ _ _ _ Hr2); last lia. intros Hws.
-  destruct (Hwp k Ef σ1 (r3 ⋅ rf)) as [Hval Hstep]; rewrite ?assoc; auto.
-  split.
-  - intros v Hv. destruct (Hval v Hv) as [r4 [??]].
-    exists (r4 â‹… r3). rewrite -assoc. eauto using uPred_mono, cmra_includedN_l.
-  - intros ??. destruct Hstep as [Hred Hpstep]; auto.
-    split; [done|]=> e2 σ2 ef ?.
-    edestruct Hpstep as (r4&r4'&?&?&?); eauto.
-    exists r4, (r4' â‹… r3); split_and?; auto.
-    + by rewrite assoc -assoc.
-    + destruct ef; simpl in *; eauto using uPred_mono, cmra_includedN_l.
-Qed.
-Next Obligation. repeat intro; eauto. Qed.
-
-Local Instance pre_wp_contractive : Contractive wp_pre.
-Proof.
-  assert (∀ n E e Φ r
-    (wp1 wp2 : coPset -c> expr Λ -c> (val Λ -c> iProp) -c> iProp),
-    (∀ i : nat, i < n → wp1 ≡{i}≡ wp2) →
-    wp_pre wp1 E e Φ n r → wp_pre wp2 E e Φ n r) as help.
-  { intros n E e Φ r wp1 wp2 HI Hwp k Ef σ1 rf ???.
-    destruct (Hwp k Ef σ1 rf) as [Hval Hstep]; auto.
-    split; first done.
-    intros ??. destruct Hstep as [Hred Hpstep]; auto.
-    split; [done|]=> e2 σ2 ef ?.
-    destruct (Hpstep e2 σ2 ef) as (r2&r2'&?&?&?); [done..|].
-    exists r2, r2'; split_and?; auto.
-    - apply HI with k; auto.
-    - destruct ef as [ef|]; simpl in *; last done.
-      apply HI with k; auto. }
-  split; split; eapply help; auto using (symmetry (R:=dist _)).
-Qed.
-
-Definition wp_fix : coPset → expr Λ → (val Λ → iProp) → iProp := 
-  fixpoint wp_pre.
-
-Lemma wp_fix_unfold E e Φ : wp_fix E e Φ ⊣⊢ wp_pre wp_fix E e Φ.
-Proof. apply (fixpoint_unfold wp_pre). Qed.
-
-Lemma wp_fix_correct E e (Φ : val Λ → iProp) : wp_fix E e Φ ⊣⊢ wp E e Φ.
-Proof.
-  split=> n r Hr. rewrite wp_eq /wp_def {2}/uPred_holds.
-  split; revert r E e Φ Hr.
-  - induction n as [n IH] using lt_wf_ind=> r1 E e Φ ? Hwp.
-    case EQ: (to_val e)=>[v|].
-    + rewrite -(of_to_val _ _ EQ) {IH}. constructor. rewrite pvs_eq.
-      intros [|k] Ef σ rf ???; first omega.
-      apply wp_fix_unfold in Hwp; last done.
-      destruct (Hwp k Ef σ rf); auto.
-    + constructor; [done|]=> k Ef σ1 rf ???.
-      apply wp_fix_unfold in Hwp; last done.
-      destruct (Hwp k Ef σ1 rf) as [Hval [Hred Hpstep]]; auto.
-      split; [done|]=> e2 σ2 ef ?.
-      destruct (Hpstep e2 σ2 ef) as (r2&r2'&?&?&?); [done..|].
-      exists r2, r2'; split_and?; auto.
-      intros ? ->. change (weakestpre.wp_pre ⊤ (cconst True%I) e' k r2'); eauto.
-  - induction n as [n IH] using lt_wf_ind=> r1 E e Φ ? Hwp.
-    apply wp_fix_unfold; [done|]=> k Ef σ1 rf ???. split.
-    + intros v Hval.
-      destruct Hwp as [??? Hpvs|]; rewrite ?to_of_val in Hval; simplify_eq/=.
-      rewrite pvs_eq in Hpvs.
-      destruct (Hpvs (S k) Ef σ1 rf) as (r2&?&?); eauto.
-    + intros Hval ?.
-      destruct Hwp as [|???? Hwp]; rewrite ?to_of_val in Hval; simplify_eq/=.
-      edestruct (Hwp k Ef σ1 rf) as [? Hpstep]; auto.
-      split; [done|]=> e2 σ2 ef ?.
-      destruct (Hpstep e2 σ2 ef) as (r2&r2'&?&?&?); auto.
-      exists r2, r2'. destruct ef; simpl; auto.
-Qed.
-End def.
diff --git a/program_logic/wsat.v b/program_logic/wsat.v
deleted file mode 100644
index 53f94b2d2ce145bbfbaab19e7ef5738e67b6fb82..0000000000000000000000000000000000000000
--- a/program_logic/wsat.v
+++ /dev/null
@@ -1,180 +0,0 @@
-From iris.prelude Require Export coPset.
-From iris.program_logic Require Export model.
-From iris.algebra Require Export cmra_big_op cmra_tactics.
-From iris.algebra Require Import updates.
-Local Hint Extern 10 (_ ≤ _) => omega.
-Local Hint Extern 10 (✓{_} _) => solve_validN.
-Local Hint Extern 1 (✓{_} gst _) => apply gst_validN.
-Local Hint Extern 1 (✓{_} wld _) => apply wld_validN.
-
-Record wsat_pre {Λ Σ} (n : nat) (E : coPset)
-    (σ : state Λ) (rs : gmap positive (iRes Λ Σ)) (r : iRes Λ Σ) := {
-  wsat_pre_valid : ✓{S n} r;
-  wsat_pre_state : pst r ≡ Excl' σ;
-  wsat_pre_dom i : is_Some (rs !! i) → i ∈ E ∧ is_Some (wld r !! i);
-  wsat_pre_wld i P :
-    i ∈ E →
-    wld r !! i ≡{S n}≡ Some (to_agree (Next (iProp_unfold P))) →
-    ∃ r', rs !! i = Some r' ∧ P n r'
-}.
-Arguments wsat_pre_valid {_ _ _ _ _ _ _} _.
-Arguments wsat_pre_state {_ _ _ _ _ _ _} _.
-Arguments wsat_pre_dom {_ _ _ _ _ _ _} _ _ _.
-Arguments wsat_pre_wld {_ _ _ _ _ _ _} _ _ _ _ _.
-
-Definition wsat {Λ Σ}
-  (n : nat) (E : coPset) (σ : state Λ) (r : iRes Λ Σ) : Prop :=
-  match n with 0 => True | S n => ∃ rs, wsat_pre n E σ rs (r ⋅ big_opM rs) end.
-Instance: Params (@wsat) 5.
-Arguments wsat : simpl never.
-
-Section wsat.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types σ : state Λ.
-Implicit Types r : iRes Λ Σ.
-Implicit Types rs : gmap positive (iRes Λ Σ).
-Implicit Types P : iProp Λ Σ.
-Implicit Types m : iGst Λ Σ.
-
-Instance wsat_ne' : Proper (dist n ==> impl) (@wsat Λ Σ n E σ).
-Proof.
-  intros [|n] E σ r1 r2 Hr; first done; intros [rs [Hdom Hv Hs Hinv]].
-  exists rs; constructor; intros until 0; setoid_rewrite <-Hr; eauto.
-Qed.
-Global Instance wsat_ne n : Proper (dist n ==> iff) (@wsat Λ Σ n E σ) | 1.
-Proof. by intros E σ w1 w2 Hw; split; apply wsat_ne'. Qed.
-Global Instance wsat_proper' n : Proper ((≡) ==> iff) (@wsat Λ Σ n E σ) | 1.
-Proof. by intros E σ w1 w2 Hw; apply wsat_ne, equiv_dist. Qed.
-Lemma wsat_proper n E1 E2 σ r1 r2 :
-  E1 = E2 → r1 ≡ r2 → wsat n E1 σ r1 → wsat n E2 σ r2.
-Proof. by move=>->->. Qed.
-Lemma wsat_le n n' E σ r : wsat n E σ r → n' ≤ n → wsat n' E σ r.
-Proof.
-  destruct n as [|n], n' as [|n']; simpl; try by (auto with lia).
-  intros [rs [Hval Hσ HE Hwld]] ?; exists rs; constructor; auto.
-  intros i P ? (P'&HiP&HP')%dist_Some_inv_r'.
-  destruct (to_agree_uninj (S n) P') as [laterP' HlaterP'].
-  { apply (lookup_validN_Some _ (wld (r â‹… big_opM rs)) i); rewrite ?HiP; auto. }
-  assert (P' ≡{S n}≡ to_agree $ Next $ iProp_unfold $
-                       iProp_fold $ later_car $ laterP') as HPiso.
-  { by rewrite iProp_unfold_fold later_eta HlaterP'. }
-  assert (P ≡{n'}≡ iProp_fold (later_car laterP')) as HPP'.
-  { apply (inj iProp_unfold), (inj Next), (inj to_agree).
-    by rewrite HP' -(dist_le _ _ _ _ HPiso). }
-  destruct (Hwld i (iProp_fold (later_car laterP'))) as (r'&?&?); auto.
-  { by rewrite HiP -HPiso. }
-  assert (✓{S n} r') by (apply (big_opM_lookup_valid _ rs i); auto).
-  exists r'; split; [done|]. apply HPP', uPred_closed with n; auto.
-Qed.
-Lemma wsat_valid n E σ r : n ≠ 0 → wsat n E σ r → ✓{n} r.
-Proof.
-  destruct n; first done.
-  intros _ [rs ?]; eapply cmra_validN_op_l, wsat_pre_valid; eauto.
-Qed.
-Lemma wsat_init k E σ m : ✓{S k} m → wsat (S k) E σ (Res ∅ (Excl' σ) m).
-Proof.
-  intros Hv. exists ∅; constructor; auto.
-  - rewrite big_opM_empty right_id.
-    split_and!; try (apply cmra_valid_validN, ra_empty_valid);
-      constructor || apply Hv.
-  - by intros i; rewrite lookup_empty=>-[??].
-  - intros i P ?; rewrite /= left_id lookup_empty; inversion_clear 1.
-Qed.
-Lemma wsat_open n E σ r i P :
-  wld r !! i ≡{S n}≡ Some (to_agree (Next (iProp_unfold P))) → i ∉ E →
-  wsat (S n) ({[i]} ∪ E) σ r → ∃ rP, wsat (S n) E σ (rP ⋅ r) ∧ P n rP.
-Proof.
-  intros HiP Hi [rs [Hval Hσ HE Hwld]].
-  destruct (Hwld i P) as (rP&?&?); [set_solver +|by apply lookup_wld_op_l|].
-  assert (rP ⋅ r ⋅ big_opM (delete i rs) ≡ r ⋅ big_opM rs) as Hr.
-  { by rewrite (comm _ rP) -assoc big_opM_delete. }
-  exists rP; split; [exists (delete i rs); constructor; rewrite ?Hr|]; auto.
-  - intros j; rewrite lookup_delete_is_Some Hr.
-    generalize (HE j); set_solver +Hi.
-  - intros j P'; rewrite Hr=> Hj ?.
-    setoid_rewrite lookup_delete_ne; last (set_solver +Hi Hj).
-    apply Hwld; [set_solver +Hj|done].
-Qed.
-Lemma wsat_close n E σ r i P rP :
-  wld rP !! i ≡{S n}≡ Some (to_agree (Next (iProp_unfold P))) → i ∉ E →
-  wsat (S n) E σ (rP ⋅ r) → P n rP → wsat (S n) ({[i]} ∪ E) σ r.
-Proof.
-  intros HiP HiE [rs [Hval Hσ HE Hwld]] ?.
-  assert (rs !! i = None) by (apply eq_None_not_Some; naive_solver).
-  assert (r ⋅ big_opM (<[i:=rP]> rs) ≡ rP ⋅ r ⋅ big_opM rs) as Hr.
-  { by rewrite (comm _ rP) -assoc big_opM_insert. }
-  exists (<[i:=rP]>rs); constructor; rewrite ?Hr; auto.
-  - intros j; rewrite Hr lookup_insert_is_Some=>-[?|[??]]; subst.
-    + split. set_solver. rewrite !lookup_op HiP !op_is_Some; eauto.
-    + destruct (HE j) as [Hj Hj']; auto; set_solver +Hj Hj'.
-  - intros j P'; rewrite Hr elem_of_union elem_of_singleton=>-[?|?]; subst.
-    + rewrite !lookup_wld_op_l ?HiP; auto=> HP.
-      apply (inj Some), (inj to_agree), (inj Next), (inj iProp_unfold) in HP.
-      exists rP; split; [rewrite lookup_insert|apply HP]; auto.
-    + intros. destruct (Hwld j P') as (r'&?&?); auto.
-      exists r'; rewrite lookup_insert_ne; naive_solver.
-Qed.
-Lemma wsat_update_pst n E σ1 σ1' r rf :
-  pst r ≡{S n}≡ Excl' σ1 → wsat (S n) E σ1' (r ⋅ rf) →
-  σ1' = σ1 ∧ ∀ σ2, wsat (S n) E σ2 (update_pst σ2 r ⋅ rf).
-Proof.
-  intros Hpst_r [rs [(?&?&?) Hpst HE Hwld]]; simpl in *.
-  assert (pst rf ⋅ pst (big_opM rs) = ∅) as Hpst'.
-  { by apply: (excl_validN_inv_l (S n) _ σ1); rewrite -Hpst_r assoc. }
-  assert (σ1' = σ1) as ->.
-  { apply leibniz_equiv, (timeless _), dist_le with (S n); auto.
-    apply (inj Excl), (inj Some).
-    by rewrite -Hpst_r -Hpst -assoc Hpst' right_id. }
-  split; [done|exists rs].
-  by constructor; first split_and!; try rewrite /= -assoc Hpst'.
-Qed.
-Lemma wsat_update_gst n E σ r rf m1 (P : iGst Λ Σ → Prop) :
-  m1 ≼{S n} gst r → m1 ~~>: P →
-  wsat (S n) E σ (r ⋅ rf) → ∃ m2, wsat (S n) E σ (update_gst m2 r ⋅ rf) ∧ P m2.
-Proof.
-  intros [mf Hr] Hup [rs [(?&?&?) Hσ HE Hwld]].
-  destruct (Hup (S n) (Some (mf â‹… gst (rf â‹… big_opM rs)))) as (m2&?&Hval'); try done.
-  { by rewrite /= (assoc _ m1) -Hr assoc. }
-  exists m2; split; [exists rs|done].
-  by constructor; first split_and!; auto.
-Qed.
-Lemma wsat_alloc n E1 E2 σ r P rP :
-  ¬set_finite E1 → P n rP → wsat (S n) (E1 ∪ E2) σ (rP ⋅ r) →
-  ∃ i, wsat (S n) (E1 ∪ E2) σ
-         (Res {[i := to_agree (Next (iProp_unfold P))]} ∅ ∅ ⋅ r) ∧
-       wld r !! i = None ∧ i ∈ E1.
-Proof.
-  intros HE1 ? [rs [Hval Hσ HE Hwld]].
-  assert (∃ i, i ∈ E1 ∧ wld r !! i = None ∧ wld rP !! i = None ∧
-                        wld (big_opM rs) !! i = None) as (i&Hi&Hri&HrPi&Hrsi).
-  { exists (coPpick (E1 ∖
-      (dom _ (wld r) ∪ (dom _ (wld rP) ∪ dom _ (wld (big_opM rs)))))).
-    rewrite -!not_elem_of_dom -?not_elem_of_union -elem_of_difference.
-    apply coPpick_elem_of=>HE'; eapply HE1, (difference_finite_inv _ _), HE'.
-    by repeat apply union_finite; apply dom_finite. }
-  assert (rs !! i = None).
-  { apply eq_None_not_Some=>?; destruct (HE i) as [_ Hri']; auto; revert Hri'.
-    rewrite /= !lookup_op !op_is_Some -!not_eq_None_Some; tauto. }
-  assert (r ⋅ big_opM (<[i:=rP]> rs) ≡ rP ⋅ r ⋅ big_opM rs) as Hr.
-  { by rewrite (comm _ rP) -assoc big_opM_insert. }
-  exists i; split_and?; [exists (<[i:=rP]>rs); constructor| |]; auto.
-  - destruct Hval as (?&?&?);  rewrite -assoc Hr.
-    split_and!; rewrite /= ?left_id; [|eauto|eauto].
-    intros j; destruct (decide (j = i)) as [->|].
-    + by rewrite !lookup_op Hri HrPi Hrsi !right_id lookup_singleton.
-    + by rewrite lookup_op lookup_singleton_ne // left_id.
-  - by rewrite -assoc Hr /= left_id.
-  - intros j; rewrite -assoc Hr; destruct (decide (j = i)) as [->|].
-    + intros _; split; first set_solver +Hi.
-      rewrite /= !lookup_op lookup_singleton !op_is_Some; eauto.
-    + rewrite lookup_insert_ne //.
-      rewrite lookup_op lookup_singleton_ne // left_id; eauto.
-  - intros j P'; rewrite -assoc Hr; destruct (decide (j=i)) as [->|].
-    + rewrite /= !lookup_op Hri HrPi Hrsi right_id lookup_singleton=>? HP.
-      apply (inj Some), (inj to_agree), (inj Next), (inj iProp_unfold) in HP.
-      exists rP; rewrite lookup_insert; split; [|apply HP]; auto.
-    + rewrite /= lookup_op lookup_singleton_ne // left_id=> ??.
-      destruct (Hwld j P') as [r' ?]; auto.
-      by exists r'; rewrite lookup_insert_ne.
-Qed.
-End wsat.
diff --git a/proofmode/class_instances.v b/proofmode/class_instances.v
index d53bae6a6a99101f4884852470bd45bd29fb6b58..27b7f8db8ec40523c704e674fec4531a01afd112 100644
--- a/proofmode/class_instances.v
+++ b/proofmode/class_instances.v
@@ -328,4 +328,18 @@ Proof. rewrite /IntoExist=> HP ?. by rewrite HP later_exist. Qed.
 Global Instance into_exist_always {A} P (Φ : A → uPred M) :
   IntoExist P Φ → IntoExist (□ P) (λ a, □ (Φ a))%I.
 Proof. rewrite /IntoExist=> HP. by rewrite HP always_exist. Qed.
+
+(* IsTrueNow *)
+Global Instance is_now_True_now_True P : IsNowTrue (â—‡ P).
+Proof. by rewrite /IsNowTrue now_True_idemp. Qed.
+Global Instance is_now_True_later P : IsNowTrue (â–· P).
+Proof. by rewrite /IsNowTrue now_True_later. Qed.
+
+(* FromViewShift *)
+Global Instance from_vs_rvs P : FromVs (|=r=> P) P.
+Proof. done. Qed.
+
+(* ElimViewShift *)
+Global Instance elim_vs_rvs_rvs P Q : ElimVs (|=r=> P) P (|=r=> Q) (|=r=> Q).
+Proof. by rewrite /ElimVs rvs_frame_r wand_elim_r rvs_trans. Qed.
 End classes.
diff --git a/proofmode/classes.v b/proofmode/classes.v
index 0c0a1cc283ed7d546b38528f35bf2385fd53130e..037de8467d3b37024c44225c83ad628925beb476 100644
--- a/proofmode/classes.v
+++ b/proofmode/classes.v
@@ -55,7 +55,13 @@ Class IntoExist {A} (P : uPred M) (Φ : A → uPred M) :=
   into_exist : P ⊢ ∃ x, Φ x.
 Global Arguments into_exist {_} _ _ {_}.
 
-Class TimelessElim (Q : uPred M) :=
-  timeless_elim `{!TimelessP P} : ▷ P ★ (P -★ Q) ⊢ Q.
-Global Arguments timeless_elim _ {_} _ {_}.
+Class IsNowTrue (Q : uPred M) := is_now_True : ◇ Q ⊢ Q.
+Global Arguments is_now_True : clear implicits.
+
+Class FromVs (P Q : uPred M) := from_vs : (|=r=> Q) ⊢ P.
+Global Arguments from_vs : clear implicits.
+
+Class ElimVs (P P' : uPred M) (Q Q' : uPred M) :=
+  elim_vs : P ★ (P' -★ Q') ⊢ Q.
+Arguments elim_vs _ _ _ _ {_}.
 End classes.
diff --git a/proofmode/coq_tactics.v b/proofmode/coq_tactics.v
index ebe55ff1941bce5d701972cd74eb1203ff7612c6..1c0d4f3ac4d9c2452260434c35f52d52232059bf 100644
--- a/proofmode/coq_tactics.v
+++ b/proofmode/coq_tactics.v
@@ -376,13 +376,14 @@ Proof.
 Qed.
 
 Lemma tac_timeless Δ Δ' i p P Q :
-  TimelessElim Q →
+  IsNowTrue Q →
   envs_lookup i Δ = Some (p, ▷ P)%I → TimelessP P →
   envs_simple_replace i p (Esnoc Enil i P) Δ = Some Δ' →
   (Δ' ⊢ Q) → Δ ⊢ Q.
 Proof.
   intros ???? HQ. rewrite envs_simple_replace_sound //; simpl.
-  by rewrite always_if_later right_id HQ timeless_elim.
+  rewrite always_if_later right_id HQ -{2}(is_now_True Q).
+  by rewrite (timelessP (â–¡?p P)) now_True_frame_r wand_elim_r.
 Qed.
 
 (** * Always *)
@@ -734,4 +735,18 @@ Proof.
   apply exist_elim=> a; destruct (HΦ a) as (Δ'&?&?).
   rewrite envs_simple_replace_sound' //; simpl. by rewrite right_id wand_elim_r.
 Qed.
+
+(** * Viewshifts *)
+Lemma tac_vs_intro Δ P Q : FromVs P Q → (Δ ⊢ Q) → Δ ⊢ P.
+Proof. rewrite /FromVs. intros <- ->. apply rvs_intro. Qed.
+
+Lemma tac_vs_elim Δ Δ' i p P' P Q Q' :
+  envs_lookup i Δ = Some (p, P) →
+  ElimVs P P' Q Q' →
+  envs_replace i p false (Esnoc Enil i P') Δ = Some Δ' →
+  (Δ' ⊢ Q') → Δ ⊢ Q.
+Proof.
+  intros ??? HΔ. rewrite envs_replace_sound //; simpl.
+  rewrite right_id HΔ always_if_elim. by apply elim_vs.
+Qed.
 End tactics.
diff --git a/proofmode/ghost_ownership.v b/proofmode/ghost_ownership.v
index 2c557444aa43cd73ce27521f0343a33329695cb7..be37d4c7d0e71df507adcecd6afc8d79b92aaa6b 100644
--- a/proofmode/ghost_ownership.v
+++ b/proofmode/ghost_ownership.v
@@ -3,7 +3,7 @@ From iris.proofmode Require Export tactics.
 From iris.program_logic Require Export ghost_ownership.
 
 Section ghost.
-Context `{inG Λ Σ A}.
+Context `{inG Σ A}.
 Implicit Types a b : A.
 
 Global Instance into_sep_own p γ a b1 b2 :
diff --git a/proofmode/invariants.v b/proofmode/invariants.v
index 46fb3cf1e01432776d1a73f861df26783c7da0b2..936545e29ecf3271f43f103e3ea169dc871999c0 100644
--- a/proofmode/invariants.v
+++ b/proofmode/invariants.v
@@ -1,89 +1,28 @@
-From iris.proofmode Require Import coq_tactics.
-From iris.proofmode Require Export pviewshifts.
+From iris.proofmode Require Export tactics pviewshifts.
 From iris.program_logic Require Export invariants.
-Import uPred.
+From iris.proofmode Require Import coq_tactics intro_patterns.
 
-Section invariants.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types N : namespace.
-Implicit Types P Q R : iProp Λ Σ.
+Tactic Notation "iInvCore" constr(N) "as" tactic(tac) constr(Hclose) :=
+  let Htmp := iFresh in
+  let patback := intro_pat.parse_one Hclose in
+  let pat := constr:(IList [[IName Htmp; patback]]) in
+  iVs (inv_open _ N with "[#]") as pat;
+    [idtac|iAssumption || fail "iInv: invariant" N "not found"|idtac];
+    [solve_ndisj|tac Htmp].
 
-Lemma tac_inv_fsa {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E N i P Q Φ :
-  IsFSA Q E fsa fsaV Φ →
-  fsaV → nclose N ⊆ E → (of_envs Δ ⊢ inv N P) →
-  envs_app false (Esnoc Enil i (▷ P)) Δ = Some Δ' →
-  (Δ' ⊢ fsa (E ∖ nclose N) (λ a, ▷ P ★ Φ a)) → Δ ⊢ Q.
-Proof.
-  intros ????? HΔ'. rewrite (is_fsa Q) -(inv_fsa fsa _ _ P) //.
-  rewrite // -always_and_sep_l. apply and_intro; first done.
-  rewrite envs_app_sound //; simpl. by rewrite right_id HΔ'.
-Qed.
-
-Lemma tac_inv_fsa_timeless {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E N i P Q Φ :
-  IsFSA Q E fsa fsaV Φ →
-  fsaV → nclose N ⊆ E → (of_envs Δ ⊢ inv N P) → TimelessP P →
-  envs_app false (Esnoc Enil i P) Δ = Some Δ' →
-  (Δ' ⊢ fsa (E ∖ nclose N) (λ a, P ★ Φ a)) → Δ ⊢ Q.
-Proof.
-  intros ?????? HΔ'. rewrite (is_fsa Q) -(inv_fsa fsa _ _ P) //.
-  rewrite // -always_and_sep_l. apply and_intro, wand_intro_l; first done.
-  trans (|={E ∖ N}=> P ★ Δ)%I; first by rewrite pvs_timeless pvs_frame_r.
-  apply (fsa_strip_pvs _).
-  rewrite envs_app_sound //; simpl. rewrite right_id HΔ' wand_elim_r.
-  apply: fsa_mono=> v. by rewrite -later_intro.
-Qed.
-End invariants.
-
-Tactic Notation "iInvCore" constr(N) "as" constr(H) :=
-  eapply tac_inv_fsa with _ _ _ _ N H _ _;
-    [let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
-     apply _ || fail "iInv: cannot viewshift in goal" P
-    |trivial with fsaV
-    |solve_ndisj
-    |iAssumption || fail "iInv: invariant" N "not found"
-    |env_cbv; reflexivity
-    |simpl (* get rid of FSAs *)].
-
-Tactic Notation "iInv" constr(N) "as" constr(pat) :=
-  let H := iFresh in iInvCore N as H; last iDestruct H as pat.
+Tactic Notation "iInv" constr(N) "as" constr(pat) constr(Hclose) :=
+   iInvCore N as (fun H => iDestruct H as pat) Hclose.
 Tactic Notation "iInv" constr(N) "as" "(" simple_intropattern(x1) ")"
-    constr(pat) :=
-  let H := iFresh in iInvCore N as H; last iDestruct H as (x1) pat.
+    constr(pat) constr(Hclose) :=
+   iInvCore N as (fun H => iDestruct H as (x1) pat) Hclose.
 Tactic Notation "iInv" constr(N) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) ")" constr(pat) :=
-  let H := iFresh in iInvCore N as H; last iDestruct H as (x1 x2) pat.
+    simple_intropattern(x2) ")" constr(pat) constr(Hclose) :=
+   iInvCore N as (fun H => iDestruct H as (x1 x2) pat) Hclose.
 Tactic Notation "iInv" constr(N) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
-  let H := iFresh in iInvCore N as H; last iDestruct H as (x1 x2 x3) pat.
+    simple_intropattern(x2) simple_intropattern(x3) ")"
+    constr(pat) constr(Hclose) :=
+   iInvCore N as (fun H => iDestruct H as (x1 x2 x3) pat) Hclose.
 Tactic Notation "iInv" constr(N) "as" "(" simple_intropattern(x1)
     simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
-    constr(pat) :=
-  let H := iFresh in iInvCore N as H; last iDestruct H as (x1 x2 x3 x4) pat.
-
-Tactic Notation "iInvCore>" constr(N) "as" constr(H) :=
-  eapply tac_inv_fsa_timeless with _ _ _ _ N H _ _;
-    [let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
-     apply _ || fail "iInv: cannot viewshift in goal" P
-    |trivial with fsaV
-    |solve_ndisj
-    |iAssumption || fail "iOpenInv: invariant" N "not found"
-    |let P := match goal with |- TimelessP ?P => P end in
-     apply _ || fail "iInv:" P "not timeless"
-    |env_cbv; reflexivity
-    |simpl (* get rid of FSAs *)].
-
-Tactic Notation "iInv>" constr(N) "as" constr(pat) :=
-  let H := iFresh in iInvCore> N as H; last iDestruct H as pat.
-Tactic Notation "iInv>" constr(N) "as" "(" simple_intropattern(x1) ")"
-    constr(pat) :=
-  let H := iFresh in iInvCore> N as H; last iDestruct H as (x1) pat.
-Tactic Notation "iInv>" constr(N) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) ")" constr(pat) :=
-  let H := iFresh in iInvCore> N as H; last iDestruct H as (x1 x2) pat.
-Tactic Notation "iInv>" constr(N) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
-  let H := iFresh in iInvCore> N as H; last iDestruct H as (x1 x2 x3) pat.
-Tactic Notation "iInv>" constr(N) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
-    constr(pat) :=
-  let H := iFresh in iInvCore> N as H; last iDestruct H as (x1 x2 x3 x4) pat.
+    constr(pat) constr(Hclose) :=
+   iInvCore N as (fun H => iDestruct H as (x1 x2 x3 x4) pat) Hclose.
diff --git a/proofmode/pviewshifts.v b/proofmode/pviewshifts.v
index 2de2ce5f9ba084d9caf79e4f22e2a370e7d17dc0..36b8a758c3b8603a7948adea41214bbb42f7311c 100644
--- a/proofmode/pviewshifts.v
+++ b/proofmode/pviewshifts.v
@@ -1,17 +1,19 @@
-From iris.proofmode Require Import coq_tactics spec_patterns.
-From iris.proofmode Require Export tactics.
+From iris.proofmode Require Import coq_tactics.
+From iris.proofmode Require Export tactics ghost_ownership.
 From iris.program_logic Require Export pviewshifts.
 Import uPred.
 
 Section pvs.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P Q : iProp Λ Σ.
+Context `{irisG Λ Σ}.
+Implicit Types P Q : iProp Σ.
 
 Global Instance from_pure_pvs E P φ : FromPure P φ → FromPure (|={E}=> P) φ.
 Proof. intros ??. by rewrite -pvs_intro (from_pure P). Qed.
+
 Global Instance from_assumption_pvs E p P Q :
-  FromAssumption p P Q → FromAssumption p P (|={E}=> Q)%I.
-Proof. rewrite /FromAssumption=>->. apply pvs_intro. Qed.
+  FromAssumption p P (|=r=> Q) → FromAssumption p P (|={E}=> Q)%I.
+Proof. rewrite /FromAssumption=>->. apply rvs_pvs. Qed.
+
 Global Instance into_wand_pvs E1 E2 R P Q :
   IntoWand R P Q → IntoWand R (|={E1,E2}=> P) (|={E1,E2}=> Q) | 100.
 Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed.
@@ -19,140 +21,38 @@ Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed.
 Global Instance from_sep_pvs E P Q1 Q2 :
   FromSep P Q1 Q2 → FromSep (|={E}=> P) (|={E}=> Q1) (|={E}=> Q2).
 Proof. rewrite /FromSep=><-. apply pvs_sep. Qed.
+
 Global Instance or_split_pvs E1 E2 P Q1 Q2 :
   FromOr P Q1 Q2 → FromOr (|={E1,E2}=> P) (|={E1,E2}=> Q1) (|={E1,E2}=> Q2).
 Proof. rewrite /FromOr=><-. apply or_elim; apply pvs_mono; auto with I. Qed.
-Global Instance exists_split_pvs {A} E1 E2 P (Φ : A → iProp Λ Σ) :
+
+Global Instance exists_split_pvs {A} E1 E2 P (Φ : A → iProp Σ) :
   FromExist P Φ → FromExist (|={E1,E2}=> P) (λ a, |={E1,E2}=> Φ a)%I.
 Proof.
   rewrite /FromExist=><-. apply exist_elim=> a. by rewrite -(exist_intro a).
 Qed.
+
 Global Instance frame_pvs E1 E2 R P Q :
   Frame R P Q → Frame R (|={E1,E2}=> P) (|={E1,E2}=> Q).
 Proof. rewrite /Frame=><-. by rewrite pvs_frame_l. Qed.
 
-Global Instance timeless_elim_pvs E1 E2 Q : TimelessElim (|={E1,E2}=> Q).
-Proof.
-  intros P ?. rewrite (pvs_timeless E1 P) pvs_frame_r.
-  by rewrite wand_elim_r pvs_trans; last set_solver.
-Qed.
-
-Class IsFSA {A} (P : iProp Λ Σ) (E : coPset)
-    (fsa : FSA Λ Σ A) (fsaV : Prop) (Φ : A → iProp Λ Σ) := {
-  is_fsa : P ⊣⊢ fsa E Φ;
-  is_fsa_is_fsa :> FrameShiftAssertion fsaV fsa;
-}.
-Global Arguments is_fsa {_} _ _ _ _ _ {_}.
-Global Instance is_fsa_pvs E P :
-  IsFSA (|={E}=> P)%I E pvs_fsa True (λ _, P).
-Proof. split. done. apply _. Qed.
-Global Instance is_fsa_fsa {A} (fsa : FSA Λ Σ A) E Φ :
-  FrameShiftAssertion fsaV fsa → IsFSA (fsa E Φ) E fsa fsaV Φ.
-Proof. done. Qed.
-
-Global Instance to_assert_pvs {A} P Q E (fsa : FSA Λ Σ A) fsaV Φ :
-  IsFSA Q E fsa fsaV Φ → IntoAssert P Q (|={E}=> P).
-Proof.
-  intros. by rewrite /IntoAssert pvs_frame_r wand_elim_r (is_fsa Q) fsa_pvs_fsa.
-Qed.
-Global Instance timeless_elim_fsa {A} Q E (fsa : FSA Λ Σ A) fsaV Φ :
-  IsFSA Q E fsa fsaV Φ → TimelessElim Q.
-Proof.
-  intros ? P ?. rewrite (is_fsa Q) -{2}fsa_pvs_fsa.
-  by rewrite (pvs_timeless _ P) pvs_frame_r wand_elim_r.
-Qed.
+Global Instance to_assert_pvs E1 E2 P Q :
+  IntoAssert P (|={E1,E2}=> Q) (|={E1}=> P).
+Proof. intros. by rewrite /IntoAssert pvs_frame_r wand_elim_r pvs_trans. Qed.
 
-Lemma tac_pvs_intro Δ E1 E2 Q : E1 = E2 → (Δ ⊢ Q) → Δ ⊢ |={E1,E2}=> Q.
-Proof. intros -> ->. apply pvs_intro. Qed.
+Global Instance is_now_True_pvs E1 E2 P : IsNowTrue (|={E1,E2}=> P).
+Proof. by rewrite /IsNowTrue now_True_pvs. Qed.
 
-Lemma tac_pvs_elim Δ Δ' E1 E2 E3 i p P' E1' E2' P Q :
-  envs_lookup i Δ = Some (p, P') → P' = (|={E1',E2'}=> P)%I →
-  (E1' = E1 ∧ E2' = E2 ∧ E2 ⊆ E1 ∪ E3
-  ∨ E2 = E2' ∪ E1 ∖ E1' ∧ E2' ⊥ E1 ∖ E1' ∧ E1' ⊆ E1 ∧ E2' ⊆ E1' ∪ E3) →
-  envs_replace i p false (Esnoc Enil i P) Δ = Some Δ' →
-  (Δ' ={E2,E3}=> Q) → Δ ={E1,E3}=> Q.
-Proof.
-  intros ? -> HE ? HQ. rewrite envs_replace_sound //; simpl.
-  rewrite always_if_elim right_id pvs_frame_r wand_elim_r HQ.
-  destruct HE as [(?&?&?)|(?&?&?&?)]; subst; first (by apply pvs_trans).
-  etrans; [eapply pvs_mask_frame'|apply pvs_trans]; auto; set_solver.
-Qed.
+Global Instance from_vs_pvs E P : FromVs (|={E}=> P) P.
+Proof. by rewrite /FromVs -rvs_pvs. Qed.
 
-Lemma tac_pvs_elim_fsa {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E i p P' P Q Φ :
-  envs_lookup i Δ = Some (p, P') → P' = (|={E}=> P)%I →
-  IsFSA Q E fsa fsaV Φ →
-  envs_replace i p false (Esnoc Enil i P) Δ = Some Δ' →
-  (Δ' ⊢ fsa E Φ) → Δ ⊢ Q.
-Proof.
-  intros ? -> ??. rewrite (is_fsa Q) -fsa_pvs_fsa.
-  eapply tac_pvs_elim; set_solver.
-Qed.
+Global Instance elim_vs_rvs_pvs E1 E2 P Q :
+  ElimVs (|=r=> P) P (|={E1,E2}=> Q) (|={E1,E2}=> Q).
+Proof. by rewrite /ElimVs (rvs_pvs E1) pvs_frame_r wand_elim_r pvs_trans. Qed.
+Global Instance elim_vs_pvs_pvs E1 E2 E3 P Q :
+  ElimVs (|={E1,E2}=> P) P (|={E1,E3}=> Q) (|={E2,E3}=> Q).
+Proof. by rewrite /ElimVs pvs_frame_r wand_elim_r pvs_trans. Qed.
 End pvs.
 
-Tactic Notation "iPvsIntro" := apply tac_pvs_intro; first try fast_done.
-
-Tactic Notation "iPvsCore" constr(H) :=
-  match goal with
-  | |- _ ⊢ |={_,_}=> _ =>
-    eapply tac_pvs_elim with _ _ H _ _ _ _ _;
-      [env_cbv; reflexivity || fail "iPvs:" H "not found"
-      |let P := match goal with |- ?P = _ => P end in
-       reflexivity || fail "iPvs:" H ":" P "not a pvs with the right mask"
-      |first
-         [left; split_and!;
-            [reflexivity|reflexivity|try (rewrite (idemp_L (∪)); reflexivity)]
-         |right; split; first reflexivity]
-      |env_cbv; reflexivity|]
-  | |- _ =>
-    eapply tac_pvs_elim_fsa with _ _ _ _ H _ _ _ _;
-      [env_cbv; reflexivity || fail "iPvs:" H "not found"
-      |let P := match goal with |- ?P = _ => P end in
-       reflexivity || fail "iPvs:" H ":" P "not a pvs with the right mask"
-      |let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
-       apply _ || fail "iPvs:" P "not a pvs"
-      |env_cbv; reflexivity|simpl]
-  end.
-
-Tactic Notation "iPvs" open_constr(lem) :=
-  iDestructCore lem as (fun H => iPvsCore H; last iDestruct H as "?").
-Tactic Notation "iPvs" open_constr(lem) "as" constr(pat) :=
-  iDestructCore lem as (fun H => iPvsCore H; last iDestruct H as pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1) ")"
-    constr(pat) :=
-  iDestructCore lem as (fun H => iPvsCore H; last iDestruct H as ( x1 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) ")" constr(pat) :=
-  iDestructCore lem as (fun H => iPvsCore H; last iDestruct H as ( x1 x2 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
-  iDestructCore lem as (fun H => iPvsCore H; last iDestruct H as ( x1 x2 x3 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
-    constr(pat) :=
-  iDestructCore lem as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) ")" constr(pat) :=
-  iDestructCore lem as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) :=
-  iDestructCore lem as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")"
-    constr(pat) :=
-  iDestructCore lem as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 x7 ) pat).
-Tactic Notation "iPvs" open_constr(lem) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7)
-    simple_intropattern(x8) ")" constr(pat) :=
-  iDestructCore lem as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat).
-
 Hint Extern 2 (of_envs _ ⊢ _) =>
-  match goal with |- _ ⊢ (|={_}=> _)%I => iPvsIntro end.
+  match goal with |- _ ⊢ |={_}=> _ => iVsIntro end.
diff --git a/proofmode/sts.v b/proofmode/sts.v
deleted file mode 100644
index 60dfe21befb38c98b0d09be253834313b4c5f2f5..0000000000000000000000000000000000000000
--- a/proofmode/sts.v
+++ /dev/null
@@ -1,46 +0,0 @@
-From iris.proofmode Require Import coq_tactics pviewshifts.
-From iris.proofmode Require Export tactics.
-From iris.program_logic Require Export sts.
-Import uPred.
-
-Section sts.
-Context `{stsG Λ Σ sts} (φ : sts.state sts → iPropG Λ Σ).
-Implicit Types P Q : iPropG Λ Σ.
-
-Lemma tac_sts_fsa {A} (fsa : FSA Λ _ A) fsaV Δ E N i γ S T Q Φ :
-  IsFSA Q E fsa fsaV Φ →
-  fsaV →
-  envs_lookup i Δ = Some (false, sts_ownS γ S T) →
-  (of_envs Δ ⊢ sts_ctx γ N φ) → nclose N ⊆ E →
-  (∀ s, s ∈ S → ∃ Δ',
-    envs_simple_replace i false (Esnoc Enil i (▷ φ s)) Δ = Some Δ' ∧
-    (Δ' ⊢ fsa (E ∖ nclose N) (λ a, ∃ s' T',
-      ■ sts.steps (s, T) (s', T') ★ ▷ φ s' ★ (sts_own γ s' T' -★ Φ a)))) →
-  Δ ⊢ Q.
-Proof.
-  intros ????? HΔ'. rewrite (is_fsa Q) -(sts_fsaS φ fsa) //.
-  rewrite // -always_and_sep_l. apply and_intro; first done.
-  rewrite envs_lookup_sound //; simpl; apply sep_mono_r.
-  apply forall_intro=>s; apply wand_intro_l.
-  rewrite -assoc; apply pure_elim_sep_l=> Hs.
-  destruct (HΔ' s) as (Δ'&?&?); clear HΔ'; auto.
-  rewrite envs_simple_replace_sound' //; simpl.
-  by rewrite right_id wand_elim_r.
-Qed.
-End sts.
-
-Tactic Notation "iSts" constr(H) "as"
-    simple_intropattern(s) simple_intropattern(Hs) :=
-  match type of H with
-  | string => eapply tac_sts_fsa with _ _ _ _ _ _ H _ _ _ _
-  | gname => eapply tac_sts_fsa with _ _ _ _ _ _ _ H _ _ _
-  | _ => fail "iSts:" H "not a string or gname"
-  end;
-    [let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
-     apply _ || fail "iSts: cannot viewshift in goal" P
-    |auto with fsaV
-    |iAssumptionCore || fail "iSts:" H "not found"
-    |iAssumption || fail "iSts: invariant not found"
-    |solve_ndisj
-    |intros s Hs; eexists; split; [env_cbv; reflexivity|simpl]].
-Tactic Notation "iSts" constr(H) "as" simple_intropattern(s) := iSts H as s ?.
diff --git a/proofmode/tactics.v b/proofmode/tactics.v
index 107eea1e11ff94cc3db02aeaced447e6999ce42d..651456ee14fb2ad3be0e3cc520fbacd05b524d97 100644
--- a/proofmode/tactics.v
+++ b/proofmode/tactics.v
@@ -547,11 +547,11 @@ Tactic Notation "iNext":=
 
 Tactic Notation "iTimeless" constr(H) :=
   eapply tac_timeless with _ H _ _;
-    [let Q := match goal with |- TimelessElim ?Q => Q end in
-     apply _ || fail "iTimeless: cannot eliminate later in goal" Q
+    [let Q := match goal with |- IsNowTrue ?Q => Q end in
+     apply _ || fail "iTimeless: cannot remove later of timeless hypothesis in goal" Q
     |env_cbv; reflexivity || fail "iTimeless:" H "not found"
     |let P := match goal with |- TimelessP ?P => P end in
-     apply _ || fail "iTimeless: " P "not timeless"
+     apply _ || fail "iTimeless:" P "not timeless"
     |env_cbv; reflexivity|].
 
 (** * Introduction tactic *)
@@ -856,6 +856,61 @@ Ltac iSimplifyEq := repeat (
   iMatchGoal ltac:(fun H P => match P with (_ = _)%I => iDestruct H as %? end)
   || simplify_eq/=).
 
+(** * View shifts *)
+Tactic Notation "iVsIntro" :=
+  eapply tac_vs_intro;
+    [let P := match goal with |- FromVs ?P _ => P end in
+     apply _ || fail "iVsIntro:" P "not a viewshift"|].
+
+Tactic Notation "iVsCore" constr(H) :=
+  eapply tac_vs_elim with _ H _ _ _ _;
+    [env_cbv; reflexivity || fail "iVs:" H "not found"
+    |let P := match goal with |- ElimVs ?P _ _ _ => P end in
+     let Q := match goal with |- ElimVs _ _ _ ?Q => Q end in
+     apply _ || fail "iVs: cannot eliminate" H ":" P "in" Q
+    |env_cbv; reflexivity|].
+
+Tactic Notation "iVs" open_constr(lem) :=
+  iDestructCore lem as (fun H => iVsCore H; last iDestruct H as "?").
+Tactic Notation "iVs" open_constr(lem) "as" constr(pat) :=
+  iDestructCore lem as (fun H => iVsCore H; last iDestruct H as pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1) ")"
+    constr(pat) :=
+  iDestructCore lem as (fun H => iVsCore H; last iDestruct H as ( x1 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) ")" constr(pat) :=
+  iDestructCore lem as (fun H => iVsCore H; last iDestruct H as ( x1 x2 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
+  iDestructCore lem as (fun H => iVsCore H; last iDestruct H as ( x1 x2 x3 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
+    constr(pat) :=
+  iDestructCore lem as (fun H =>
+    iVsCore H; last iDestruct H as ( x1 x2 x3 x4 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
+    simple_intropattern(x5) ")" constr(pat) :=
+  iDestructCore lem as (fun H =>
+    iVsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
+    simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) :=
+  iDestructCore lem as (fun H =>
+    iVsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
+    simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")"
+    constr(pat) :=
+  iDestructCore lem as (fun H =>
+    iVsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 x7 ) pat).
+Tactic Notation "iVs" open_constr(lem) "as" "(" simple_intropattern(x1)
+    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
+    simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7)
+    simple_intropattern(x8) ")" constr(pat) :=
+  iDestructCore lem as (fun H =>
+    iVsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat).
+
 (* Make sure that by and done solve trivial things in proof mode *)
 Hint Extern 0 (of_envs _ ⊢ _) => by iPureIntro.
 Hint Extern 0 (of_envs _ ⊢ _) => iAssumption.
@@ -866,11 +921,12 @@ Hint Resolve uPred.eq_refl'. (* Maybe make an [iReflexivity] tactic *)
 but then [eauto] mysteriously fails. See bug 4762 *)
 Hint Extern 1 (of_envs _ ⊢ _) =>
   match goal with
-  | |- _ ⊢ (_ ∧ _)%I => iSplit
-  | |- _ ⊢ (_ ★ _)%I => iSplit
-  | |- _ ⊢ (▷ _)%I => iNext
-  | |- _ ⊢ (□ _)%I => iClear "*"; iAlways
-  | |- _ ⊢ (∃ _, _)%I => iExists _
+  | |- _ ⊢ _ ∧ _ => iSplit
+  | |- _ ⊢ _ ★ _ => iSplit
+  | |- _ ⊢ ▷ _ => iNext
+  | |- _ ⊢ □ _ => iClear "*"; iAlways
+  | |- _ ⊢ ∃ _, _ => iExists _
+  | |- _ ⊢ |=r=> _ => iVsIntro
   end.
 Hint Extern 1 (of_envs _ ⊢ _) =>
   match goal with |- _ ⊢ (_ ∨ _)%I => iLeft end.
diff --git a/proofmode/weakestpre.v b/proofmode/weakestpre.v
index e7249a464a4880bf6186676c4dd352c7fab5566b..91dd17e2c89d2136fb10cb9dfec0c1373dc61dad 100644
--- a/proofmode/weakestpre.v
+++ b/proofmode/weakestpre.v
@@ -1,17 +1,36 @@
+From iris.proofmode Require Export classes pviewshifts.
 From iris.proofmode Require Import coq_tactics.
-From iris.proofmode Require Export pviewshifts.
 From iris.program_logic Require Export weakestpre.
 Import uPred.
 
 Section weakestpre.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P Q : iProp Λ Σ.
-Implicit Types Φ : val Λ → iProp Λ Σ.
+Context `{irisG Λ Σ}.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val Λ → iProp Σ.
 
 Global Instance frame_wp E e R Φ Ψ :
   (∀ v, Frame R (Φ v) (Ψ v)) → Frame R (WP e @ E {{ Φ }}) (WP e @ E {{ Ψ }}).
 Proof. rewrite /Frame=> HR. rewrite wp_frame_l. apply wp_mono, HR. Qed.
-Global Instance is_fsa_wp E e Φ :
-  IsFSA (WP e @ E {{ Φ }})%I E (wp_fsa e) (language.atomic e) Φ.
-Proof. split. done. apply _. Qed.
+
+Global Instance to_assert_wp E e P Φ :
+  IntoAssert P (WP e @ E {{ Ψ }}) (|={E}=> P).
+Proof. intros. by rewrite /IntoAssert pvs_frame_r wand_elim_r pvs_wp. Qed.
+
+Global Instance is_now_True_wp E e Φ : IsNowTrue (WP e @ E {{ Φ }}).
+Proof. by rewrite /IsNowTrue -{2}pvs_wp -now_True_pvs -pvs_intro. Qed.
+
+Global Instance elim_vs_rvs_wp E e P Φ :
+  ElimVs (|=r=> P) P (WP e @ E {{ Φ }}) (WP e @ E {{ Φ }}).
+Proof. by rewrite /ElimVs (rvs_pvs E) pvs_frame_r wand_elim_r pvs_wp. Qed.
+
+Global Instance elim_vs_pvs_wp E e P Φ :
+  ElimVs (|={E}=> P) P (WP e @ E {{ Φ }}) (WP e @ E {{ Φ }}).
+Proof. by rewrite /ElimVs pvs_frame_r wand_elim_r pvs_wp. Qed.
+
+(* lower precedence, if possible, it should always pick elim_vs_pvs_wp *)
+Global Instance elim_vs_pvs_wp_atomic E1 E2 e P Φ :
+  atomic e →
+  ElimVs (|={E1,E2}=> P) P 
+         (WP e @ E1 {{ Φ }}) (WP e @ E2 {{ v, |={E2,E1}=> Φ v }})%I | 100.
+Proof. intros. by rewrite /ElimVs pvs_frame_r wand_elim_r wp_atomic. Qed.
 End weakestpre.
diff --git a/tests/barrier_client.v b/tests/barrier_client.v
index ae0d5df08505a7e75841cb1f226ebb37d937ce72..e422df09b6d93e80c9e5bc2769df632e9af866d5 100644
--- a/tests/barrier_client.v
+++ b/tests/barrier_client.v
@@ -14,10 +14,9 @@ Definition client : expr :=
 Global Opaque worker client.
 
 Section client.
-  Context {Σ : gFunctors} `{!heapG Σ, !barrierG Σ, !spawnG Σ} (N : namespace).
-  Local Notation iProp := (iPropG heap_lang Σ).
+  Context `{!heapG Σ, !barrierG Σ, !spawnG Σ} (N : namespace).
 
-  Definition y_inv (q : Qp) (l : loc) : iProp :=
+  Definition y_inv (q : Qp) (l : loc) : iProp Σ :=
     (∃ f : val, l ↦{q} f ★ □ ∀ n : Z, WP f #n {{ v, v = #(n + 42) }})%I.
 
   Lemma y_inv_split q l : y_inv q l ⊢ (y_inv (q/2) l ★ y_inv (q/2) l).
@@ -50,13 +49,14 @@ Section client.
       iSplitL; [|by iIntros (_ _) "_ >"].
       iDestruct (recv_weaken with "[] Hr") as "Hr".
       { iIntros "Hy". by iApply (y_inv_split with "Hy"). }
-      iPvs (recv_split with "Hr") as "[H1 H2]"; first done.
+      iVs (recv_split with "Hr") as "[H1 H2]"; first done.
       iApply (wp_par (λ _, True%I) (λ _, True%I)). iFrame "Hh".
       iSplitL "H1"; [|iSplitL "H2"; [|by iIntros (_ _) "_ >"]];
         iApply worker_safe; by iSplit.
 Qed.
 End client.
 
+(*
 Section ClosedProofs.
   Definition Σ : gFunctors := #[ heapGF ; barrierGF ; spawnGF ].
   Notation iProp := (iPropG heap_lang Σ).
@@ -64,9 +64,10 @@ Section ClosedProofs.
   Lemma client_safe_closed σ : {{ ownP σ : iProp }} client {{ v, True }}.
   Proof.
     iIntros "! Hσ".
-    iPvs (heap_alloc with "Hσ") as (h) "[#Hh _]"; first done.
+    iVs (heap_alloc with "Hσ") as (h) "[#Hh _]"; first done.
     iApply (client_safe (nroot .@ "barrier")); auto with ndisj.
   Qed.
 
   Print Assumptions client_safe_closed.
 End ClosedProofs.
+*)
\ No newline at end of file
diff --git a/tests/heap_lang.v b/tests/heap_lang.v
index b259757ffa4bd204caa68a5f4ef2a24e58138166..932aac1d352e38b0c45df8482fde182d70172d4f 100644
--- a/tests/heap_lang.v
+++ b/tests/heap_lang.v
@@ -1,25 +1,23 @@
 (** This file is essentially a bunch of testcases. *)
 From iris.program_logic Require Import ownership hoare auth.
 From iris.heap_lang Require Import proofmode notation.
-Import uPred.
 
 Section LangTests.
-  Definition add : expr := (#21 + #21)%E.
+  Definition add : expr := #21 + #21.
   Goal ∀ σ, head_step add σ (#42) σ None.
   Proof. intros; do_head_step done. Qed.
-  Definition rec_app : expr := ((rec: "f" "x" := "f" "x") #0)%E.
+  Definition rec_app : expr := (rec: "f" "x" := "f" "x") #0.
   Goal ∀ σ, head_step rec_app σ rec_app σ None.
   Proof. intros. rewrite /rec_app. do_head_step done. Qed.
-  Definition lam : expr := (λ: "x", "x" + #21)%E.
+  Definition lam : expr := λ: "x", "x" + #21.
   Goal ∀ σ, head_step (lam #21)%E σ add σ None.
   Proof. intros. rewrite /lam. do_head_step done. Qed.
 End LangTests.
 
 Section LiftingTests.
   Context `{heapG Σ}.
-  Local Notation iProp := (iPropG heap_lang Σ).
-  Implicit Types P Q : iPropG heap_lang Σ.
-  Implicit Types Φ : val → iPropG heap_lang Σ.
+  Implicit Types P Q : iProp Σ.
+  Implicit Types Φ : val → iProp Σ.
 
   Definition heap_e  : expr :=
     let: "x" := ref #1 in "x" <- !"x" + #1 ;; !"x".
@@ -71,18 +69,19 @@ Section LiftingTests.
   Qed.
 
   Lemma Pred_user E :
-    (True : iProp) ⊢ WP let: "x" := Pred #42 in Pred "x" @ E {{ v, v = #40 }}.
+    True ⊢ WP let: "x" := Pred #42 in Pred "x" @ E {{ v, v = #40 }}.
   Proof. iIntros "". wp_apply Pred_spec. wp_let. by wp_apply Pred_spec. Qed.
 End LiftingTests.
 
+(*
 Section ClosedProofs.
-  Definition Σ : gFunctors := #[ heapGF ].
-  Notation iProp := (iPropG heap_lang Σ).
+  Definition Σ : gFunctors := #[ irisPreGF; heapGF ].
 
-  Lemma heap_e_closed σ : {{ ownP σ : iProp }} heap_e {{ v, v = #2 }}.
+  Lemma heap_e_closed σ : {{ ownP σ : iProp Σ }} heap_e {{ v, v = #2 }}.
   Proof.
     iProof. iIntros "! Hσ".
-    iPvs (heap_alloc with "Hσ") as (h) "[? _]"; first solve_ndisj.
+    iVs (heap_alloc with "Hσ") as (h) "[? _]"; first solve_ndisj.
     by iApply heap_e_spec; first solve_ndisj.
   Qed.
 End ClosedProofs.
+*)
\ No newline at end of file
diff --git a/tests/joining_existentials.v b/tests/joining_existentials.v
index 97f37f828f7bf005268b7053cab3a29b267f0cb1..e5d2f1f4e5381ac04797db34d7444a546f1de6bb 100644
--- a/tests/joining_existentials.v
+++ b/tests/joining_existentials.v
@@ -1,21 +1,20 @@
-From iris.algebra Require Import csum.
+From iris.algebra Require Import excl agree csum.
 From iris.program_logic Require Import hoare.
 From iris.heap_lang.lib.barrier Require Import proof specification.
 From iris.heap_lang Require Import notation par proofmode.
 From iris.proofmode Require Import invariants.
-Import uPred.
 
-Definition one_shotR (Λ : language) (Σ : gFunctors) (F : cFunctor) :=
-  csumR (exclR unitC) (agreeR $ laterC $ F (iPrePropG Λ Σ)).
-Definition Pending {Λ Σ F} : one_shotR Λ Σ F := Cinl (Excl ()).
-Definition Shot {Λ Σ} {F : cFunctor} (x : F (iPropG Λ Σ)) : one_shotR Λ Σ F :=
+Definition one_shotR (Σ : gFunctors) (F : cFunctor) :=
+  csumR (exclR unitC) (agreeR $ laterC $ F (iPreProp Σ)).
+Definition Pending {Σ F} : one_shotR Σ F := Cinl (Excl ()).
+Definition Shot {Σ} {F : cFunctor} (x : F (iProp Σ)) : one_shotR Σ F :=
   Cinr $ to_agree $ Next $ cFunctor_map F (iProp_fold, iProp_unfold) x.
 
-Class oneShotG (Λ : language) (Σ : gFunctors) (F : cFunctor) :=
-  one_shot_inG :> inG Λ Σ (one_shotR Λ Σ F).
+Class oneShotG (Σ : gFunctors) (F : cFunctor) :=
+  one_shot_inG :> inG Σ (one_shotR Σ F).
 Definition oneShotGF (F : cFunctor) : gFunctor :=
   GFunctor (csumRF (exclRF unitC) (agreeRF (â–¶ F))).
-Instance inGF_oneShotG  `{inGF Λ Σ (oneShotGF F)} : oneShotG Λ Σ F.
+Instance inGF_oneShotG  `{inGF Σ (oneShotGF F)} : oneShotG Σ F.
 Proof. apply: inGF_inG. Qed.
 
 Definition client eM eW1 eW2 : expr :=
@@ -24,15 +23,14 @@ Definition client eM eW1 eW2 : expr :=
 Global Opaque client.
 
 Section proof.
-Context `{!heapG Σ, !barrierG Σ, !spawnG Σ, !oneShotG heap_lang Σ F}.
+Context `{!heapG Σ, !barrierG Σ, !spawnG Σ, !oneShotG Σ F}.
 Context (N : namespace).
-Local Notation iProp := (iPropG heap_lang Σ).
-Local Notation X := (F iProp).
+Local Notation X := (F (iProp Σ)).
 
-Definition barrier_res γ (Φ : X → iProp) : iProp :=
+Definition barrier_res γ (Φ : X → iProp Σ) : iProp Σ :=
   (∃ x, own γ (Shot x) ★ Φ x)%I.
 
-Lemma worker_spec e γ l (Φ Ψ : X → iProp) `{!Closed [] e} :
+Lemma worker_spec e γ l (Φ Ψ : X → iProp Σ) `{!Closed [] e} :
   recv N l (barrier_res γ Φ) ★ (∀ x, {{ Φ x }} e {{ _, Ψ x }})
   ⊢ WP wait #l ;; e {{ _, barrier_res γ Ψ }}.
 Proof.
@@ -42,7 +40,7 @@ Proof.
   iIntros (v) "?"; iExists x; by iSplit.
 Qed.
 
-Context (P : iProp) (Φ Φ1 Φ2 Ψ Ψ1 Ψ2 : X -n> iProp).
+Context (P : iProp Σ) (Φ Φ1 Φ2 Ψ Ψ1 Ψ2 : X -n> iProp Σ).
 Context {Φ_split : ∀ x, Φ x ⊢ (Φ1 x ★ Φ2 x)}.
 Context {Ψ_join  : ∀ x, (Ψ1 x ★ Ψ2 x) ⊢ Ψ x}.
 
@@ -58,7 +56,7 @@ Proof.
   iDestruct "Hγ" as (x) "[#Hγ Hx]"; iDestruct "Hγ'" as (x') "[#Hγ' Hx']".
   iAssert (▷ (x ≡ x'))%I as "Hxx" .
   { iCombine "Hγ" "Hγ'" as "Hγ2". iClear "Hγ Hγ'".
-    rewrite own_valid csum_validI /= agree_validI agree_equivI later_equivI /=.
+    rewrite own_valid csum_validI /= agree_validI agree_equivI uPred.later_equivI /=.
     rewrite -{2}[x]cFunctor_id -{2}[x']cFunctor_id.
     rewrite (ne_proper (cFunctor_map F) (cid, cid) (_ â—Ž _, _ â—Ž _)); last first.
     { by split; intro; simpl; symmetry; apply iProp_fold_unfold. }
@@ -76,7 +74,7 @@ Lemma client_spec_new eM eW1 eW2 `{!Closed [] eM, !Closed [] eW1, !Closed [] eW2
   ⊢ WP client eM eW1 eW2 {{ _, ∃ γ, barrier_res γ Ψ }}.
 Proof.
   iIntros (HN) "/= (#Hh&HP&#He&#He1&#He2)"; rewrite /client.
-  iPvs (own_alloc (Pending : one_shotR heap_lang Σ F)) as (γ) "Hγ". done.
+  iVs (own_alloc (Pending : one_shotR Σ F)) as (γ) "Hγ"; first done.
   wp_apply (newbarrier_spec N (barrier_res γ Φ)); auto.
   iFrame "Hh". iIntros (l) "[Hr Hs]".
   set (workers_post (v : val) := (barrier_res γ Ψ1 ★ barrier_res γ Ψ2)%I).
@@ -84,12 +82,12 @@ Proof.
   iSplitL "HP Hs Hγ"; [|iSplitL "Hr"].
   - wp_focus eM. iApply wp_wand_l; iSplitR "HP"; [|by iApply "He"].
     iIntros (v) "HP"; iDestruct "HP" as (x) "HP". wp_let.
-    iPvs (own_update with "Hγ") as "Hx".
+    iVs (own_update with "Hγ") as "Hx".
     { by apply (cmra_update_exclusive (Shot x)). }
     iApply signal_spec; iFrame "Hs"; iSplit; last done.
     iExists x; auto.
   - iDestruct (recv_weaken with "[] Hr") as "Hr"; first by iApply P_res_split.
-    iPvs (recv_split with "Hr") as "[H1 H2]"; first done.
+    iVs (recv_split with "Hr") as "[H1 H2]"; first done.
     wp_apply (wp_par (λ _, barrier_res γ Ψ1)%I
                      (λ _, barrier_res γ Ψ2)%I); iFrame "Hh".
     iSplitL "H1"; [|iSplitL "H2"].
diff --git a/tests/list_reverse.v b/tests/list_reverse.v
index 3655049b00c50949945f3965b84e13137b8110d2..dafe2070bee7fca57a914073eb8719f472b7eee1 100644
--- a/tests/list_reverse.v
+++ b/tests/list_reverse.v
@@ -5,10 +5,9 @@ From iris.heap_lang Require Import proofmode notation.
 
 Section list_reverse.
 Context `{!heapG Σ}.
-Notation iProp := (iPropG heap_lang Σ).
 Implicit Types l : loc.
 
-Fixpoint is_list (hd : val) (xs : list val) : iProp :=
+Fixpoint is_list (hd : val) (xs : list val) : iProp Σ :=
   match xs with
   | [] => hd = NONEV
   | x :: xs => ∃ l hd', hd = SOMEV #l ★ l ↦ (x,hd') ★ is_list hd' xs
@@ -26,7 +25,7 @@ Definition rev : val :=
     end.
 Global Opaque rev.
 
-Lemma rev_acc_wp hd acc xs ys (Φ : val → iProp) :
+Lemma rev_acc_wp hd acc xs ys (Φ : val → iProp Σ) :
   heap_ctx ★ is_list hd xs ★ is_list acc ys ★
     (∀ w, is_list w (reverse xs ++ ys) -★ Φ w)
   ⊢ WP rev hd acc {{ Φ }}.
@@ -42,7 +41,7 @@ Proof.
     iIntros (w). rewrite cons_middle assoc -reverse_cons. iApply "HΦ".
 Qed.
 
-Lemma rev_wp hd xs (Φ : val → iProp) :
+Lemma rev_wp hd xs (Φ : val → iProp Σ) :
   heap_ctx ★ is_list hd xs ★ (∀ w, is_list w (reverse xs) -★ Φ w)
   ⊢ WP rev hd (InjL #()) {{ Φ }}.
 Proof.
diff --git a/tests/one_shot.v b/tests/one_shot.v
index 66479873d92a06b8be0450c81d1574cb3d08357e..9dbb7194fad5e94cb23b2f74ae5f9b5a551b8107 100644
--- a/tests/one_shot.v
+++ b/tests/one_shot.v
@@ -1,8 +1,7 @@
-From iris.algebra Require Import dec_agree csum.
+From iris.algebra Require Import excl dec_agree csum.
 From iris.program_logic Require Import hoare.
 From iris.heap_lang Require Import assert proofmode notation.
 From iris.proofmode Require Import invariants ghost_ownership.
-Import uPred.
 
 Definition one_shot_example : val := λ: <>,
   let: "x" := ref NONE in (
@@ -24,19 +23,18 @@ Definition one_shotR := csumR (exclR unitC) (dec_agreeR Z).
 Definition Pending : one_shotR := (Cinl (Excl ()) : one_shotR).
 Definition Shot (n : Z) : one_shotR := (Cinr (DecAgree n) : one_shotR).
 
-Class one_shotG Σ := one_shot_inG :> inG heap_lang Σ one_shotR.
+Class one_shotG Σ := one_shot_inG :> inG Σ one_shotR.
 Definition one_shotGF : gFunctorList := [GFunctor (constRF one_shotR)].
-Instance inGF_one_shotG Σ : inGFs heap_lang Σ one_shotGF → one_shotG Σ.
+Instance inGF_one_shotG Σ : inGFs Σ one_shotGF → one_shotG Σ.
 Proof. intros [? _]; apply: inGF_inG. Qed.
 
 Section proof.
 Context `{!heapG Σ, !one_shotG Σ} (N : namespace) (HN : heapN ⊥ N).
-Local Notation iProp := (iPropG heap_lang Σ).
 
-Definition one_shot_inv (γ : gname) (l : loc) : iProp :=
+Definition one_shot_inv (γ : gname) (l : loc) : iProp Σ :=
   (l ↦ NONEV ★ own γ Pending ∨ ∃ n : Z, l ↦ SOMEV #n ★ own γ (Shot n))%I.
 
-Lemma wp_one_shot (Φ : val → iProp) :
+Lemma wp_one_shot (Φ : val → iProp Σ) :
   heap_ctx ★ (∀ f1 f2 : val,
     (∀ n : Z, □ WP f1 #n {{ w, w = #true ∨ w = #false }}) ★
     □ WP f2 #() {{ g, □ WP g #() {{ _, True }} }} -★ Φ (f1,f2)%V)
@@ -44,43 +42,49 @@ Lemma wp_one_shot (Φ : val → iProp) :
 Proof.
   iIntros "[#? Hf] /=".
   rewrite /one_shot_example. wp_seq. wp_alloc l as "Hl". wp_let.
-  iPvs (own_alloc Pending) as (γ) "Hγ"; first done.
-  iPvs (inv_alloc N _ (one_shot_inv γ l) with "[Hl Hγ]") as "#HN"; first done.
+  iVs (own_alloc Pending) as (γ) "Hγ"; first done.
+  iVs (inv_alloc N _ (one_shot_inv γ l) with "[Hl Hγ]") as "#HN".
   { iNext. iLeft. by iSplitL "Hl". }
-  iPvsIntro. iApply "Hf"; iSplit.
+  iVsIntro. iApply "Hf"; iSplit.
   - iIntros (n) "!". wp_let.
-    iInv> N as "[[Hl Hγ]|H]"; last iDestruct "H" as (m) "[Hl Hγ]".
-    + wp_cas_suc. iSplitL; [|by iLeft].
-      iPvs (own_update with "Hγ") as "Hγ".
+    iInv N as "H" "Hclose"; iTimeless "H".
+    iDestruct "H" as "[[Hl Hγ]|H]"; last iDestruct "H" as (m) "[Hl Hγ]".
+    + wp_cas_suc. iVs (own_update with "Hγ") as "Hγ".
       { by apply cmra_update_exclusive with (y:=Shot n). }
-      iPvsIntro; iRight; iExists n; by iSplitL "Hl".
-    + wp_cas_fail. rewrite /one_shot_inv; eauto 10.
-  - iIntros "!". wp_seq. wp_focus (! _)%E. iInv> N as "Hγ".
+      iVs ("Hclose" with "[-]"); last eauto.
+      iNext; iRight; iExists n; by iFrame.
+    + wp_cas_fail. iVs ("Hclose" with "[-]"); last eauto.
+      rewrite /one_shot_inv; eauto 10.
+  - iIntros "!". wp_seq. wp_focus (! _)%E.
+    iInv N as "Hγ" "Hclose"; iTimeless "Hγ".
     iAssert (∃ v, l ↦ v ★ ((v = NONEV ★ own γ Pending) ∨
-       ∃ n : Z, v = SOMEV #n ★ own γ (Shot n)))%I with "[-]" as "Hv".
+       ∃ n : Z, v = SOMEV #n ★ own γ (Shot n)))%I with "[Hγ]" as "Hv".
     { iDestruct "Hγ" as "[[Hl Hγ]|Hl]"; last iDestruct "Hl" as (m) "[Hl Hγ]".
       + iExists NONEV. iFrame. eauto.
       + iExists (SOMEV #m). iFrame. eauto. }
-    iDestruct "Hv" as (v) "[Hl Hv]". wp_load; iPvsIntro.
+    iDestruct "Hv" as (v) "[Hl Hv]". wp_load.
     iAssert (one_shot_inv γ l ★ (v = NONEV ∨ ∃ n : Z,
-      v = SOMEV #n ★ own γ (Shot n)))%I with "[-]" as "[$ #Hv]".
+      v = SOMEV #n ★ own γ (Shot n)))%I with "[Hl Hv]" as "[Hinv #Hv]".
     { iDestruct "Hv" as "[[% ?]|Hv]"; last iDestruct "Hv" as (m) "[% ?]"; subst.
       + iSplit. iLeft; by iSplitL "Hl". eauto.
       + iSplit. iRight; iExists m; by iSplitL "Hl". eauto. }
-    wp_let. iPvsIntro. iIntros "!". wp_seq.
+    iVs ("Hclose" with "[Hinv]") as "_"; eauto; iVsIntro.
+    wp_let. iVsIntro. iIntros "!". wp_seq.
     iDestruct "Hv" as "[%|Hv]"; last iDestruct "Hv" as (m) "[% Hγ']"; subst.
     { by wp_match. }
     wp_match. wp_focus (! _)%E.
-    iInv> N as "[[Hl Hγ]|Hinv]"; last iDestruct "Hinv" as (m') "[Hl Hγ]".
+    iInv N as "H" "Hclose"; iTimeless "H".
+    iDestruct "H" as "[[Hl Hγ]|H]"; last iDestruct "H" as (m') "[Hl Hγ]".
     { iCombine "Hγ" "Hγ'" as "Hγ". by iDestruct (own_valid with "Hγ") as %?. }
-    wp_load; iPvsIntro.
+    wp_load.
     iCombine "Hγ" "Hγ'" as "Hγ".
     iDestruct (own_valid with "#Hγ") as %[=->]%dec_agree_op_inv.
-    iSplitL "Hl"; [iRight; by eauto|].
-    wp_match. iApply wp_assert. wp_op=>?; simplify_eq/=; eauto.
+    iVs ("Hclose" with "[Hl]") as "_".
+    { iNext; iRight; by eauto. }
+    iVsIntro. wp_match. iApply wp_assert. wp_op=>?; simplify_eq/=; eauto.
 Qed.
 
-Lemma hoare_one_shot (Φ : val → iProp) :
+Lemma ht_one_shot (Φ : val → iProp Σ) :
   heap_ctx ⊢ {{ True }} one_shot_example #()
     {{ ff,
       (∀ n : Z, {{ True }} Fst ff #n {{ w, w = #true ∨ w = #false }}) ★
diff --git a/tests/program_logic.v b/tests/program_logic.v
deleted file mode 100644
index f9f66bfd748528b32b98193ecb2553a7b53fd01e..0000000000000000000000000000000000000000
--- a/tests/program_logic.v
+++ /dev/null
@@ -1,22 +0,0 @@
-(** This file tests a bunch of things. *)
-From iris.program_logic Require Import model saved_prop.
-
-Module ModelTest. (* Make sure we got the notations right. *)
-  Definition iResTest {Λ : language} {Σ : iFunctor}
-    (w : iWld Λ Σ) (p : iPst Λ) (g : iGst Λ Σ) : iRes Λ Σ := Res w p g.
-End ModelTest.
-
-Module SavedPropTest.
-  (* Test if we can really go "crazy higher order" *)
-  Section sec.
-    Definition F : cFunctor := ( ∙ -n> ∙ ).
-    Definition Σ : gFunctors := #[ savedPropGF F ].
-    Context {Λ : language}.
-    Notation iProp := (iPropG Λ Σ).
-
-    Local Instance : savedPropG Λ Σ F := _.
-
-    Definition own_pred γ (φ : iProp -n> iProp) : iProp :=
-      saved_prop_own γ φ.
-  End sec.
-End SavedPropTest.
diff --git a/tests/proofmode.v b/tests/proofmode.v
index 7ac9b89561e967ad0703eb8cb86ee56e2c76f4f0..f8d2711313a342c13daa5b0f6f99ff5b4b24e108 100644
--- a/tests/proofmode.v
+++ b/tests/proofmode.v
@@ -82,20 +82,11 @@ Proof.
 Qed.
 
 Section iris.
-  Context {Λ : language} {Σ : iFunctor}.
+  Context `{irisG Λ Σ}.
   Implicit Types E : coPset.
-  Implicit Types P Q : iProp Λ Σ.
+  Implicit Types P Q : iProp Σ.
 
-  Lemma demo_7 E1 E2 E P :
-    E1 ⊆ E2 → E ⊆ E1 →
-    (|={E1,E}=> ▷ P) ={E2,E ∪ E2 ∖ E1}=> ▷ P.
-  Proof.
-    iIntros (? ?) "Hpvs".
-    iPvs "Hpvs"; first set_solver.
-    done.
-  Qed.
-
-  Lemma demo_8 N E P Q R :
+  Lemma demo_7 N E P Q R :
     nclose N ⊆ E →
     (True -★ P -★ inv N Q -★ True -★ R) ⊢ P -★ ▷ Q ={E}=★ R.
   Proof.
diff --git a/tests/tree_sum.v b/tests/tree_sum.v
index 6d07b098cde1be59189a4e083999129f6db43a54..d188e64b6549a5b16602d635edcb59e9c11e7c0d 100644
--- a/tests/tree_sum.v
+++ b/tests/tree_sum.v
@@ -5,7 +5,7 @@ Inductive tree :=
   | leaf : Z → tree
   | node : tree → tree → tree.
 
-Fixpoint is_tree `{!heapG Σ} (v : val) (t : tree) : iPropG heap_lang Σ :=
+Fixpoint is_tree `{!heapG Σ} (v : val) (t : tree) : iProp Σ :=
   match t with
   | leaf n => v = InjLV #n
   | node tl tr =>
@@ -33,7 +33,7 @@ Definition sum' : val := λ: "t",
 
 Global Opaque sum_loop sum'.
 
-Lemma sum_loop_wp `{!heapG Σ} v t l (n : Z) (Φ : val → iPropG heap_lang Σ) :
+Lemma sum_loop_wp `{!heapG Σ} v t l (n : Z) (Φ : val → iProp Σ) :
   heap_ctx ★ l ↦ #n ★ is_tree v t
     ★ (l ↦ #(sum t + n) -★ is_tree v t -★ Φ #())
   ⊢ WP sum_loop v #l {{ Φ }}.