diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml
index f24ff612b2834e2e4593af3fe8fdaa2eb5b5359a..78e1f98a2f0ee1a1dcf80500db406dcf09177723 100644
--- a/.gitlab-ci.yml
+++ b/.gitlab-ci.yml
@@ -6,11 +6,11 @@ buildjob:
   script:
   - coqc -v
   - 'time make -j8 TIMED=y 2>&1 | tee build-log.txt'
-  - 'grep Axiom build-log.txt && exit 1'
+  - 'fgrep Axiom build-log.txt && exit 1'
   - 'cat build-log.txt  | egrep "[a-zA-Z0-9_/-]+ \(user: [0-9]" | tee build-time.txt'
+  - make validate
   only:
   - master
-  - jh_simplified_resources
   artifacts:
     paths:
     - build-time.txt
diff --git a/CHANGELOG.md b/CHANGELOG.md
index 3853fde1690ea8e268ab7956ecda3c129f0d6197..7dfe2cbeb0d509d4daf856523d33fff75fdbfbfe 100644
--- a/CHANGELOG.md
+++ b/CHANGELOG.md
@@ -1,19 +1,38 @@
 In this changelog, we document "large-ish" changes to Iris that affect even the
 way the logic is used on paper.  We also mention some significant changes in the
 Coq development, but not every API-breaking change is listed.  Changes marked
-[#] still need to be ported to the Iris Documentation LaTeX file.
+[#] still need to be ported to the Iris Documentation LaTeX file(s).
 
-## Iris 2.0
+## Iris 3.0
+
+* [#] View shifts are radically simplified to just internalize frame-preserving
+  updates.  Weakestpre is defined inside the logic, and invariants and view
+  shifts with masks are also coded up inside Iris.  Adequacy of weakestpre
+  is proven in the logic.
+* [#] With invariants and the physical state being handled in the logic, there
+  is no longer any reason to demand the CMRA unit to be discrete.
+* [#] The language can now fork off multiple threads at once.
 
-This version accompanies the final ICFP paper.
+## Iris 2.0
 
-* [# algebra] Make the core of an RA or CMRA a partial function.
 * [heap_lang] No longer use dependent types for expressions.  Instead, values
   carry a proof of closedness.  Substitution, closedness and value-ness proofs
   are performed by computation after reflecting into a term langauge that knows
   about values and closed expressions.
 * [program_logic/language] The language does not define its own "atomic"
   predicate.  Instead, atomicity is defined as reducing in one step to a value.
+* [program_logic] Due to a lack of maintenance and usefulness, lifting lemmas
+  for Hoare triples are removed.
+
+## Iris 2.0-rc2
+
+This version matches the final ICFP paper.
+
+* [algebra] Make the core of an RA or CMRA a partial function.
+* [program_logic/lifting] Lifting lemmas no longer round-trip through a
+  user-chosen predicate to define the configurations we can reduce to; they
+  directly relate to the operational semantics.  This is equivalent and
+  much simpler to read.
 
 ## Iris 2.0-rc1
 
diff --git a/Makefile b/Makefile
index 6c554b2751754c364a05e7cad9ad4634c68da0a4..ff1ccd39b120b9d1ff5e72d079d8484bfbfb551b 100644
--- a/Makefile
+++ b/Makefile
@@ -1,6 +1,6 @@
 # Makefile originally taken from coq-club
 
-%: Makefile.coq
+%: Makefile.coq phony
 	+make -f Makefile.coq $@
 
 all: Makefile.coq
@@ -17,4 +17,6 @@ _CoqProject: ;
 
 Makefile: ;
 
-.PHONY: all clean
+phony: ;
+
+.PHONY: all clean phony
diff --git a/ProofMode.md b/ProofMode.md
index 5a086022abc189210c9db55f233ac297e6876661..59846cef5fe77688ae9fb15a7bda9dc511c72212 100644
--- a/ProofMode.md
+++ b/ProofMode.md
@@ -1,13 +1,20 @@
 Tactic overview
 ===============
 
+Many of the tactics below apply to more goals than described in this document
+since the behavior of these tactics can be tuned via instances of the type
+classes in the file `proofmode/classes`. Most notable, many of the tactics can
+be applied when the to be introduced or to be eliminated connective appears
+under a later, a primitive view shift, or in the conclusion of a weakest
+precondition connective.
+
 Applying hypotheses and lemmas
 ------------------------------
 
 - `iExact "H"`  : finish the goal if the conclusion matches the hypothesis `H`
 - `iAssumption` : finish the goal if the conclusion matches any hypothesis
-- `iApply trm` : match the conclusion of the current goal against the
-   conclusion of `tetrmrm` and generates goals for the premises of `trm`. See
+- `iApply pm_trm` : match the conclusion of the current goal against the
+   conclusion of `pm_trm` and generates goals for the premises of `pm_trm`. See
    proof mode terms below.
 
 Context management
@@ -23,9 +30,10 @@ Context management
   `x1 ... xn` into universal quantifiers. The symbol `★` can be used to revert
   the entire spatial context.
 - `iRename "H1" into "H2"` : rename the hypothesis `H1` into `H2`.
-- `iSpecialize trm` : instantiate universal quantifiers and eliminate
-  implications/wands of a hypothesis `trm`. See proof mode terms below.
-- `iPoseProof trm as "H"` : put `trm` into the context as a new hypothesis `H`.
+- `iSpecialize pm_trm` : instantiate universal quantifiers and eliminate
+  implications/wands of a hypothesis `pm_trm`. See proof mode terms below.
+- `iPoseProof pm_trm as "H"` : put `pm_trm` into the context as a new hypothesis
+  `H`.
 - `iAssert P with "spat" as "ipat"` : create a new goal with conclusion `P` and
   put `P` in the context of the original goal. The specialization pattern
   `spat` specifies which hypotheses will be consumed by proving `P` and the
@@ -52,17 +60,21 @@ Elimination of logical connectives
 ----------------------------------
 
 - `iExFalso` : Ex falso sequitur quod libet.
-- `iDestruct trm as (x1 ... xn) "spat1 ... spatn"` : elimination of existential
-  quantifiers using Coq introduction patterns `x1 ... xn` and elimination of
-  object level connectives using the proof mode introduction patterns
-  `ipat1 ... ipatn`.
-- `iDestruct trm as %cpat` : elimination of a pure hypothesis using the Coq
+- `iDestruct pm_trm as (x1 ... xn) "spat1 ... spatn"` : elimination of
+  existential quantifiers using Coq introduction patterns `x1 ... xn` and
+  elimination of object level connectives using the proof mode introduction
+  patterns `ipat1 ... ipatn`.
+- `iDestruct pm_trm as %cpat` : elimination of a pure hypothesis using the Coq
   introduction pattern `cpat`.
 
 Separating logic specific tactics
 ---------------------------------
 
-- `iFrame "H0 ... Hn"` : cancel the hypotheses `H0 ... Hn` in the goal. 
+- `iFrame "H0 ... Hn"` : cancel the hypotheses `H0 ... Hn` in the goal. The
+  symbol `★` can be used to frame as much of the spatial context as possible,
+  and the symbol `#` can be used to repeatedly frame as much of the persistent
+  context as possible. When without arguments, it attempts to frame all spatial
+  hypotheses.
 - `iCombine "H1" "H2" as "H"` : turns `H1 : P1` and `H2 : P2` into
   `H : P1 ★ P2`.
 
@@ -75,20 +87,20 @@ The later modality
 Rewriting
 ---------
 
-- `iRewrite trm` : rewrite an equality in the conclusion.
-- `iRewrite trm in "H"` : rewrite an equality in the hypothesis `H`.
+- `iRewrite pm_trm` : rewrite an equality in the conclusion.
+- `iRewrite pm_trm in "H"` : rewrite an equality in the hypothesis `H`.
 
 Iris
 ----
 
-- `iPvsIntro` : introduction of a primitive view shift. Generates a goal if
-  the masks are not syntactically equal.
-- `iPvs trm as (x1 ... xn) "ipat"` : runs a primitive view shift `trm`.
+- `iVsIntro` : introduction of a raw or primitive view shift.
+- `iVs pm_trm as (x1 ... xn) "ipat"` : run a raw or primitive view shift
+  `pm_trm` (if the goal permits, i.e. it is a raw or primitive view shift, or
+   a weakest precondition).
 - `iInv N as (x1 ... xn) "ipat"` : open the invariant `N`.
-- `iInv> N as (x1 ... xn) "ipat"` : open the invariant `N` and establish that
-  it is timeless so no laters have to be added.
-- `iTimeless "H"` : strip a later of a timeless hypotheses `H` in case the
-  conclusion is a primitive view shifts or weakest precondition.
+- `iTimeless "H"` : strip a later of a timeless hypothesis `H` (if the goal
+   permits, i.e. it is a later, True now, raw or primitive view shift, or a
+   weakest precondition).
 
 Miscellaneous
 -------------
@@ -111,20 +123,24 @@ introduction patterns:
 - `?` : create an anonymous hypothesis.
 - `_` : remove the hypothesis.
 - `$` : frame the hypothesis in the goal.
-- `# ipat` : move the hypothesis to the persistent context.
-- `%` : move the hypothesis to the pure Coq context (anonymously).
 - `[ipat ipat]` : (separating) conjunction elimination.
 - `[ipat|ipat]` : disjunction elimination.
 - `[]` : false elimination.
+- `%` : move the hypothesis to the pure Coq context (anonymously).
+- `# ipat` : move the hypothesis to the persistent context.
+- `> ipat` : remove a later of a timeless hypothesis (if the goal permits).
+- `==> ipat` : run a view shift (if the goal permits).
 
 Apart from this, there are the following introduction patterns that can only
 appear at the top level:
 
-- `!` : introduce a box (provided that the spatial context is empty).
-- `>` : introduce a later (which strips laters from all hypotheses).
 - `{H1 ... Hn}` : clear `H1 ... Hn`.
 - `{$H1 ... $Hn}` : frame `H1 ... Hn` (this pattern can be mixed with the
   previous pattern, e.g., `{$H1 H2 $H3}`).
+- `!%` : introduce a pure goal (and leave the proof mode).
+- `!#` : introduce an always modality (given that the spatial context is empty).
+- `!>` : introduce a later (which strips laters from all hypotheses).
+- `!==>` : introduce a view shift.
 - `/=` : perform `simpl`.
 - `*` : introduce all universal quantifiers.
 - `**` : introduce all universal quantifiers, as well as all arrows and wands.
@@ -135,7 +151,7 @@ For example, given:
 
 You can write
 
-        iIntros (x) "% ! $ [[] | #[HQ HR]] /= >".
+        iIntros (x) "% !# $ [[] | #[HQ HR]] /= !>".
 
 which results in:
 
@@ -161,7 +177,7 @@ so called specification patterns to express this splitting:
 - `[H1 ... Hn]` : generate a goal with the spatial hypotheses `H1 ... Hn` and
   all persistent hypotheses. The hypotheses `H1 ... Hn` will be consumed.
 - `[-H1 ... Hn]`  : negated form of the above pattern
-- `=>[H1 ... Hn]` : same as the above pattern, but can only be used if the goal
+- `==>[H1 ... Hn]` : same as the above pattern, but can only be used if the goal
   is a primitive view shift, in which case the view shift will be kept in the
   goal of the premise too.
 - `[#]` : This pattern can be used when eliminating `P -★ Q` when either `P` or
@@ -186,7 +202,7 @@ Many of the proof mode tactics (such as `iDestruct`, `iApply`, `iRewrite`) can
 take both hypotheses and lemmas, and allow one to instantiate universal
 quantifiers and implications/wands of these hypotheses/lemmas on the fly.
 
-The syntax for the arguments, called _proof mode terms_, of these tactics is:
+The syntax for the arguments of these tactics, called _proof mode terms_, is:
 
         (H $! t1 ... tn with "spat1 .. spatn")
 
diff --git a/_CoqProject b/_CoqProject
index 473926602eadb226fce69a2b79883589ea58f292..54a1a2d1594cda76579af944d8b8b6cc6dd24295 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -66,25 +66,25 @@ 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/viewshifts.v
 program_logic/language.v
 program_logic/ectx_language.v
 program_logic/ectxi_language.v
 program_logic/ectx_lifting.v
 program_logic/ghost_ownership.v
-program_logic/global_functor.v
 program_logic/saved_prop.v
 program_logic/auth.v
 program_logic/sts.v
 program_logic/namespaces.v
 program_logic/boxes.v
+program_logic/counter_examples.v
+program_logic/iris.v
+program_logic/thread_local.v
+program_logic/cancelable_invariants.v
 heap_lang/lang.v
 heap_lang/tactics.v
 heap_lang/wp_tactics.v
@@ -96,20 +96,24 @@ heap_lang/lib/spawn.v
 heap_lang/lib/par.v
 heap_lang/lib/assert.v
 heap_lang/lib/lock.v
+heap_lang/lib/spin_lock.v
+heap_lang/lib/ticket_lock.v
 heap_lang/lib/counter.v
 heap_lang/lib/barrier/barrier.v
 heap_lang/lib/barrier/specification.v
 heap_lang/lib/barrier/protocol.v
 heap_lang/lib/barrier/proof.v
 heap_lang/proofmode.v
+heap_lang/adequacy.v
+tests/atomic.v
 tests/heap_lang.v
-tests/program_logic.v
 tests/one_shot.v
 tests/joining_existentials.v
 tests/proofmode.v
 tests/barrier_client.v
 tests/list_reverse.v
 tests/tree_sum.v
+tests/counter.v
 proofmode/coq_tactics.v
 proofmode/pviewshifts.v
 proofmode/environments.v
@@ -120,5 +124,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/algebra/agree.v b/algebra/agree.v
index a85d59e8f795968981919624de76fdc98b47ae0c..e138d03685b32c584d179ad407dad57fa092325a 100644
--- a/algebra/agree.v
+++ b/algebra/agree.v
@@ -108,7 +108,7 @@ Proof.
     symmetry; apply dist_le with n; try apply Hx; auto.
   - intros x. apply agree_idemp.
   - by intros n x y [(?&?&?) ?].
-  - intros n x y1 y2 Hval Hx; exists (x,x); simpl; split.
+  - intros n x y1 y2 Hval Hx; exists x, x; simpl; split.
     + by rewrite agree_idemp.
     + by move: Hval; rewrite Hx; move=> /agree_op_inv->; rewrite agree_idemp.
 Qed.
diff --git a/algebra/auth.v b/algebra/auth.v
index d9fa09e152742fe889072b811e75fc6ffa7f44da..7d04ae801c93cb5a38f1e18e61e1980a117cb7ae 100644
--- a/algebra/auth.v
+++ b/algebra/auth.v
@@ -134,10 +134,10 @@ Proof.
      naive_solver eauto using cmra_validN_op_l, cmra_validN_includedN.
   - intros n x y1 y2 ? [??]; simpl in *.
     destruct (cmra_extend n (authoritative x) (authoritative y1)
-      (authoritative y2)) as (ea&?&?&?); auto using authoritative_validN.
+      (authoritative y2)) as (ea1&ea2&?&?&?); auto using authoritative_validN.
     destruct (cmra_extend n (auth_own x) (auth_own y1) (auth_own y2))
-      as (b&?&?&?); auto using auth_own_validN.
-    by exists (Auth (ea.1) (b.1), Auth (ea.2) (b.2)).
+      as (b1&b2&?&?&?); auto using auth_own_validN.
+    by exists (Auth ea1 b1), (Auth ea2 b2).
 Qed.
 Canonical Structure authR := CMRAT (auth A) auth_cofe_mixin auth_cmra_mixin.
 
@@ -156,7 +156,6 @@ Proof.
   split; simpl.
   - apply (@ucmra_unit_valid A).
   - by intros x; constructor; rewrite /= left_id.
-  - apply _.
   - do 2 constructor; simpl; apply (persistent_core _).
 Qed.
 Canonical Structure authUR :=
@@ -191,11 +190,17 @@ Proof.
   exists bf2. rewrite -assoc.
   apply (Hab' _ (Some _)); auto. by rewrite /= assoc.
 Qed.
+
 Lemma auth_update_no_frame a b : a ~l~> b @ Some ∅ → ● a ⋅ ◯ a ~~> ● b ⋅ ◯ b.
 Proof.
   intros. rewrite -{1}(right_id _ _ a) -{1}(right_id _ _ b).
   by apply auth_update.
 Qed.
+Lemma auth_update_no_frag af b : ∅ ~l~> b @ Some af → ● af ~~> ● (b ⋅ af) ⋅ ◯ b.
+Proof.
+  intros. rewrite -{1}(left_id _ _ af) -{1}(right_id _ _ (● _)).
+  by apply auth_update.
+Qed.
 End cmra.
 
 Arguments authR : clear implicits.
@@ -235,6 +240,28 @@ Definition authC_map {A B} (f : A -n> B) : authC A -n> authC B :=
 Lemma authC_map_ne A B n : Proper (dist n ==> dist n) (@authC_map A B).
 Proof. intros f f' Hf [[[a|]|] b]; repeat constructor; apply Hf. Qed.
 
+Program Definition authRF (F : urFunctor) : rFunctor := {|
+  rFunctor_car A B := authR (urFunctor_car F A B);
+  rFunctor_map A1 A2 B1 B2 fg := authC_map (urFunctor_map F fg)
+|}.
+Next Obligation.
+  by intros F A1 A2 B1 B2 n f g Hfg; apply authC_map_ne, urFunctor_ne.
+Qed.
+Next Obligation.
+  intros F A B x. rewrite /= -{2}(auth_map_id x).
+  apply auth_map_ext=>y; apply urFunctor_id.
+Qed.
+Next Obligation.
+  intros F A1 A2 A3 B1 B2 B3 f g f' g' x. rewrite /= -auth_map_compose.
+  apply auth_map_ext=>y; apply urFunctor_compose.
+Qed.
+
+Instance authRF_contractive F :
+  urFunctorContractive F → rFunctorContractive (authRF F).
+Proof.
+  by intros ? A1 A2 B1 B2 n f g Hfg; apply authC_map_ne, urFunctor_contractive.
+Qed.
+
 Program Definition authURF (F : urFunctor) : urFunctor := {|
   urFunctor_car A B := authUR (urFunctor_car F A B);
   urFunctor_map A1 A2 B1 B2 fg := authC_map (urFunctor_map F fg)
diff --git a/algebra/cmra.v b/algebra/cmra.v
index a388fc743e13b18b44abb721d60e912c754d373b..ef47985f98ce4f1ea717204f004ec33ede12886f 100644
--- a/algebra/cmra.v
+++ b/algebra/cmra.v
@@ -53,7 +53,7 @@ Record CMRAMixin A `{Dist A, Equiv A, PCore A, Op A, Valid A, ValidN A} := {
   mixin_cmra_validN_op_l n x y : ✓{n} (x ⋅ y) → ✓{n} x;
   mixin_cmra_extend n x y1 y2 :
     ✓{n} x → x ≡{n}≡ y1 ⋅ y2 →
-    { z | x ≡ z.1 ⋅ z.2 ∧ z.1 ≡{n}≡ y1 ∧ z.2 ≡{n}≡ y2 }
+    ∃ z1 z2, x ≡ z1 ⋅ z2 ∧ z1 ≡{n}≡ y1 ∧ z2 ≡{n}≡ y2
 }.
 
 (** Bundeled version *)
@@ -120,7 +120,7 @@ Section cmra_mixin.
   Proof. apply (mixin_cmra_validN_op_l _ (cmra_mixin A)). Qed.
   Lemma cmra_extend n x y1 y2 :
     ✓{n} x → x ≡{n}≡ y1 ⋅ y2 →
-    { z | x ≡ z.1 ⋅ z.2 ∧ z.1 ≡{n}≡ y1 ∧ z.2 ≡{n}≡ y2 }.
+    ∃ z1 z2, x ≡ z1 ⋅ z2 ∧ z1 ≡{n}≡ y1 ∧ z2 ≡{n}≡ y2.
   Proof. apply (mixin_cmra_extend _ (cmra_mixin A)). Qed.
 End cmra_mixin.
 
@@ -153,7 +153,6 @@ Arguments core' _ _ _ /.
 Record UCMRAMixin A `{Dist A, Equiv A, PCore A, Op A, Valid A, Empty A} := {
   mixin_ucmra_unit_valid : ✓ ∅;
   mixin_ucmra_unit_left_id : LeftId (≡) ∅ (⋅);
-  mixin_ucmra_unit_timeless : Timeless ∅;
   mixin_ucmra_pcore_unit : pcore ∅ ≡ Some ∅
 }.
 
@@ -201,8 +200,6 @@ Section ucmra_mixin.
   Proof. apply (mixin_ucmra_unit_valid _ (ucmra_mixin A)). Qed.
   Global Instance ucmra_unit_left_id : LeftId (≡) ∅ (@op A _).
   Proof. apply (mixin_ucmra_unit_left_id _ (ucmra_mixin A)). Qed.
-  Global Instance ucmra_unit_timeless : Timeless (∅ : A).
-  Proof. apply (mixin_ucmra_unit_timeless _ (ucmra_mixin A)). Qed.
   Lemma ucmra_pcore_unit : pcore (∅:A) ≡ Some ∅.
   Proof. apply (mixin_ucmra_pcore_unit _ (ucmra_mixin A)). Qed.
 End ucmra_mixin.
@@ -324,7 +321,7 @@ Proof. by apply cmra_pcore_dup' with x. Qed.
 
 (** ** Exclusive elements *)
 Lemma exclusiveN_l n x `{!Exclusive x} y : ✓{n} (x ⋅ y) → False.
-Proof. intros ?%cmra_validN_le%exclusive0_l; auto with arith. Qed.
+Proof. intros. eapply (exclusive0_l x y), cmra_validN_le; eauto with lia. Qed.
 Lemma exclusiveN_r n x `{!Exclusive x} y : ✓{n} (y ⋅ x) → False.
 Proof. rewrite comm. by apply exclusiveN_l. Qed.
 Lemma exclusive_l x `{!Exclusive x} y : ✓ (x ⋅ y) → False.
@@ -332,7 +329,7 @@ Proof. by move /cmra_valid_validN /(_ 0) /exclusive0_l. Qed.
 Lemma exclusive_r x `{!Exclusive x} y : ✓ (y ⋅ x) → False.
 Proof. rewrite comm. by apply exclusive_l. Qed.
 Lemma exclusiveN_opM n x `{!Exclusive x} my : ✓{n} (x ⋅? my) → my = None.
-Proof. destruct my. move=> /(exclusiveN_l _ x) []. done. Qed.
+Proof. destruct my as [y|]. move=> /(exclusiveN_l _ x) []. done. Qed.
 
 (** ** Order *)
 Lemma cmra_included_includedN n x y : x ≼ y → x ≼{n} y.
@@ -472,16 +469,16 @@ End total_core.
 Lemma cmra_timeless_included_l x y : Timeless x → ✓{0} y → x ≼{0} y → x ≼ y.
 Proof.
   intros ?? [x' ?].
-  destruct (cmra_extend 0 y x x') as ([z z']&Hy&Hz&Hz'); auto; simpl in *.
+  destruct (cmra_extend 0 y x x') as (z&z'&Hy&Hz&Hz'); auto; simpl in *.
   by exists z'; rewrite Hy (timeless x z).
 Qed.
-Lemma cmra_timeless_included_r n x y : Timeless y → x ≼{0} y → x ≼{n} y.
-Proof. intros ? [x' ?]. exists x'. by apply equiv_dist, (timeless y). Qed.
+Lemma cmra_timeless_included_r x y : Timeless y → x ≼{0} y → x ≼ y.
+Proof. intros ? [x' ?]. exists x'. by apply (timeless y). Qed.
 Lemma cmra_op_timeless x1 x2 :
   ✓ (x1 ⋅ x2) → Timeless x1 → Timeless x2 → Timeless (x1 ⋅ x2).
 Proof.
   intros ??? z Hz.
-  destruct (cmra_extend 0 z x1 x2) as ([y1 y2]&Hz'&?&?); auto; simpl in *.
+  destruct (cmra_extend 0 z x1 x2) as (y1&y2&Hz'&?&?); auto; simpl in *.
   { rewrite -?Hz. by apply cmra_valid_validN. }
   by rewrite Hz' (timeless x1 y1) // (timeless x2 y2).
 Qed.
@@ -504,8 +501,6 @@ Section ucmra.
   Context {A : ucmraT}.
   Implicit Types x y z : A.
 
-  Global Instance ucmra_unit_inhabited : Inhabited A := populate ∅.
-
   Lemma ucmra_unit_validN n : ✓{n} (∅:A).
   Proof. apply cmra_valid_validN, ucmra_unit_valid. Qed.
   Lemma ucmra_unit_leastN n x : ∅ ≼{n} x.
@@ -542,7 +537,7 @@ Section cmra_total.
   Context (validN_op_l : ∀ n (x y : A), ✓{n} (x ⋅ y) → ✓{n} x).
   Context (extend : ∀ n (x y1 y2 : A),
     ✓{n} x → x ≡{n}≡ y1 ⋅ y2 →
-    { z | x ≡ z.1 ⋅ z.2 ∧ z.1 ≡{n}≡ y1 ∧ z.2 ≡{n}≡ y2 }).
+    ∃ z1 z2, x ≡ z1 ⋅ z2 ∧ z1 ≡{n}≡ y1 ∧ z2 ≡{n}≡ y2).
   Lemma cmra_total_mixin : CMRAMixin A.
   Proof.
     split; auto.
@@ -692,7 +687,7 @@ Section discrete.
   Proof.
     destruct ra_mix; split; try done.
     - intros x; split; first done. by move=> /(_ 0).
-    - intros n x y1 y2 ??; by exists (y1,y2).
+    - intros n x y1 y2 ??; by exists y1, y2.
   Qed.
 End discrete.
 
@@ -883,9 +878,9 @@ Section prod.
       exists (z1,z2). by rewrite prod_included prod_pcore_Some.
     - intros n x y [??]; split; simpl in *; eauto using cmra_validN_op_l.
     - intros n x y1 y2 [??] [??]; simpl in *.
-      destruct (cmra_extend n (x.1) (y1.1) (y2.1)) as (z1&?&?&?); auto.
-      destruct (cmra_extend n (x.2) (y1.2) (y2.2)) as (z2&?&?&?); auto.
-      by exists ((z1.1,z2.1),(z1.2,z2.2)).
+      destruct (cmra_extend n (x.1) (y1.1) (y2.1)) as (z11&z12&?&?&?); auto.
+      destruct (cmra_extend n (x.2) (y1.2) (y2.2)) as (z21&z22&?&?&?); auto.
+      by exists (z11,z21), (z12,z22).
   Qed.
   Canonical Structure prodR :=
     CMRAT (A * B) prod_cofe_mixin prod_cmra_mixin.
@@ -924,7 +919,6 @@ Section prod_unit.
     split.
     - split; apply ucmra_unit_valid.
     - by split; rewrite /=left_id.
-    - by intros ? [??]; split; apply (timeless _).
     - rewrite prod_pcore_Some'; split; apply (persistent _).
   Qed.
   Canonical Structure prodUR :=
@@ -1050,13 +1044,12 @@ Section option.
         eauto using cmra_validN_op_l.
     - intros n mx my1 my2.
       destruct mx as [x|], my1 as [y1|], my2 as [y2|]; intros Hx Hx';
-        try (by exfalso; inversion Hx'; auto).
-      + destruct (cmra_extend n x y1 y2) as ([z1 z2]&?&?&?); auto.
-        { by inversion_clear Hx'. }
-        by exists (Some z1, Some z2); repeat constructor.
-      + by exists (Some x,None); inversion Hx'; repeat constructor.
-      + by exists (None,Some x); inversion Hx'; repeat constructor.
-      + exists (None,None); repeat constructor.
+        inversion_clear Hx'; auto.
+      + destruct (cmra_extend n x y1 y2) as (z1&z2&?&?&?); auto.
+        by exists (Some z1), (Some z2); repeat constructor.
+      + by exists (Some x), None; repeat constructor.
+      + by exists None, (Some x); repeat constructor.
+      + exists None, None; repeat constructor.
   Qed.
   Canonical Structure optionR :=
     CMRAT (option A) option_cofe_mixin option_cmra_mixin.
@@ -1066,7 +1059,7 @@ Section option.
 
   Instance option_empty : Empty (option A) := None.
   Lemma option_ucmra_mixin : UCMRAMixin optionR.
-  Proof. split. done. by intros []. by inversion_clear 1. done. Qed.
+  Proof. split. done. by intros []. done. Qed.
   Canonical Structure optionUR :=
     UCMRAT (option A) option_cofe_mixin option_cmra_mixin option_ucmra_mixin.
 
diff --git a/algebra/cofe.v b/algebra/cofe.v
index 3d43b09e74edd725186d30275af741ed72d3bb97..da126ca99a9fa35ba1b94cb84d083f44072c4a42 100644
--- a/algebra/cofe.v
+++ b/algebra/cofe.v
@@ -188,12 +188,21 @@ Definition fixpoint_eq : @fixpoint = @fixpoint_def := proj2_sig fixpoint_aux.
 
 Section fixpoint.
   Context {A : cofeT} `{Inhabited A} (f : A → A) `{!Contractive f}.
+
   Lemma fixpoint_unfold : fixpoint f ≡ f (fixpoint f).
   Proof.
     apply equiv_dist=>n.
     rewrite fixpoint_eq /fixpoint_def (conv_compl n (fixpoint_chain f)) //.
     induction n as [|n IH]; simpl; eauto using contractive_0, contractive_S.
   Qed.
+
+  Lemma fixpoint_unique (x : A) : x ≡ f x → x ≡ fixpoint f.
+  Proof.
+    rewrite !equiv_dist=> Hx n. induction n as [|n IH]; simpl in *.
+    - rewrite Hx fixpoint_unfold; eauto using contractive_0.
+    - rewrite Hx fixpoint_unfold. apply (contractive_S _), IH.
+  Qed.
+
   Lemma fixpoint_ne (g : A → A) `{!Contractive g} n :
     (∀ z, f z ≡{n}≡ g z) → fixpoint f ≡{n}≡ fixpoint g.
   Proof.
@@ -208,6 +217,8 @@ Section fixpoint.
 End fixpoint.
 
 (** Function space *)
+(* We make [cofe_fun] a definition so that we can register it as a canonical
+structure. *)
 Definition cofe_fun (A : Type) (B : cofeT) := A → B.
 
 Section cofe_fun.
@@ -673,8 +684,6 @@ Inductive later (A : Type) : Type := Next { later_car : A }.
 Add Printing Constructor later.
 Arguments Next {_} _.
 Arguments later_car {_} _.
-Lemma later_eta {A} (x : later A) : Next (later_car x) = x.
-Proof. by destruct x. Qed.
 
 Section later.
   Context {A : cofeT}.
diff --git a/algebra/cofe_solver.v b/algebra/cofe_solver.v
index 515028163eaa79d7aa529cfeb42c6c6b8e0ffe08..20a9ff0f01bfb9788a0525f29c873d107950da65 100644
--- a/algebra/cofe_solver.v
+++ b/algebra/cofe_solver.v
@@ -106,17 +106,17 @@ Proof. by assert (k = j) by lia; subst; rewrite !coerce_id. Qed.
 Lemma coerce_f {k j} (H : S k = S j) (x : A k) :
   coerce H (f k x) = f j (coerce (Nat.succ_inj _ _ H) x).
 Proof. by assert (k = j) by lia; subst; rewrite !coerce_id. Qed.
-Lemma gg_gg {k i i1 i2 j} (H1 : k = i + j) (H2 : k = i2 + (i1 + j)) (x : A k) :
+Lemma gg_gg {k i i1 i2 j} : ∀ (H1: k = i + j) (H2: k = i2 + (i1 + j)) (x: A k),
   gg i (coerce H1 x) = gg i1 (gg i2 (coerce H2 x)).
 Proof.
-  assert (i = i2 + i1) by lia; simplify_eq/=. revert j x H1.
+  intros ? -> x. assert (i = i2 + i1) as -> by lia. revert j x H1.
   induction i2 as [|i2 IH]; intros j X H1; simplify_eq/=;
     [by rewrite coerce_id|by rewrite g_coerce IH].
 Qed.
-Lemma ff_ff {k i i1 i2 j} (H1 : i + k = j) (H2 : i1 + (i2 + k) = j) (x : A k) :
+Lemma ff_ff {k i i1 i2 j} : ∀ (H1: i + k = j) (H2: i1 + (i2 + k) = j) (x: A k),
   coerce H1 (ff i x) = coerce H2 (ff i1 (ff i2 x)).
 Proof.
-  assert (i = i1 + i2) by lia; simplify_eq/=.
+  intros ? <- x. assert (i = i1 + i2) as -> by lia.
   induction i1 as [|i1 IH]; simplify_eq/=;
     [by rewrite coerce_id|by rewrite coerce_f IH].
 Qed.
diff --git a/algebra/csum.v b/algebra/csum.v
index a1d053517d42620852e5384082054203d5233705..d8a4d66ad172423785dd71629045f3323e0c93e2 100644
--- a/algebra/csum.v
+++ b/algebra/csum.v
@@ -209,15 +209,13 @@ Proof.
       exists (Cinr cb'). rewrite csum_included; eauto 10.
   - intros n [a1|b1|] [a2|b2|]; simpl; eauto using cmra_validN_op_l; done.
   - intros n [a|b|] y1 y2 Hx Hx'.
-    + destruct y1 as [a1|b1|], y2 as [a2|b2|]; try (exfalso; by inversion_clear Hx').
-      apply (inj Cinl) in Hx'.
-      destruct (cmra_extend n a a1 a2) as ([z1 z2]&?&?&?); auto.
-      exists (Cinl z1, Cinl z2). by repeat constructor.
-    + destruct y1 as [a1|b1|], y2 as [a2|b2|]; try (exfalso; by inversion_clear Hx').
-      apply (inj Cinr) in Hx'.
-      destruct (cmra_extend n b b1 b2) as ([z1 z2]&?&?&?); auto.
-      exists (Cinr z1, Cinr z2). by repeat constructor.
-    + by exists (CsumBot, CsumBot); destruct y1, y2; inversion_clear Hx'.
+    + destruct y1 as [a1|b1|], y2 as [a2|b2|]; inversion_clear Hx'.
+      destruct (cmra_extend n a a1 a2) as (z1&z2&?&?&?); auto.
+      exists (Cinl z1), (Cinl z2). by repeat constructor.
+    + destruct y1 as [a1|b1|], y2 as [a2|b2|]; inversion_clear Hx'.
+      destruct (cmra_extend n b b1 b2) as (z1&z2&?&?&?); auto.
+      exists (Cinr z1), (Cinr z2). by repeat constructor.
+    + by exists CsumBot, CsumBot; destruct y1, y2; inversion_clear Hx'.
 Qed.
 Canonical Structure csumR :=
   CMRAT (csum A B) csum_cofe_mixin csum_cmra_mixin.
diff --git a/algebra/dra.v b/algebra/dra.v
index 059729ff008246b40701209fa586afc270d88133..a3b51580b7b72aebc2140d628349929c153724dd 100644
--- a/algebra/dra.v
+++ b/algebra/dra.v
@@ -167,10 +167,10 @@ Proof.
   - intros [x px ?]; split; naive_solver eauto using dra_core_idemp.
   - intros [x px ?] [y py ?] [[z pz ?] [? Hy]]; simpl in *.
     destruct (dra_core_mono x z) as (z'&Hz').
-    unshelve eexists (Validity z' (px ∧ py ∧ pz) _); [|split; simpl].
+    unshelve eexists (Validity z' (px ∧ py ∧ pz) _).
     { intros (?&?&?); apply Hz'; tauto. }
-    + tauto.
-    + intros. rewrite Hy //. tauto.
+    split; simpl; first tauto.
+    intros. rewrite Hy //. tauto.
   - by intros [x px ?] [y py ?] (?&?&?).
 Qed.
 Canonical Structure validityR : cmraT :=
diff --git a/algebra/excl.v b/algebra/excl.v
index d6967f1d138e05dc0fea08357e3b7c2bb7758d06..63c884b67a0ac2a91d301fae0cb2e87128908f14 100644
--- a/algebra/excl.v
+++ b/algebra/excl.v
@@ -88,11 +88,7 @@ Proof.
   - by intros [?|] [?|] [?|]; constructor.
   - by intros [?|] [?|]; constructor.
   - by intros n [?|] [?|].
-  - intros n x y1 y2 ? Hx.
-    by exists match y1, y2 with
-      | Excl a1, Excl a2 => (Excl a1, Excl a2)
-      | ExclBot, _ => (ExclBot, y2) | _, ExclBot => (y1, ExclBot)
-      end; destruct y1, y2; inversion_clear Hx; repeat constructor.
+  - intros n x [?|] [?|] ?; inversion_clear 1; eauto.
 Qed.
 Canonical Structure exclR :=
   CMRAT (excl A) excl_cofe_mixin excl_cmra_mixin.
diff --git a/algebra/frac.v b/algebra/frac.v
index 47fd97642304649210775ce3d4d22324fef5991f..2a10b1cd166c725eb0cc7a8498bed1632838ef46 100644
--- a/algebra/frac.v
+++ b/algebra/frac.v
@@ -1,6 +1,5 @@
 From Coq.QArith Require Import Qcanon.
 From iris.algebra Require Export cmra.
-From iris.algebra Require Import upred.
 
 Notation frac := Qp (only parsing).
 
@@ -20,12 +19,6 @@ Qed.
 Canonical Structure fracR := discreteR frac frac_ra_mixin.
 End frac.
 
-(** Internalized properties *)
-Lemma frac_equivI {M} (x y : frac) : x ≡ y ⊣⊢ (x = y : uPred M).
-Proof. by uPred.unseal. Qed.
-Lemma frac_validI {M} (x : frac) : ✓ x ⊣⊢ (■ (x ≤ 1)%Qc : uPred M).
-Proof. by uPred.unseal. Qed.
-
 (** Exclusive *)
 Global Instance frac_full_exclusive : Exclusive 1%Qp.
 Proof.
diff --git a/algebra/gmap.v b/algebra/gmap.v
index c0e59099a938be9bfa46d84683d27cbc8e32a458..3315eba08543b940b0d42be7b638e61f69dbef10 100644
--- a/algebra/gmap.v
+++ b/algebra/gmap.v
@@ -62,7 +62,7 @@ Proof.
     [by constructor|by apply lookup_ne].
 Qed.
 
-Instance gmap_empty_timeless : Timeless (∅ : gmap K A).
+Global Instance gmap_empty_timeless : Timeless (∅ : gmap K A).
 Proof.
   intros m Hm i; specialize (Hm i); rewrite lookup_empty in Hm |- *.
   inversion_clear Hm; constructor.
@@ -137,21 +137,31 @@ Proof.
     rewrite !lookup_core. by apply cmra_core_mono.
   - intros n m1 m2 Hm i; apply cmra_validN_op_l with (m2 !! i).
     by rewrite -lookup_op.
-  - intros n m m1 m2 Hm Hm12.
-    assert (∀ i, m !! i ≡{n}≡ m1 !! i ⋅ m2 !! i) as Hm12'
-      by (by intros i; rewrite -lookup_op).
-    set (f i := cmra_extend n (m !! i) (m1 !! i) (m2 !! i) (Hm i) (Hm12' i)).
-    set (f_proj i := proj1_sig (f i)).
-    exists (map_imap (λ i _, (f_proj i).1) m, map_imap (λ i _, (f_proj i).2) m);
-      repeat split; intros i; rewrite /= ?lookup_op !lookup_imap.
-    + destruct (m !! i) as [x|] eqn:Hx; rewrite !Hx /=; [|constructor].
-      rewrite -Hx; apply (proj2_sig (f i)).
-    + destruct (m !! i) as [x|] eqn:Hx; rewrite /=; [apply (proj2_sig (f i))|].
-      move: (Hm12' i). rewrite Hx -!timeless_iff.
-      rewrite !(symmetry_iff _ None) !equiv_None op_None; tauto.
-    + destruct (m !! i) as [x|] eqn:Hx; simpl; [apply (proj2_sig (f i))|].
-      move: (Hm12' i). rewrite Hx -!timeless_iff.
-      rewrite !(symmetry_iff _ None) !equiv_None op_None; tauto.
+  - intros n m. induction m as [|i x m Hi IH] using map_ind=> m1 m2 Hmv Hm.
+    { exists ∅. exists ∅. (* FIXME: exists ∅, ∅. results in a TC loop in Coq 8.6 *)
+      split_and!=> -i; symmetry; symmetry in Hm; move: Hm=> /(_ i);
+        rewrite !lookup_op !lookup_empty ?dist_None op_None; intuition. }
+    destruct (IH (delete i m1) (delete i m2)) as (m1'&m2'&Hm'&Hm1'&Hm2').
+    { intros j; move: Hmv=> /(_ j). destruct (decide (i = j)) as [->|].
+      + intros _. by rewrite Hi.
+      + by rewrite lookup_insert_ne. }
+    { intros j; move: Hm=> /(_ j); destruct (decide (i = j)) as [->|].
+      + intros _. by rewrite lookup_op !lookup_delete Hi.
+      + by rewrite !lookup_op lookup_insert_ne // !lookup_delete_ne. }
+    destruct (cmra_extend n (Some x) (m1 !! i) (m2 !! i)) as (y1&y2&?&?&?).
+    { move: Hmv=> /(_ i). by rewrite lookup_insert. } 
+    { move: Hm=> /(_ i). by rewrite lookup_insert lookup_op. }
+    exists (partial_alter (λ _, y1) i m1'), (partial_alter (λ _, y2) i m2').
+    split_and!.
+    + intros j. destruct (decide (i = j)) as [->|].
+      * by rewrite lookup_insert lookup_op !lookup_partial_alter.
+      * by rewrite lookup_insert_ne // Hm' !lookup_op !lookup_partial_alter_ne.
+    + intros j. destruct (decide (i = j)) as [->|].
+      * by rewrite lookup_partial_alter.
+      * by rewrite lookup_partial_alter_ne // Hm1' lookup_delete_ne.
+    + intros j. destruct (decide (i = j)) as [->|].
+      * by rewrite lookup_partial_alter.
+      * by rewrite lookup_partial_alter_ne // Hm2' lookup_delete_ne.
 Qed.
 Canonical Structure gmapR :=
   CMRAT (gmap K A) gmap_cofe_mixin gmap_cmra_mixin.
@@ -164,7 +174,6 @@ Proof.
   split.
   - by intros i; rewrite lookup_empty.
   - by intros m i; rewrite /= lookup_op lookup_empty (left_id_L None _).
-  - apply gmap_empty_timeless.
   - constructor=> i. by rewrite lookup_omap lookup_empty.
 Qed.
 Canonical Structure gmapUR :=
@@ -334,8 +343,8 @@ Section freshness.
     ✓ u → LeftId (≡) u (⋅) →
     u ~~>: P → ∅ ~~>: λ m, ∃ y, m = {[ i := y ]} ∧ P y.
   Proof. eauto using alloc_unit_singleton_updateP. Qed.
-  Lemma alloc_unit_singleton_update u i (y : A) :
-    ✓ u → LeftId (≡) u (⋅) → u ~~> y → ∅ ~~> {[ i := y ]}.
+  Lemma alloc_unit_singleton_update (u : A) i (y : A) :
+    ✓ u → LeftId (≡) u (⋅) → u ~~> y → (∅:gmap K A) ~~> {[ i := y ]}.
   Proof.
     rewrite !cmra_update_updateP;
       eauto using alloc_unit_singleton_updateP with subst.
@@ -371,7 +380,7 @@ Proof.
 Qed.
 
 Lemma alloc_unit_singleton_local_update i x mf :
-  mf ≫= (!! i) = None → ✓ x → ∅ ~l~> {[ i := x ]} @ mf.
+  mf ≫= (!! i) = None → ✓ x → (∅:gmap K A) ~l~> {[ i := x ]} @ mf.
 Proof.
   intros Hi; apply alloc_singleton_local_update. by rewrite lookup_opM Hi.
 Qed.
diff --git a/algebra/gset.v b/algebra/gset.v
index c60bd02cd2eb8cd673f0a4ef89341ecfc098ddd9..f0c1c709950bfb3ca13ae23bb615e2cbb7fea158 100644
--- a/algebra/gset.v
+++ b/algebra/gset.v
@@ -54,26 +54,54 @@ Section gset.
   Canonical Structure gset_disjUR :=
     discreteUR (gset_disj K) gset_disj_ra_mixin gset_disj_ucmra_mixin.
 
-  Context `{Fresh K (gset K), !FreshSpec K (gset K)}.
   Arguments op _ _ _ _ : simpl never.
 
-  Lemma gset_alloc_updateP_strong (Q : gset_disj K → Prop) (I : gset K) X :
-    (∀ i, i ∉ X → i ∉ I → Q (GSet ({[i]} ∪ X))) → GSet X ~~>: Q.
+  Lemma gset_alloc_updateP_strong P (Q : gset_disj K → Prop) X :
+    (∀ Y, X ⊆ Y → ∃ j, j ∉ Y ∧ P j) →
+    (∀ i, i ∉ X → P i → Q (GSet ({[i]} ∪ X))) → GSet X ~~>: Q.
   Proof.
-    intros HQ; apply cmra_discrete_updateP=> ? /gset_disj_valid_inv_l [Y [->?]].
-    destruct (exist_fresh (X ∪ Y ∪ I)) as [i ?].
+    intros Hfresh HQ.
+    apply cmra_discrete_updateP=> ? /gset_disj_valid_inv_l [Y [->?]].
+    destruct (Hfresh (X ∪ Y)) as (i&?&?); first set_solver.
     exists (GSet ({[ i ]} ∪ X)); split.
     - apply HQ; set_solver by eauto.
     - apply gset_disj_valid_op. set_solver by eauto.
   Qed.
-  Lemma gset_alloc_updateP (Q : gset_disj K → Prop) X :
-    (∀ i, i ∉ X → Q (GSet ({[i]} ∪ X))) → GSet X ~~>: Q.
-  Proof. move=>??. eapply gset_alloc_updateP_strong with (I:=∅); by eauto. Qed.
-  Lemma gset_alloc_updateP_strong' (I : gset K) X :
-    GSet X ~~>: λ Y, ∃ i, Y = GSet ({[ i ]} ∪ X) ∧ i ∉ I ∧ i ∉ X.
+  Lemma gset_alloc_updateP_strong' P X :
+    (∀ Y, X ⊆ Y → ∃ j, j ∉ Y ∧ P j) →
+    GSet X ~~>: λ Y, ∃ i, Y = GSet ({[ i ]} ∪ X) ∧ i ∉ X ∧ P i.
   Proof. eauto using gset_alloc_updateP_strong. Qed.
-  Lemma gset_alloc_updateP' X : GSet X ~~>: λ Y, ∃ i, Y = GSet ({[ i ]} ∪ X) ∧ i ∉ X.
-  Proof. eauto using gset_alloc_updateP. Qed.
+
+  Lemma gset_alloc_empty_updateP_strong P (Q : gset_disj K → Prop) :
+    (∀ Y : gset K, ∃ j, j ∉ Y ∧ P j) →
+    (∀ i, P i → Q (GSet {[i]})) → GSet ∅ ~~>: Q.
+  Proof.
+    intros. apply (gset_alloc_updateP_strong P); eauto.
+    intros i; rewrite right_id_L; auto.
+  Qed.
+  Lemma gset_alloc_empty_updateP_strong' P :
+    (∀ Y : gset K, ∃ j, j ∉ Y ∧ P j) →
+    GSet ∅ ~~>: λ Y, ∃ i, Y = GSet {[ i ]} ∧ P i.
+  Proof. eauto using gset_alloc_empty_updateP_strong. Qed.
+
+  Section fresh_updates.
+    Context `{Fresh K (gset K), !FreshSpec K (gset K)}.
+
+    Lemma gset_alloc_updateP (Q : gset_disj K → Prop) X :
+      (∀ i, i ∉ X → Q (GSet ({[i]} ∪ X))) → GSet X ~~>: Q.
+    Proof.
+      intro; eapply gset_alloc_updateP_strong with (λ _, True); eauto.
+      intros Y ?; exists (fresh Y); eauto using is_fresh.
+    Qed.
+    Lemma gset_alloc_updateP' X : GSet X ~~>: λ Y, ∃ i, Y = GSet ({[ i ]} ∪ X) ∧ i ∉ X.
+    Proof. eauto using gset_alloc_updateP. Qed.
+
+    Lemma gset_alloc_empty_updateP (Q : gset_disj K → Prop) :
+      (∀ i, Q (GSet {[i]})) → GSet ∅ ~~>: Q.
+    Proof. intro. apply gset_alloc_updateP. intros i; rewrite right_id_L; auto. Qed.
+    Lemma gset_alloc_empty_updateP' : GSet ∅ ~~>: λ Y, ∃ i, Y = GSet {[ i ]}.
+    Proof. eauto using gset_alloc_empty_updateP. Qed.
+  End fresh_updates.
 
   Lemma gset_alloc_local_update X i Xf :
     i ∉ X → i ∉ Xf → GSet X ~l~> GSet ({[i]} ∪ X) @ Some (GSet Xf).
@@ -83,6 +111,12 @@ Section gset.
     - intros mZ ?%gset_disj_valid_op HXf.
       rewrite -gset_disj_union -?assoc ?HXf ?cmra_opM_assoc; set_solver.
   Qed.
+  Lemma gset_alloc_empty_local_update i Xf :
+    i ∉ Xf → GSet ∅ ~l~> GSet {[i]} @ Some (GSet Xf).
+  Proof.
+    intros. rewrite -(right_id_L _ _ {[i]}).
+    apply gset_alloc_local_update; set_solver.
+  Qed.
 End gset.
 
 Arguments gset_disjR _ {_ _}.
diff --git a/algebra/iprod.v b/algebra/iprod.v
index 312017cea454555ce4d2e965ba03bdc6a7e4e26d..181e5d80de9b0a417e2818a7ae79cb87e67c8289 100644
--- a/algebra/iprod.v
+++ b/algebra/iprod.v
@@ -117,9 +117,12 @@ Section iprod_cmra.
       by rewrite iprod_lookup_core; apply cmra_core_mono, Hf.
     - intros n f1 f2 Hf x; apply cmra_validN_op_l with (f2 x), Hf.
     - intros n f f1 f2 Hf Hf12.
-      set (g x := cmra_extend n (f x) (f1 x) (f2 x) (Hf x) (Hf12 x)).
-      exists ((λ x, (proj1_sig (g x)).1), (λ x, (proj1_sig (g x)).2)).
-      split_and?; intros x; apply (proj2_sig (g x)).
+      destruct (finite_choice (λ x (yy : B x * B x),
+        f x ≡ yy.1 ⋅ yy.2 ∧ yy.1 ≡{n}≡ f1 x ∧ yy.2 ≡{n}≡ f2 x)) as [gg Hgg].
+      { intros x. specialize (Hf12 x).
+        destruct (cmra_extend n (f x) (f1 x) (f2 x)) as (y1&y2&?&?&?); eauto.
+        exists (y1,y2); eauto. }
+      exists (λ x, gg x.1), (λ x, gg x.2). split_and!=> -?; naive_solver.
   Qed.
   Canonical Structure iprodR :=
     CMRAT (iprod B) iprod_cofe_mixin iprod_cmra_mixin.
@@ -132,12 +135,15 @@ Section iprod_cmra.
     split.
     - intros x; apply ucmra_unit_valid.
     - by intros f x; rewrite iprod_lookup_op left_id.
-    - intros f Hf x. by apply: timeless.
     - constructor=> x. apply persistent_core, _.
   Qed.
   Canonical Structure iprodUR :=
     UCMRAT (iprod B) iprod_cofe_mixin iprod_cmra_mixin iprod_ucmra_mixin.
 
+  Global Instance iprod_empty_timeless :
+    (∀ i, Timeless (∅ : B i)) → Timeless (∅ : iprod B).
+  Proof. intros ? f Hf x. by apply: timeless. Qed.
+
   (** Internalized properties *)
   Lemma iprod_equivI {M} g1 g2 : g1 ≡ g2 ⊣⊢ (∀ i, g1 i ≡ g2 i : uPred M).
   Proof. by uPred.unseal. Qed.
@@ -193,7 +199,8 @@ Section iprod_singleton.
   Proof. intros; by rewrite /iprod_singleton iprod_lookup_insert_ne. Qed.
 
   Global Instance iprod_singleton_timeless x (y : B x) :
-    Timeless y → Timeless (iprod_singleton x y) := _.
+    (∀ i, Timeless (∅ : B i)) →  Timeless y → Timeless (iprod_singleton x y).
+  Proof. apply _. Qed.
 
   Lemma iprod_singleton_validN n x (y : B x) : ✓{n} iprod_singleton x y ↔ ✓{n} y.
   Proof.
diff --git a/algebra/list.v b/algebra/list.v
index 3ab4d27c259f3b75a130ae1163f0814ff01d522b..373ea50573750d3f3ef3af6bf95436a26b60d9b1 100644
--- a/algebra/list.v
+++ b/algebra/list.v
@@ -81,13 +81,16 @@ End cofe.
 Arguments listC : clear implicits.
 
 (** Functor *)
+Lemma list_fmap_ext_ne {A} {B : cofeT} (f g : A → B) (l : list A) n :
+  (∀ x, f x ≡{n}≡ g x) → f <$> l ≡{n}≡ g <$> l.
+Proof. intros Hf. by apply Forall2_fmap, Forall_Forall2, Forall_true. Qed.
 Instance list_fmap_ne {A B : cofeT} (f : A → B) n:
   Proper (dist n ==> dist n) f → Proper (dist n ==> dist n) (fmap (M:=list) f).
-Proof. intros Hf l k ?; by eapply Forall2_fmap, Forall2_impl; eauto. Qed. 
+Proof. intros Hf l k ?; by eapply Forall2_fmap, Forall2_impl; eauto. Qed.
 Definition listC_map {A B} (f : A -n> B) : listC A -n> listC B :=
   CofeMor (fmap f : listC A → listC B).
 Instance listC_map_ne A B n : Proper (dist n ==> dist n) (@listC_map A B).
-Proof. intros f f' ? l; by apply Forall2_fmap, Forall_Forall2, Forall_true. Qed.
+Proof. intros f g ? l. by apply list_fmap_ext_ne. Qed.
 
 Program Definition listCF (F : cFunctor) : cFunctor := {|
   cFunctor_car A B := listC (cFunctor_car F A B);
@@ -190,15 +193,14 @@ Section cmra.
       rewrite !list_lookup_core. by apply cmra_core_mono.
     - intros n l1 l2. rewrite !list_lookup_validN.
       setoid_rewrite list_lookup_op. eauto using cmra_validN_op_l.
-    - intros n l. induction l as [|x l IH]=> -[|y1 l1] [|y2 l2] Hl Hl';
-        try (by exfalso; inversion_clear Hl').
-      + by exists ([], []).
-      + by exists ([], x :: l).
-      + by exists (x :: l, []).
-      + destruct (IH l1 l2) as ([l1' l2']&?&?&?),
-          (cmra_extend n x y1 y2) as ([y1' y2']&?&?&?);
-          [inversion_clear Hl; inversion_clear Hl'; auto ..|]; simplify_eq/=.
-        exists (y1' :: l1', y2' :: l2'); repeat constructor; auto.
+    - intros n l.
+      induction l as [|x l IH]=> -[|y1 l1] [|y2 l2] Hl; inversion_clear 1.
+      + by exists [], [].
+      + exists [], (x :: l); by repeat constructor.
+      + exists (x :: l), []; by repeat constructor.
+      + inversion_clear Hl. destruct (IH l1 l2) as (l1'&l2'&?&?&?),
+          (cmra_extend n x y1 y2) as (y1'&y2'&?&?&?); simplify_eq/=; auto.
+        exists (y1' :: l1'), (y2' :: l2'); repeat constructor; auto.
   Qed.
   Canonical Structure listR := CMRAT (list A) list_cofe_mixin list_cmra_mixin.
 
@@ -208,7 +210,6 @@ Section cmra.
     split.
     - constructor.
     - by intros l.
-    - by inversion_clear 1.
     - by constructor.
   Qed.
   Canonical Structure listUR :=
diff --git a/algebra/sts.v b/algebra/sts.v
index ee8570ccdf2f38826f155bf53a0738af122dbd40..6d7346321eb8b5dac2caa05409bff2163b7fb5d4 100644
--- a/algebra/sts.v
+++ b/algebra/sts.v
@@ -30,8 +30,10 @@ Inductive step : relation (state sts * tokens sts) :=
      tok s1 ∪ T1 ≡ tok s2 ∪ T2 → step (s1,T1) (s2,T2).
 Notation steps := (rtc step).
 Inductive frame_step (T : tokens sts) (s1 s2 : state sts) : Prop :=
+  (* Probably equivalent definition: (\mathcal{L}(s') ⊥ T) ∧ s \rightarrow s' *)
   | Frame_step T1 T2 :
      T1 ⊥ tok s1 ∪ T → step (s1,T1) (s2,T2) → frame_step T s1 s2.
+Notation frame_steps T := (rtc (frame_step T)).
 
 (** ** Closure under frame steps *)
 Record closed (S : states sts) (T : tokens sts) : Prop := Closed {
@@ -39,7 +41,7 @@ Record closed (S : states sts) (T : tokens sts) : Prop := Closed {
   closed_step s1 s2 : s1 ∈ S → frame_step T s1 s2 → s2 ∈ S
 }.
 Definition up (s : state sts) (T : tokens sts) : states sts :=
-  {[ s' | rtc (frame_step T) s s' ]}.
+  {[ s' | frame_steps T s s' ]}.
 Definition up_set (S : states sts) (T : tokens sts) : states sts :=
   S ≫= λ s, up s T.
 
@@ -86,7 +88,7 @@ Qed.
 
 (** ** Properties of closure under frame steps *)
 Lemma closed_steps S T s1 s2 :
-  closed S T → s1 ∈ S → rtc (frame_step T) s1 s2 → s2 ∈ S.
+  closed S T → s1 ∈ S → frame_steps T s1 s2 → s2 ∈ S.
 Proof. induction 3; eauto using closed_step. Qed.
 Lemma closed_op T1 T2 S1 S2 :
   closed S1 T1 → closed S2 T2 → closed (S1 ∩ S2) (T1 ∪ T2).
@@ -160,6 +162,7 @@ Proof. move=> ?? s [s' [? ?]]. eauto using closed_steps. Qed.
 End sts.
 
 Notation steps := (rtc step).
+Notation frame_steps T := (rtc (frame_step T)).
 
 (* The type of bounds we can give to the state of an STS. This is the type
    that we equip with an RA structure. *)
diff --git a/algebra/upred.v b/algebra/upred.v
index a8fe0398abf3342910385840c277f2de42a4863b..5a547af138f9e084cf9c01fd4df918824c2d47e1 100644
--- a/algebra/upred.v
+++ b/algebra/upred.v
@@ -1,4 +1,4 @@
-From iris.algebra Require Export cmra.
+From iris.algebra Require Export cmra updates.
 Local Hint Extern 1 (_ ≼ _) => etrans; [eassumption|].
 Local Hint Extern 1 (_ ≼ _) => etrans; [|eassumption].
 Local Hint Extern 10 (_ ≤ _) => omega.
@@ -264,6 +264,21 @@ Definition uPred_valid {M A} := proj1_sig uPred_valid_aux M A.
 Definition uPred_valid_eq :
   @uPred_valid = @uPred_valid_def := proj2_sig uPred_valid_aux.
 
+Program Definition uPred_rvs_def {M} (Q : uPred M) : uPred M :=
+  {| uPred_holds n x := ∀ k yf,
+      k ≤ n → ✓{k} (x ⋅ yf) → ∃ x', ✓{k} (x' ⋅ yf) ∧ Q k x' |}.
+Next Obligation.
+  intros M Q n x1 x2 HQ [x3 Hx] k yf Hk.
+  rewrite (dist_le _ _ _ _ Hx); last lia. intros Hxy.
+  destruct (HQ k (x3 â‹… yf)) as (x'&?&?); [auto|by rewrite assoc|].
+  exists (x' â‹… x3); split; first by rewrite -assoc.
+  apply uPred_mono with x'; eauto using cmra_includedN_l.
+Qed.
+Next Obligation. naive_solver. Qed.
+Definition uPred_rvs_aux : { x | x = @uPred_rvs_def }. by eexists. Qed.
+Definition uPred_rvs {M} := proj1_sig uPred_rvs_aux M.
+Definition uPred_rvs_eq : @uPred_rvs = @uPred_rvs_def := proj2_sig uPred_rvs_aux.
+
 Notation "P ⊢ Q" := (uPred_entails P%I Q%I)
   (at level 99, Q at level 200, right associativity) : C_scope.
 Notation "(⊢)" := uPred_entails (only parsing) : C_scope.
@@ -295,6 +310,12 @@ Notation "â–· P" := (uPred_later P)
   (at level 20, right associativity) : uPred_scope.
 Infix "≡" := uPred_eq : uPred_scope.
 Notation "✓ x" := (uPred_valid x) (at level 20) : uPred_scope.
+Notation "|=r=> Q" := (uPred_rvs Q)
+  (at level 99, Q at level 200, format "|=r=>  Q") : uPred_scope.
+Notation "P =r=> Q" := (P ⊢ |=r=> Q)
+  (at level 99, Q at level 200, only parsing) : C_scope.
+Notation "P =r=★ Q" := (P -★ |=r=> Q)%I
+  (at level 99, Q at level 200, format "P  =r=★  Q") : uPred_scope.
 
 Definition uPred_iff {M} (P Q : uPred M) : uPred M := ((P → Q) ∧ (Q → P))%I.
 Instance: Params (@uPred_iff) 1.
@@ -307,7 +328,13 @@ Arguments uPred_always_if _ !_ _/.
 Notation "â–¡? p P" := (uPred_always_if p P)
   (at level 20, p at level 0, P at level 20, format "â–¡? p  P").
 
-Class TimelessP {M} (P : uPred M) := timelessP : ▷ P ⊢ (P ∨ ▷ False).
+Definition uPred_except_last {M} (P : uPred M) : uPred M := ▷ False ∨ P.
+Notation "â—‡ P" := (uPred_except_last P)
+  (at level 20, right associativity) : uPred_scope.
+Instance: Params (@uPred_except_last) 1.
+Typeclasses Opaque uPred_except_last.
+
+Class TimelessP {M} (P : uPred M) := timelessP : ▷ P ⊢ ◇ P.
 Arguments timelessP {_} _ {_}.
 
 Class PersistentP {M} (P : uPred M) := persistentP : P ⊢ □ P.
@@ -317,7 +344,7 @@ Module uPred.
 Definition unseal :=
   (uPred_pure_eq, uPred_and_eq, uPred_or_eq, uPred_impl_eq, uPred_forall_eq,
   uPred_exist_eq, uPred_eq_eq, uPred_sep_eq, uPred_wand_eq, uPred_always_eq,
-  uPred_later_eq, uPred_ownM_eq, uPred_valid_eq).
+  uPred_later_eq, uPred_ownM_eq, uPred_valid_eq, uPred_rvs_eq).
 Ltac unseal := rewrite !unseal /=.
 
 Section uPred_logic.
@@ -333,11 +360,12 @@ Hint Immediate uPred_in_entails.
 Global Instance: PreOrder (@uPred_entails M).
 Proof.
   split.
-  * by intros P; split=> x i.
-  * by intros P Q Q' HP HQ; split=> x i ??; apply HQ, HP.
+  - by intros P; split=> x i.
+  - by intros P Q Q' HP HQ; split=> x i ??; apply HQ, HP.
 Qed.
 Global Instance: AntiSymm (⊣⊢) (@uPred_entails M).
 Proof. intros P Q HPQ HQP; split=> x n; by split; [apply HPQ|apply HQP]. Qed.
+
 Lemma equiv_spec P Q : (P ⊣⊢ Q) ↔ (P ⊢ Q) ∧ (Q ⊢ P).
 Proof.
   split; [|by intros [??]; apply (anti_symm (⊢))].
@@ -405,8 +433,8 @@ Global Instance eq_ne (A : cofeT) n :
   Proper (dist n ==> dist n ==> dist n) (@uPred_eq M A).
 Proof.
   intros x x' Hx y y' Hy; split=> n' z; unseal; split; intros; simpl in *.
-  * by rewrite -(dist_le _ _ _ _ Hx) -?(dist_le _ _ _ _ Hy); auto.
-  * by rewrite (dist_le _ _ _ _ Hx) ?(dist_le _ _ _ _ Hy); auto.
+  - by rewrite -(dist_le _ _ _ _ Hx) -?(dist_le _ _ _ _ Hy); auto.
+  - by rewrite (dist_le _ _ _ _ Hx) ?(dist_le _ _ _ _ Hy); auto.
 Qed.
 Global Instance eq_proper (A : cofeT) :
   Proper ((≡) ==> (≡) ==> (⊣⊢)) (@uPred_eq M A) := ne_proper_2 _.
@@ -437,7 +465,7 @@ Proof.
   intros n P Q HPQ; unseal; split=> -[|n'] x ??; simpl; [done|].
   apply (HPQ n'); eauto using cmra_validN_S.
 Qed.
-Global Instance later_proper :
+Global Instance later_proper' :
   Proper ((⊣⊢) ==> (⊣⊢)) (@uPred_later M) := ne_proper _.
 Global Instance always_ne n : Proper (dist n ==> dist n) (@uPred_always M).
 Proof.
@@ -460,10 +488,14 @@ Proof.
 Qed.
 Global Instance valid_proper {A : cmraT} :
   Proper ((≡) ==> (⊣⊢)) (@uPred_valid M A) := ne_proper _.
-Global Instance iff_ne n : Proper (dist n ==> dist n ==> dist n) (@uPred_iff M).
-Proof. unfold uPred_iff; solve_proper. Qed.
-Global Instance iff_proper :
-  Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@uPred_iff M) := ne_proper_2 _.
+Global Instance rvs_ne n : Proper (dist n ==> dist n) (@uPred_rvs M).
+Proof.
+  intros P Q HPQ.
+  unseal; split=> n' x; split; intros HP k yf ??;
+    destruct (HP k yf) as (x'&?&?); auto;
+    exists x'; split; auto; apply HPQ; eauto using cmra_validN_op_l.
+Qed.
+Global Instance rvs_proper : Proper ((≡) ==> (≡)) (@uPred_rvs M) := ne_proper _.
 
 (** Introduction and elimination rules *)
 Lemma pure_intro φ P : φ → P ⊢ ■ φ.
@@ -523,6 +555,11 @@ Proof.
 Qed.
 
 (* Derived logical stuff *)
+Global Instance iff_ne n : Proper (dist n ==> dist n ==> dist n) (@uPred_iff M).
+Proof. unfold uPred_iff; solve_proper. Qed.
+Global Instance iff_proper :
+  Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@uPred_iff M) := ne_proper_2 _.
+
 Lemma False_elim P : False ⊢ P.
 Proof. by apply (pure_elim False). Qed.
 Lemma True_intro P : P ⊢ True.
@@ -944,6 +981,7 @@ Proof. intros; rewrite -always_and_sep_l'; auto. Qed.
 Lemma always_entails_r' P Q : (P ⊢ □ Q) → P ⊢ P ★ □ Q.
 Proof. intros; rewrite -always_and_sep_r'; auto. Qed.
 
+(* Conditional always *)
 Global Instance always_if_ne n p : Proper (dist n ==> dist n) (@uPred_always_if M p).
 Proof. solve_proper. Qed.
 Global Instance always_if_proper p : Proper ((⊣⊢) ==> (⊣⊢)) (@uPred_always_if M p).
@@ -971,59 +1009,117 @@ Lemma later_mono P Q : (P ⊢ Q) → ▷ P ⊢ ▷ Q.
 Proof.
   unseal=> HP; split=>-[|n] x ??; [done|apply HP; eauto using cmra_validN_S].
 Qed.
-Lemma later_intro P : P ⊢ ▷ P.
-Proof.
-  unseal; split=> -[|n] x ??; simpl in *; [done|].
-  apply uPred_closed with (S n); eauto using cmra_validN_S.
-Qed.
 Lemma löb P : (▷ P → P) ⊢ P.
 Proof.
   unseal; split=> n x ? HP; induction n as [|n IH]; [by apply HP|].
   apply HP, IH, uPred_closed with (S n); eauto using cmra_validN_S.
 Qed.
+Lemma later_pure φ : ▷ ■ φ ⊢ ▷ False ∨ ■ φ.
+Proof. unseal; split=> -[|n] x ?? /=; auto. Qed.
 Lemma later_and P Q : ▷ (P ∧ Q) ⊣⊢ ▷ P ∧ ▷ Q.
 Proof. unseal; split=> -[|n] x; by split. Qed.
 Lemma later_or P Q : ▷ (P ∨ Q) ⊣⊢ ▷ P ∨ ▷ Q.
 Proof. unseal; split=> -[|n] x; simpl; tauto. Qed.
 Lemma later_forall {A} (Φ : A → uPred M) : (▷ ∀ a, Φ a) ⊣⊢ (∀ a, ▷ Φ a).
 Proof. unseal; by split=> -[|n] x. Qed.
-Lemma later_exist_1 {A} (Φ : A → uPred M) : (∃ a, ▷ Φ a) ⊢ (▷ ∃ a, Φ a).
-Proof. unseal; by split=> -[|[|n]] x. Qed.
-Lemma later_exist_2 `{Inhabited A} (Φ : A → uPred M) : (▷ ∃ a, Φ a) ⊢ ∃ a, ▷ Φ a.
-Proof. unseal; split=> -[|[|n]] x; done || by exists inhabitant. Qed.
+Lemma later_exist_false {A} (Φ : A → uPred M) :
+  (▷ ∃ a, Φ a) ⊢ ▷ False ∨ (∃ a, ▷ Φ a).
+Proof. unseal; split=> -[|[|n]] x /=; eauto. Qed.
 Lemma later_sep P Q : ▷ (P ★ Q) ⊣⊢ ▷ P ★ ▷ Q.
 Proof.
   unseal; split=> n x ?; split.
   - destruct n as [|n]; simpl.
     { by exists x, (core x); rewrite cmra_core_r. }
     intros (x1&x2&Hx&?&?); destruct (cmra_extend n x x1 x2)
-      as ([y1 y2]&Hx'&Hy1&Hy2); eauto using cmra_validN_S; simpl in *.
+      as (y1&y2&Hx'&Hy1&Hy2); eauto using cmra_validN_S; simpl in *.
     exists y1, y2; split; [by rewrite Hx'|by rewrite Hy1 Hy2].
   - destruct n as [|n]; simpl; [done|intros (x1&x2&Hx&?&?)].
     exists x1, x2; eauto using dist_S.
 Qed.
+Lemma later_false_excluded_middle P : ▷ P ⊢ ▷ False ∨ (▷ False → P).
+Proof.
+  unseal; split=> -[|n] x ? /= HP; [by left|right].
+  intros [|n'] x' ????; [|done].
+  eauto using uPred_closed, uPred_mono, cmra_included_includedN.
+Qed.
 
 (* Later derived *)
+Lemma later_proper P Q : (P ⊣⊢ Q) → ▷ P ⊣⊢ ▷ Q.
+Proof. by intros ->. Qed.
+Hint Resolve later_mono later_proper.
 Global Instance later_mono' : Proper ((⊢) ==> (⊢)) (@uPred_later M).
 Proof. intros P Q; apply later_mono. Qed.
 Global Instance later_flip_mono' :
   Proper (flip (⊢) ==> flip (⊢)) (@uPred_later M).
 Proof. intros P Q; apply later_mono. Qed.
+
+Lemma later_intro P : P ⊢ ▷ P.
+Proof.
+  rewrite -(and_elim_l (▷ P) P) -(löb (▷ P ∧ P)) later_and.
+  apply impl_intro_l; auto.
+Qed.
 Lemma later_True : ▷ True ⊣⊢ True.
 Proof. apply (anti_symm (⊢)); auto using later_intro. Qed.
 Lemma later_impl P Q : ▷ (P → Q) ⊢ ▷ P → ▷ Q.
-Proof.
-  apply impl_intro_l; rewrite -later_and.
-  apply later_mono, impl_elim with P; auto.
-Qed.
+Proof. apply impl_intro_l; rewrite -later_and; eauto using impl_elim. Qed.
 Lemma later_exist `{Inhabited A} (Φ : A → uPred M) :
   ▷ (∃ a, Φ a) ⊣⊢ (∃ a, ▷ Φ a).
-Proof. apply: anti_symm; eauto using later_exist_2, later_exist_1. Qed.
+Proof.
+  apply: anti_symm; [|apply exist_elim; eauto using exist_intro].
+  rewrite later_exist_false. apply or_elim; last done.
+  rewrite -(exist_intro inhabitant); auto.
+Qed.
 Lemma later_wand P Q : ▷ (P -★ Q) ⊢ ▷ P -★ ▷ Q.
-Proof. apply wand_intro_r;rewrite -later_sep; apply later_mono,wand_elim_l. Qed.
+Proof. apply wand_intro_r; rewrite -later_sep; eauto using wand_elim_l. Qed.
 Lemma later_iff P Q : ▷ (P ↔ Q) ⊢ ▷ P ↔ ▷ Q.
 Proof. by rewrite /uPred_iff later_and !later_impl. Qed.
 
+(* True now *)
+Global Instance except_last_ne n : Proper (dist n ==> dist n) (@uPred_except_last M).
+Proof. solve_proper. Qed.
+Global Instance except_last_proper : Proper ((⊣⊢) ==> (⊣⊢)) (@uPred_except_last M).
+Proof. solve_proper. Qed.
+Global Instance except_last_mono' : Proper ((⊢) ==> (⊢)) (@uPred_except_last M).
+Proof. solve_proper. Qed.
+Global Instance except_last_flip_mono' :
+  Proper (flip (⊢) ==> flip (⊢)) (@uPred_except_last M).
+Proof. solve_proper. Qed.
+
+Lemma except_last_intro P : P ⊢ ◇ P.
+Proof. rewrite /uPred_except_last; auto. Qed.
+Lemma except_last_mono P Q : (P ⊢ Q) → ◇ P ⊢ ◇ Q.
+Proof. by intros ->. Qed.
+Lemma except_last_idemp P : ◇ ◇ P ⊢ ◇ P.
+Proof. rewrite /uPred_except_last; auto. Qed.
+
+Lemma except_last_True : ◇ True ⊣⊢ True.
+Proof. rewrite /uPred_except_last. apply (anti_symm _); auto. Qed.
+Lemma except_last_or P Q : ◇ (P ∨ Q) ⊣⊢ ◇ P ∨ ◇ Q.
+Proof. rewrite /uPred_except_last. apply (anti_symm _); auto. Qed.
+Lemma except_last_and P Q : ◇ (P ∧ Q) ⊣⊢ ◇ P ∧ ◇ Q.
+Proof. by rewrite /uPred_except_last or_and_l. Qed.
+Lemma except_last_sep P Q : ◇ (P ★ Q) ⊣⊢ ◇ P ★ ◇ Q.
+Proof.
+  rewrite /uPred_except_last. apply (anti_symm _).
+  - apply or_elim; last by auto.
+    by rewrite -!or_intro_l -always_pure -always_later -always_sep_dup'.
+  - rewrite sep_or_r sep_elim_l sep_or_l; auto.
+Qed.
+Lemma except_last_forall {A} (Φ : A → uPred M) : ◇ (∀ a, Φ a) ⊢ ∀ a, ◇ Φ a.
+Proof. apply forall_intro=> a. by rewrite (forall_elim a). Qed.
+Lemma except_last_exist {A} (Φ : A → uPred M) : (∃ a, ◇ Φ a) ⊢ ◇ ∃ a, Φ a.
+Proof. apply exist_elim=> a. by rewrite (exist_intro a). Qed.
+Lemma except_last_later P : ◇ ▷ P ⊢ ▷ P.
+Proof. by rewrite /uPred_except_last -later_or False_or. Qed.
+Lemma except_last_always P : ◇ □ P ⊣⊢ □ ◇ P.
+Proof. by rewrite /uPred_except_last always_or always_later always_pure. Qed.
+Lemma except_last_always_if p P : ◇ □?p P ⊣⊢ □?p ◇ P.
+Proof. destruct p; simpl; auto using except_last_always. Qed.
+Lemma except_last_frame_l P Q : P ★ ◇ Q ⊢ ◇ (P ★ Q).
+Proof. by rewrite {1}(except_last_intro P) except_last_sep. Qed.
+Lemma except_last_frame_r P Q : ◇ P ★ Q ⊢ ◇ (P ★ Q).
+Proof. by rewrite {1}(except_last_intro Q) except_last_sep. Qed.
+
 (* Own *)
 Lemma ownM_op (a1 a2 : M) :
   uPred_ownM (a1 ⋅ a2) ⊣⊢ uPred_ownM a1 ★ uPred_ownM a2.
@@ -1040,10 +1136,14 @@ Proof.
   split=> n x /=; split; [by apply always_elim|unseal; intros Hx]; simpl.
   rewrite -(persistent_core a). by apply cmra_core_monoN.
 Qed.
-Lemma ownM_something : True ⊢ ∃ a, uPred_ownM a.
-Proof. unseal; split=> n x ??. by exists x; simpl. Qed.
 Lemma ownM_empty : True ⊢ uPred_ownM ∅.
 Proof. unseal; split=> n x ??; by  exists x; rewrite left_id. Qed.
+Lemma later_ownM a : Timeless a → ▷ uPred_ownM a ⊢ ▷ False ∨ uPred_ownM a.
+Proof.
+  unseal=> Ha; split=> -[|n] x ?? /=; [by left|right].
+  apply cmra_included_includedN, cmra_timeless_included_l,
+    cmra_includedN_le with n; eauto using cmra_validN_le.
+Qed.
 
 (* Valid *)
 Lemma ownM_valid (a : M) : uPred_ownM a ⊢ ✓ a.
@@ -1065,6 +1165,63 @@ Proof. by intros; rewrite ownM_valid valid_elim. Qed.
 Global Instance ownM_mono : Proper (flip (≼) ==> (⊢)) (@uPred_ownM M).
 Proof. intros a b [b' ->]. rewrite ownM_op. eauto. Qed.
 
+(* Viewshifts *)
+Lemma rvs_intro P : P =r=> P.
+Proof.
+  unseal. split=> n x ? HP k yf ?; exists x; split; first done.
+  apply uPred_closed with n; eauto using cmra_validN_op_l.
+Qed.
+Lemma rvs_mono P Q : (P ⊢ Q) → (|=r=> P) =r=> Q.
+Proof.
+  unseal. intros HPQ; split=> n x ? HP k yf ??.
+  destruct (HP k yf) as (x'&?&?); eauto.
+  exists x'; split; eauto using uPred_in_entails, cmra_validN_op_l.
+Qed.
+Lemma rvs_trans P : (|=r=> |=r=> P) =r=> P.
+Proof. unseal; split; naive_solver. Qed.
+Lemma rvs_frame_r P R : (|=r=> P) ★ R =r=> P ★ R.
+Proof.
+  unseal; split; intros n x ? (x1&x2&Hx&HP&?) k yf ??.
+  destruct (HP k (x2 â‹… yf)) as (x'&?&?); eauto.
+  { by rewrite assoc -(dist_le _ _ _ _ Hx); last lia. }
+  exists (x' â‹… x2); split; first by rewrite -assoc.
+  exists x', x2; split_and?; auto.
+  apply uPred_closed with n; eauto 3 using cmra_validN_op_l, cmra_validN_op_r.
+Qed.
+Lemma rvs_ownM_updateP x (Φ : M → Prop) :
+  x ~~>: Φ → uPred_ownM x =r=> ∃ y, ■ Φ y ∧ uPred_ownM y.
+Proof.
+  unseal=> Hup; split=> n x2 ? [x3 Hx] k yf ??.
+  destruct (Hup k (Some (x3 â‹… yf))) as (y&?&?); simpl in *.
+  { rewrite /= assoc -(dist_le _ _ _ _ Hx); auto. }
+  exists (y â‹… x3); split; first by rewrite -assoc.
+  exists y; eauto using cmra_includedN_l.
+Qed.
+
+(** * Derived rules *)
+Global Instance rvs_mono' : Proper ((⊢) ==> (⊢)) (@uPred_rvs M).
+Proof. intros P Q; apply rvs_mono. Qed.
+Global Instance rvs_flip_mono' : Proper (flip (⊢) ==> flip (⊢)) (@uPred_rvs M).
+Proof. intros P Q; apply rvs_mono. Qed.
+Lemma rvs_frame_l R Q : (R ★ |=r=> Q) =r=> R ★ Q.
+Proof. rewrite !(comm _ R); apply rvs_frame_r. Qed.
+Lemma rvs_wand_l P Q : (P -★ Q) ★ (|=r=> P) =r=> Q.
+Proof. by rewrite rvs_frame_l wand_elim_l. Qed.
+Lemma rvs_wand_r P Q : (|=r=> P) ★ (P -★ Q) =r=> Q.
+Proof. by rewrite rvs_frame_r wand_elim_r. Qed.
+Lemma rvs_sep P Q : (|=r=> P) ★ (|=r=> Q) =r=> P ★ Q.
+Proof. by rewrite rvs_frame_r rvs_frame_l rvs_trans. Qed.
+Lemma rvs_ownM_update x y : x ~~> y → uPred_ownM x ⊢ |=r=> uPred_ownM y.
+Proof.
+  intros; rewrite (rvs_ownM_updateP _ (y =)); last by apply cmra_update_updateP.
+  by apply rvs_mono, exist_elim=> y'; apply pure_elim_l=> ->.
+Qed.
+Lemma except_last_rvs P : ◇ (|=r=> P) ⊢ (|=r=> ◇ P).
+Proof.
+  rewrite /uPred_except_last. apply or_elim; auto using rvs_mono.
+  by rewrite -rvs_intro -or_intro_l.
+Qed.
+
 (* Products *)
 Lemma prod_equivI {A B : cofeT} (x y : A * B) : x ≡ y ⊣⊢ x.1 ≡ y.1 ∧ x.2 ≡ y.2.
 Proof. by unseal. Qed.
@@ -1072,8 +1229,7 @@ Lemma prod_validI {A B : cmraT} (x : A * B) : ✓ x ⊣⊢ ✓ x.1 ∧ ✓ x.2.
 Proof. by unseal. Qed.
 
 (* Later *)
-Lemma later_equivI {A : cofeT} (x y : later A) :
-  x ≡ y ⊣⊢ ▷ (later_car x ≡ later_car y).
+Lemma later_equivI {A : cofeT} (x y : A) : Next x ≡ Next y ⊣⊢ ▷ (x ≡ y).
 Proof. by unseal. Qed.
 
 (* Discrete *)
@@ -1096,74 +1252,60 @@ Lemma option_validI {A : cmraT} (mx : option A) :
   ✓ mx ⊣⊢ match mx with Some x => ✓ x | None => True end.
 Proof. uPred.unseal. by destruct mx. Qed.
 
-(* Timeless *)
-Lemma timelessP_spec P : TimelessP P ↔ ∀ n x, ✓{n} x → P 0 x → P n x.
-Proof.
-  split.
-  - intros HP n x ??; induction n as [|n]; auto.
-    move: HP; rewrite /TimelessP; unseal=> /uPred_in_entails /(_ (S n) x).
-    by destruct 1; auto using cmra_validN_S.
-  - move=> HP; rewrite /TimelessP; unseal; split=> -[|n] x /=; auto; left.
-    apply HP, uPred_closed with n; eauto using cmra_validN_le.
-Qed.
-
+(* Timeless instances *)
 Global Instance pure_timeless φ : TimelessP (■ φ : uPred M)%I.
-Proof. by apply timelessP_spec; unseal => -[|n] x. Qed.
+Proof. by rewrite /TimelessP later_pure. Qed.
 Global Instance valid_timeless {A : cmraT} `{CMRADiscrete A} (a : A) :
-   TimelessP (✓ a : uPred M)%I.
-Proof.
-  apply timelessP_spec; unseal=> n x /=. by rewrite -!cmra_discrete_valid_iff.
-Qed.
+  TimelessP (✓ a : uPred M)%I.
+Proof. rewrite /TimelessP !discrete_valid. apply (timelessP _). Qed.
 Global Instance and_timeless P Q: TimelessP P → TimelessP Q → TimelessP (P ∧ Q).
-Proof. by intros ??; rewrite /TimelessP later_and or_and_r; apply and_mono. Qed.
+Proof. intros; rewrite /TimelessP except_last_and later_and; auto. Qed.
 Global Instance or_timeless P Q : TimelessP P → TimelessP Q → TimelessP (P ∨ Q).
-Proof.
-  intros; rewrite /TimelessP later_or (timelessP _) (timelessP Q); eauto 10.
-Qed.
+Proof. intros; rewrite /TimelessP except_last_or later_or; auto. Qed.
 Global Instance impl_timeless P Q : TimelessP Q → TimelessP (P → Q).
 Proof.
-  rewrite !timelessP_spec; unseal=> HP [|n] x ? HPQ [|n'] x' ????; auto.
-  apply HP, HPQ, uPred_closed with (S n'); eauto using cmra_validN_le.
+  rewrite /TimelessP=> HQ. rewrite later_false_excluded_middle.
+  apply or_mono, impl_intro_l; first done.
+  rewrite -{2}(löb Q); apply impl_intro_l.
+  rewrite HQ /uPred_except_last !and_or_r. apply or_elim; last auto.
+  by rewrite assoc (comm _ _ P) -assoc !impl_elim_r.
 Qed.
 Global Instance sep_timeless P Q: TimelessP P → TimelessP Q → TimelessP (P ★ Q).
-Proof.
-  intros; rewrite /TimelessP later_sep (timelessP P) (timelessP Q).
-  apply wand_elim_l', or_elim; apply wand_intro_r; auto.
-  apply wand_elim_r', or_elim; apply wand_intro_r; auto.
-  rewrite ?(comm _ Q); auto.
-Qed.
+Proof. intros; rewrite /TimelessP except_last_sep later_sep; auto. Qed.
 Global Instance wand_timeless P Q : TimelessP Q → TimelessP (P -★ Q).
 Proof.
-  rewrite !timelessP_spec; unseal=> HP [|n] x ? HPQ [|n'] x' ???; auto.
-  apply HP, HPQ, uPred_closed with (S n');
-    eauto 3 using cmra_validN_le, cmra_validN_op_r.
+  rewrite /TimelessP=> HQ. rewrite later_false_excluded_middle.
+  apply or_mono, wand_intro_l; first done.
+  rewrite -{2}(löb Q); apply impl_intro_l.
+  rewrite HQ /uPred_except_last !and_or_r. apply or_elim; last auto.
+  rewrite -(always_pure) -always_later always_and_sep_l'.
+  by rewrite assoc (comm _ _ P) -assoc -always_and_sep_l' impl_elim_r wand_elim_r.
 Qed.
 Global Instance forall_timeless {A} (Ψ : A → uPred M) :
   (∀ x, TimelessP (Ψ x)) → TimelessP (∀ x, Ψ x).
-Proof. by setoid_rewrite timelessP_spec; unseal=> HΨ n x ?? a; apply HΨ. Qed.
+Proof.
+  rewrite /TimelessP=> HQ. rewrite later_false_excluded_middle.
+  apply or_mono; first done. apply forall_intro=> x.
+  rewrite -(löb (Ψ x)); apply impl_intro_l.
+  rewrite HQ /uPred_except_last !and_or_r. apply or_elim; last auto.
+  by rewrite impl_elim_r (forall_elim x).
+Qed.
 Global Instance exist_timeless {A} (Ψ : A → uPred M) :
   (∀ x, TimelessP (Ψ x)) → TimelessP (∃ x, Ψ x).
 Proof.
-  by setoid_rewrite timelessP_spec; unseal=> HΨ [|n] x ?;
-    [|intros [a ?]; exists a; apply HΨ].
+  rewrite /TimelessP=> ?. rewrite later_exist_false. apply or_elim.
+  - rewrite /uPred_except_last; auto.
+  - apply exist_elim=> x. rewrite -(exist_intro x); auto.
 Qed.
 Global Instance always_timeless P : TimelessP P → TimelessP (□ P).
-Proof.
-  intros ?; rewrite /TimelessP.
-  by rewrite -always_pure -!always_later -always_or; apply always_mono.
-Qed.
+Proof. intros; rewrite /TimelessP except_last_always -always_later; auto. Qed.
 Global Instance always_if_timeless p P : TimelessP P → TimelessP (□?p P).
 Proof. destruct p; apply _. Qed.
 Global Instance eq_timeless {A : cofeT} (a b : A) :
   Timeless a → TimelessP (a ≡ b : uPred M)%I.
-Proof.
-  intro; apply timelessP_spec; unseal=> n x ??; by apply equiv_dist, timeless.
-Qed.
+Proof. intros. rewrite /TimelessP !timeless_eq. apply (timelessP _). Qed.
 Global Instance ownM_timeless (a : M) : Timeless a → TimelessP (uPred_ownM a).
-Proof.
-  intro; apply timelessP_spec; unseal=> n x ??; apply cmra_included_includedN,
-    cmra_timeless_included_l; eauto using cmra_validN_le.
-Qed.
+Proof. apply later_ownM. Qed.
 
 (* Persistence *)
 Global Instance pure_persistent φ : PersistentP (■ φ : uPred M)%I.
@@ -1216,6 +1358,20 @@ Lemma always_entails_l P Q `{!PersistentP Q} : (P ⊢ Q) → P ⊢ Q ★ P.
 Proof. by rewrite -(always_always Q); apply always_entails_l'. Qed.
 Lemma always_entails_r P Q `{!PersistentP Q} : (P ⊢ Q) → P ⊢ P ★ Q.
 Proof. by rewrite -(always_always Q); apply always_entails_r'. Qed.
+
+(* Soundness results *)
+Lemma adequacy φ n : (True ⊢ Nat.iter n (λ P, |=r=> ▷ P) (■ φ)) → φ.
+Proof.
+  cut (∀ x, ✓{n} x → Nat.iter n (λ P, |=r=> ▷ P)%I (■ φ)%I n x → φ).
+  { intros help H. eapply (help ∅); eauto using ucmra_unit_validN.
+    eapply H; try unseal; eauto using ucmra_unit_validN. }
+  unseal. induction n as [|n IH]=> x Hx Hvs; auto.
+  destruct (Hvs (S n) ∅) as (x'&?&?); rewrite ?right_id; auto.
+  eapply IH with x'; eauto using cmra_validN_S, cmra_validN_op_l.
+Qed.
+
+Theorem soundness : ¬ (True ⊢ False).
+Proof. exact (adequacy False 0). Qed.
 End uPred_logic.
 
 (* Hint DB for the logic *)
diff --git a/algebra/upred_big_op.v b/algebra/upred_big_op.v
index 85e2b334696d2fa2a0af9d1c6bee29394ef9840f..c9ae5148c36f5826aab548cc3ec81227ba668597 100644
--- a/algebra/upred_big_op.v
+++ b/algebra/upred_big_op.v
@@ -6,6 +6,9 @@ Import uPred.
 
 - The operators [ [★] Ps ] and [ [∧] Ps ] fold [★] and [∧] over the list [Ps].
   This operator is not a quantifier, so it binds strongly.
+- The operator [ [★ list] k ↦ x ∈ l, P ] asserts that [P] holds separately for
+  each element [x] at position [x] in the list [l]. This operator is a
+  quantifier, and thus has the same precedence as [∀] and [∃].
 - The operator [ [★ map] k ↦ x ∈ m, P ] asserts that [P] holds separately for
   each [k ↦ x] in the map [m]. This operator is a quantifier, and thus has the
   same precedence as [∀] and [∃].
@@ -25,11 +28,22 @@ Instance: Params (@uPred_big_sep) 1.
 Notation "'[★]' Ps" := (uPred_big_sep Ps) (at level 20) : uPred_scope.
 
 (** * Other big ops *)
-(** We use a type class to obtain overloaded notations *)
+Definition uPred_big_sepL {M A} (l : list A)
+  (Φ : nat → A → uPred M) : uPred M := [★] (imap Φ l).
+Instance: Params (@uPred_big_sepL) 2.
+Typeclasses Opaque uPred_big_sepL.
+Notation "'[★' 'list' ] k ↦ x ∈ l , P" := (uPred_big_sepL l (λ k x, P))
+  (at level 200, l at level 10, k, x at level 1, right associativity,
+   format "[★  list ]  k ↦ x  ∈  l ,  P") : uPred_scope.
+Notation "'[★' 'list' ] x ∈ l , P" := (uPred_big_sepL l (λ _ x, P))
+  (at level 200, l at level 10, x at level 1, right associativity,
+   format "[★  list ]  x  ∈  l ,  P") : uPred_scope.
+
 Definition uPred_big_sepM {M} `{Countable K} {A}
     (m : gmap K A) (Φ : K → A → uPred M) : uPred M :=
   [★] (curry Φ <$> map_to_list m).
 Instance: Params (@uPred_big_sepM) 6.
+Typeclasses Opaque uPred_big_sepM.
 Notation "'[★' 'map' ] k ↦ x ∈ m , P" := (uPred_big_sepM m (λ k x, P))
   (at level 200, m at level 10, k, x at level 1, right associativity,
    format "[★  map ]  k ↦ x  ∈  m ,  P") : uPred_scope.
@@ -37,22 +51,27 @@ Notation "'[★' 'map' ] k ↦ x ∈ m , P" := (uPred_big_sepM m (λ k x, P))
 Definition uPred_big_sepS {M} `{Countable A}
   (X : gset A) (Φ : A → uPred M) : uPred M := [★] (Φ <$> elements X).
 Instance: Params (@uPred_big_sepS) 5.
+Typeclasses Opaque uPred_big_sepS.
 Notation "'[★' 'set' ] x ∈ X , P" := (uPred_big_sepS X (λ x, P))
   (at level 200, X at level 10, x at level 1, right associativity,
    format "[★  set ]  x  ∈  X ,  P") : uPred_scope.
 
-(** * Persistence of lists of uPreds *)
+(** * Persistence and timelessness of lists of uPreds *)
 Class PersistentL {M} (Ps : list (uPred M)) :=
   persistentL : Forall PersistentP Ps.
 Arguments persistentL {_} _ {_}.
 
+Class TimelessL {M} (Ps : list (uPred M)) :=
+  timelessL : Forall TimelessP Ps.
+Arguments timelessL {_} _ {_}.
+
 (** * Properties *)
 Section big_op.
 Context {M : ucmraT}.
 Implicit Types Ps Qs : list (uPred M).
 Implicit Types A : Type.
 
-(** ** Big ops over lists *)
+(** ** Generic big ops over lists of upreds *)
 Global Instance big_and_proper : Proper ((≡) ==> (⊣⊢)) (@uPred_big_and M).
 Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
 Global Instance big_sep_proper : Proper ((≡) ==> (⊣⊢)) (@uPred_big_sep M).
@@ -105,6 +124,206 @@ Proof. induction 1; simpl; auto with I. Qed.
 Lemma big_sep_elem_of Ps P : P ∈ Ps → [★] Ps ⊢ P.
 Proof. induction 1; simpl; auto with I. Qed.
 
+(** ** Persistence *)
+Global Instance big_and_persistent Ps : PersistentL Ps → PersistentP ([∧] Ps).
+Proof. induction 1; apply _. Qed.
+Global Instance big_sep_persistent Ps : PersistentL Ps → PersistentP ([★] Ps).
+Proof. induction 1; apply _. Qed.
+
+Global Instance nil_persistent : PersistentL (@nil (uPred M)).
+Proof. constructor. Qed.
+Global Instance cons_persistent P Ps :
+  PersistentP P → PersistentL Ps → PersistentL (P :: Ps).
+Proof. by constructor. Qed.
+Global Instance app_persistent Ps Ps' :
+  PersistentL Ps → PersistentL Ps' → PersistentL (Ps ++ Ps').
+Proof. apply Forall_app_2. Qed.
+
+Global Instance fmap_persistent {A} (f : A → uPred M) xs :
+  (∀ x, PersistentP (f x)) → PersistentL (f <$> xs).
+Proof. intros. apply Forall_fmap, Forall_forall; auto. Qed.
+Global Instance zip_with_persistent {A B} (f : A → B → uPred M) xs ys :
+  (∀ x y, PersistentP (f x y)) → PersistentL (zip_with f xs ys).
+Proof.
+  unfold PersistentL=> ?; revert ys; induction xs=> -[|??]; constructor; auto.
+Qed.
+Global Instance imap_persistent {A} (f : nat → A → uPred M) xs :
+  (∀ i x, PersistentP (f i x)) → PersistentL (imap f xs).
+Proof.
+  rewrite /PersistentL /imap=> ?. generalize 0. induction xs; constructor; auto.
+Qed.
+
+(** ** Timelessness *)
+Global Instance big_and_timeless Ps : TimelessL Ps → TimelessP ([∧] Ps).
+Proof. induction 1; apply _. Qed.
+Global Instance big_sep_timeless Ps : TimelessL Ps → TimelessP ([★] Ps).
+Proof. induction 1; apply _. Qed.
+
+Global Instance nil_timeless : TimelessL (@nil (uPred M)).
+Proof. constructor. Qed.
+Global Instance cons_timeless P Ps :
+  TimelessP P → TimelessL Ps → TimelessL (P :: Ps).
+Proof. by constructor. Qed.
+Global Instance app_timeless Ps Ps' :
+  TimelessL Ps → TimelessL Ps' → TimelessL (Ps ++ Ps').
+Proof. apply Forall_app_2. Qed.
+
+Global Instance fmap_timeless {A} (f : A → uPred M) xs :
+  (∀ x, TimelessP (f x)) → TimelessL (f <$> xs).
+Proof. intros. apply Forall_fmap, Forall_forall; auto. Qed.
+Global Instance zip_with_timeless {A B} (f : A → B → uPred M) xs ys :
+  (∀ x y, TimelessP (f x y)) → TimelessL (zip_with f xs ys).
+Proof.
+  unfold TimelessL=> ?; revert ys; induction xs=> -[|??]; constructor; auto.
+Qed.
+Global Instance imap_timeless {A} (f : nat → A → uPred M) xs :
+  (∀ i x, TimelessP (f i x)) → TimelessL (imap f xs).
+Proof.
+  rewrite /TimelessL /imap=> ?. generalize 0. induction xs; constructor; auto.
+Qed.
+
+(** ** Big ops over lists *)
+Section list.
+  Context {A : Type}.
+  Implicit Types l : list A.
+  Implicit Types Φ Ψ : nat → A → uPred M.
+
+  Lemma big_sepL_mono Φ Ψ l :
+    (∀ k y, l !! k = Some y → Φ k y ⊢ Ψ k y) →
+    ([★ list] k ↦ y ∈ l, Φ k y) ⊢ [★ list] k ↦ y ∈ l, Ψ k y.
+  Proof.
+    intros HΦ. apply big_sep_mono'.
+    revert Φ Ψ HΦ. induction l as [|x l IH]=> Φ Ψ HΦ; first constructor.
+    rewrite !imap_cons; constructor; eauto.
+  Qed.
+  Lemma big_sepL_proper Φ Ψ l :
+    (∀ k y, l !! k = Some y → Φ k y ⊣⊢ Ψ k y) →
+    ([★ list] k ↦ y ∈ l, Φ k y) ⊣⊢ ([★ list] k ↦ y ∈ l, Ψ k y).
+  Proof.
+    intros ?; apply (anti_symm (⊢)); apply big_sepL_mono;
+      eauto using equiv_entails, equiv_entails_sym, lookup_weaken.
+  Qed.
+
+  Global Instance big_sepL_ne l n :
+    Proper (pointwise_relation _ (pointwise_relation _ (dist n)) ==> (dist n))
+           (uPred_big_sepL (M:=M) l).
+  Proof.
+    intros Φ Ψ HΦ. apply big_sep_ne.
+    revert Φ Ψ HΦ. induction l as [|x l IH]=> Φ Ψ HΦ; first constructor.
+    rewrite !imap_cons; constructor. by apply HΦ. apply IH=> n'; apply HΦ.
+  Qed.
+  Global Instance big_sepL_proper' l :
+    Proper (pointwise_relation _ (pointwise_relation _ (⊣⊢)) ==> (⊣⊢))
+           (uPred_big_sepL (M:=M) l).
+  Proof. intros Φ1 Φ2 HΦ. by apply big_sepL_proper; intros; last apply HΦ. Qed.
+  Global Instance big_sepL_mono' l :
+    Proper (pointwise_relation _ (pointwise_relation _ (⊢)) ==> (⊢))
+           (uPred_big_sepL (M:=M) l).
+  Proof. intros Φ1 Φ2 HΦ. by apply big_sepL_mono; intros; last apply HΦ. Qed.
+
+  Lemma big_sepL_nil Φ : ([★ list] k↦y ∈ nil, Φ k y) ⊣⊢ True.
+  Proof. done. Qed.
+
+  Lemma big_sepL_cons Φ x l :
+    ([★ list] k↦y ∈ x :: l, Φ k y) ⊣⊢ Φ 0 x ★ [★ list] k↦y ∈ l, Φ (S k) y.
+  Proof. by rewrite /uPred_big_sepL imap_cons. Qed.
+
+  Lemma big_sepL_singleton Φ x : ([★ list] k↦y ∈ [x], Φ k y) ⊣⊢ Φ 0 x.
+  Proof. by rewrite big_sepL_cons big_sepL_nil right_id. Qed.
+
+  Lemma big_sepL_app Φ l1 l2 :
+    ([★ list] k↦y ∈ l1 ++ l2, Φ k y)
+    ⊣⊢ ([★ list] k↦y ∈ l1, Φ k y) ★ ([★ list] k↦y ∈ l2, Φ (length l1 + k) y).
+  Proof. by rewrite /uPred_big_sepL imap_app big_sep_app. Qed.
+
+  Lemma big_sepL_lookup Φ l i x :
+    l !! i = Some x → ([★ list] k↦y ∈ l, Φ k y) ⊢ Φ i x.
+  Proof.
+    intros. rewrite -(take_drop_middle l i x) // big_sepL_app big_sepL_cons.
+    rewrite Nat.add_0_r take_length_le; eauto using lookup_lt_Some, Nat.lt_le_incl.
+    by rewrite sep_elim_r sep_elim_l.
+  Qed.
+
+  Lemma big_sepL_elem_of (Φ : A → uPred M) l x :
+    x ∈ l → ([★ list] y ∈ l, Φ y) ⊢ Φ x.
+  Proof.
+    intros [i ?]%elem_of_list_lookup; eauto using (big_sepL_lookup (λ _, Φ)).
+  Qed.
+
+  Lemma big_sepL_fmap {B} (f : A → B) (Φ : nat → B → uPred M) l :
+    ([★ list] k↦y ∈ f <$> l, Φ k y) ⊣⊢ ([★ list] k↦y ∈ l, Φ k (f y)).
+  Proof. by rewrite /uPred_big_sepL imap_fmap. Qed.
+
+  Lemma big_sepL_sepL Φ Ψ l :
+    ([★ list] k↦x ∈ l, Φ k x ★ Ψ k x)
+    ⊣⊢ ([★ list] k↦x ∈ l, Φ k x) ★ ([★ list] k↦x ∈ l, Ψ k x).
+  Proof.
+    revert Φ Ψ; induction l as [|x l IH]=> Φ Ψ.
+    { by rewrite !big_sepL_nil left_id. }
+    rewrite !big_sepL_cons IH.
+    by rewrite -!assoc (assoc _ (Ψ _ _)) [(Ψ _ _ ★ _)%I]comm -!assoc.
+  Qed.
+
+  Lemma big_sepL_later Φ l :
+    ▷ ([★ list] k↦x ∈ l, Φ k x) ⊣⊢ ([★ list] k↦x ∈ l, ▷ Φ k x).
+  Proof.
+    revert Φ. induction l as [|x l IH]=> Φ.
+    { by rewrite !big_sepL_nil later_True. }
+    by rewrite !big_sepL_cons later_sep IH.
+  Qed.
+
+  Lemma big_sepL_always Φ l :
+    (□ [★ list] k↦x ∈ l, Φ k x) ⊣⊢ ([★ list] k↦x ∈ l, □ Φ k x).
+  Proof.
+    revert Φ. induction l as [|x l IH]=> Φ.
+    { by rewrite !big_sepL_nil always_pure. }
+    by rewrite !big_sepL_cons always_sep IH.
+  Qed.
+
+  Lemma big_sepL_always_if p Φ l :
+    □?p ([★ list] k↦x ∈ l, Φ k x) ⊣⊢ ([★ list] k↦x ∈ l, □?p Φ k x).
+  Proof. destruct p; simpl; auto using big_sepL_always. Qed.
+
+  Lemma big_sepL_forall Φ l :
+    (∀ k x, PersistentP (Φ k x)) →
+    ([★ list] k↦x ∈ l, Φ k x) ⊣⊢ (∀ k x, l !! k = Some x → Φ k x).
+  Proof.
+    intros HΦ. apply (anti_symm _).
+    { apply forall_intro=> k; apply forall_intro=> x.
+      apply impl_intro_l, pure_elim_l=> ?; by apply big_sepL_lookup. }
+    revert Φ HΦ. induction l as [|x l IH]=> Φ HΦ.
+    { rewrite big_sepL_nil; auto with I. }
+    rewrite big_sepL_cons. rewrite -always_and_sep_l; apply and_intro.
+    - by rewrite (forall_elim 0) (forall_elim x) pure_equiv // True_impl.
+    - rewrite -IH. apply forall_intro=> k; by rewrite (forall_elim (S k)).
+  Qed.
+
+  Lemma big_sepL_impl Φ Ψ l :
+    □ (∀ k x, l !! k = Some x → Φ k x → Ψ k x) ∧ ([★ list] k↦x ∈ l, Φ k x)
+    ⊢ [★ list] k↦x ∈ l, Ψ k x.
+  Proof.
+    rewrite always_and_sep_l. do 2 setoid_rewrite always_forall.
+    setoid_rewrite always_impl; setoid_rewrite always_pure.
+    rewrite -big_sepL_forall -big_sepL_sepL. apply big_sepL_mono; auto=> k x ?.
+    by rewrite -always_wand_impl always_elim wand_elim_l.
+  Qed.
+
+  Global Instance big_sepL_nil_persistent Φ :
+    PersistentP ([★ list] k↦x ∈ [], Φ k x).
+  Proof. rewrite /uPred_big_sepL. apply _. Qed.
+  Global Instance big_sepL_persistent Φ l :
+    (∀ k x, PersistentP (Φ k x)) → PersistentP ([★ list] k↦x ∈ l, Φ k x).
+  Proof. rewrite /uPred_big_sepL. apply _. Qed.
+
+  Global Instance big_sepL_nil_timeless Φ :
+    TimelessP ([★ list] k↦x ∈ [], Φ k x).
+  Proof. rewrite /uPred_big_sepL. apply _. Qed.
+  Global Instance big_sepL_timeless Φ l :
+    (∀ k x, TimelessP (Φ k x)) → TimelessP ([★ list] k↦x ∈ l, Φ k x).
+  Proof. rewrite /uPred_big_sepL. apply _. Qed.
+End list.
+
+
 (** ** Big ops over finite maps *)
 Section gmap.
   Context `{Countable K} {A : Type}.
@@ -254,8 +473,23 @@ Section gmap.
     rewrite -big_sepM_forall -big_sepM_sepM. apply big_sepM_mono; auto=> k x ?.
     by rewrite -always_wand_impl always_elim wand_elim_l.
   Qed.
+
+  Global Instance big_sepM_empty_persistent Φ :
+    PersistentP ([★ map] k↦x ∈ ∅, Φ k x).
+  Proof. rewrite /uPred_big_sepM map_to_list_empty. apply _. Qed.
+  Global Instance big_sepM_persistent Φ m :
+    (∀ k x, PersistentP (Φ k x)) → PersistentP ([★ map] k↦x ∈ m, Φ k x).
+  Proof. intros. apply big_sep_persistent, fmap_persistent=>-[??] /=; auto. Qed.
+
+  Global Instance big_sepM_nil_timeless Φ :
+    TimelessP ([★ map] k↦x ∈ ∅, Φ k x).
+  Proof. rewrite /uPred_big_sepM map_to_list_empty. apply _. Qed.
+  Global Instance big_sepM_timeless Φ m :
+    (∀ k x, TimelessP (Φ k x)) → TimelessP ([★ map] k↦x ∈ m, Φ k x).
+  Proof. intro. apply big_sep_timeless, fmap_timeless=> -[??] /=; auto. Qed.
 End gmap.
 
+
 (** ** Big ops over finite sets *)
 Section gset.
   Context `{Countable A}.
@@ -373,25 +607,17 @@ Section gset.
     rewrite -big_sepS_forall -big_sepS_sepS. apply big_sepS_mono; auto=> x ?.
     by rewrite -always_wand_impl always_elim wand_elim_l.
   Qed.
-End gset.
 
-(** ** Persistence *)
-Global Instance big_and_persistent Ps : PersistentL Ps → PersistentP ([∧] Ps).
-Proof. induction 1; apply _. Qed.
-Global Instance big_sep_persistent Ps : PersistentL Ps → PersistentP ([★] Ps).
-Proof. induction 1; apply _. Qed.
-
-Global Instance nil_persistent : PersistentL (@nil (uPred M)).
-Proof. constructor. Qed.
-Global Instance cons_persistent P Ps :
-  PersistentP P → PersistentL Ps → PersistentL (P :: Ps).
-Proof. by constructor. Qed.
-Global Instance app_persistent Ps Ps' :
-  PersistentL Ps → PersistentL Ps' → PersistentL (Ps ++ Ps').
-Proof. apply Forall_app_2. Qed.
-Global Instance zip_with_persistent {A B} (f : A → B → uPred M) xs ys :
-  (∀ x y, PersistentP (f x y)) → PersistentL (zip_with f xs ys).
-Proof.
-  unfold PersistentL=> ?; revert ys; induction xs=> -[|??]; constructor; auto.
-Qed.
+  Global Instance big_sepS_empty_persistent Φ : PersistentP ([★ set] x ∈ ∅, Φ x).
+  Proof. rewrite /uPred_big_sepS elements_empty. apply _. Qed.
+  Global Instance big_sepS_persistent Φ X :
+    (∀ x, PersistentP (Φ x)) → PersistentP ([★ set] x ∈ X, Φ x).
+  Proof. rewrite /uPred_big_sepS. apply _. Qed.
+
+  Global Instance big_sepS_nil_timeless Φ : TimelessP ([★ set] x ∈ ∅, Φ x).
+  Proof. rewrite /uPred_big_sepS elements_empty. apply _. Qed.
+  Global Instance big_sepS_timeless Φ X :
+    (∀ x, TimelessP (Φ x)) → TimelessP ([★ set] x ∈ X, Φ x).
+  Proof. rewrite /uPred_big_sepS. apply _. Qed.
+End gset.
 End big_op.
diff --git a/docs/algebra.tex b/docs/algebra.tex
index c22f0baa7a314ce5dd47baf87a313cb3484f29a8..e9b3208e454bd5e1316827bbea1a5be8f0b3c999 100644
--- a/docs/algebra.tex
+++ b/docs/algebra.tex
@@ -15,7 +15,7 @@ This definition varies slightly from the original one in~\cite{catlogic}.
     \All n. (\nequiv{n}) ~& \text{is an equivalence relation} \tagH{cofe-equiv} \\
     \All n, m.& n \geq m \Ra (\nequiv{n}) \subseteq (\nequiv{m}) \tagH{cofe-mono} \\
     \All x, y.& x = y \Lra (\All n. x \nequiv{n} y) \tagH{cofe-limit} \\
-    \All n, c.& \lim(c) \nequiv{n} c(n+1) \tagH{cofe-compl}
+    \All n, c.& \lim(c) \nequiv{n} c(n) \tagH{cofe-compl}
   \end{align*}
 \end{defn}
 
@@ -35,17 +35,17 @@ In order to solve the recursive domain equation in \Sref{sec:model} it is also e
   A function $f : \cofe \to \cofeB$ between two COFEs is \emph{non-expansive} (written $f : \cofe \nfn \cofeB$) if
   \[\All n, x \in \cofe, y \in \cofe. x \nequiv{n} y \Ra f(x) \nequiv{n} f(y) \]
   It is \emph{contractive} if
-  \[ \All n, x \in \cofe, y \in \cofe. (\All m < n. x \nequiv{m} y) \Ra f(x) \nequiv{n} f(x) \]
+  \[ \All n, x \in \cofe, y \in \cofe. (\All m < n. x \nequiv{m} y) \Ra f(x) \nequiv{n} f(y) \]
 \end{defn}
 Intuitively, applying a non-expansive function to some data will not suddenly introduce differences between seemingly equal data.
 Elements that cannot be distinguished by programs within $n$ steps remain indistinguishable after applying $f$.
-The reason that contractive functions are interesting is that for every contractive $f : \cofe \to \cofe$ with $\cofe$ inhabited, there exists a fixed-point $\fix(f)$ such that $\fix(f) = f(\fix(f))$.
+The reason that contractive functions are interesting is that for every contractive $f : \cofe \to \cofe$ with $\cofe$ inhabited, there exists a \emph{unique} fixed-point $\fix(f)$ such that $\fix(f) = f(\fix(f))$.
 
 \begin{defn}
   The category $\COFEs$ consists of COFEs as objects, and non-expansive functions as arrows.
 \end{defn}
 
-Note that $\COFEs$ is cartesian closed:
+Note that $\COFEs$ is cartesian closed. In particular:
 \begin{defn}
   Given two COFEs $\cofe$ and $\cofeB$, the set of non-expansive functions $\set{f : \cofe \nfn \cofeB}$ is itself a COFE with
   \begin{align*}
@@ -59,6 +59,7 @@ Note that $\COFEs$ is cartesian closed:
 \end{defn}
 The function space $(-) \nfn (-)$ is a locally non-expansive bifunctor.
 Note that the composition of non-expansive (bi)functors is non-expansive, and the composition of a non-expansive and a contractive (bi)functor is contractive.
+The reason contractive (bi)functors are interesting is that by America and Rutten's theorem~\cite{America-Rutten:JCSS89,birkedal:metric-space}, they have a unique\footnote{Uniqueness is not proven in Coq.} fixed-point.
 
 \subsection{RA}
 
@@ -211,7 +212,7 @@ Furthermore, discrete CMRAs can be turned into RAs by ignoring their COFE struct
 \end{defn}
 Note that every object/arrow in $\CMRAs$ is also an object/arrow of $\COFEs$.
 The notion of a locally non-expansive (or contractive) bifunctor naturally generalizes to bifunctors between these categories.
-
+%TODO: Discuss how we probably have a commuting square of functors between Set, RA, CMRA, COFE.
 
 %%% Local Variables: 
 %%% mode: latex
diff --git a/docs/constructions.tex b/docs/constructions.tex
index cf4114d7b525dea873d4bb0b180db09f46df1f42..eceb03316643585a2ee0c4b72ebe51d8a1645982 100644
--- a/docs/constructions.tex
+++ b/docs/constructions.tex
@@ -3,7 +3,7 @@
 
 \subsection{Next (type-level later)}
 
-Given a COFE $\cofe$, we define $\latert\cofe$ as follows:
+Given a COFE $\cofe$, we define $\latert\cofe$ as follows (using a datatype-like notation to define the type):
 \begin{align*}
   \latert\cofe \eqdef{}& \latertinj(x:\cofe) \\
   \latertinj(x) \nequiv{n} \latertinj(y) \eqdef{}& n = 0 \lor x \nequiv{n-1} y
@@ -62,7 +62,7 @@ Frame-preserving updates on the $M_i$ lift to the product:
 \subsection{Sum}
 \label{sec:summ}
 
-The \emph{sum CMRA} $\monoid_1 \csumm \monoid_2$ for any CMRAs $\monoid_1$ and $\monoid_2$ is defined as:
+The \emph{sum CMRA} $\monoid_1 \csumm \monoid_2$ for any CMRAs $\monoid_1$ and $\monoid_2$ is defined as (again, we use a datatype-like notation):
 \begin{align*}
   \monoid_1 \csumm \monoid_2 \eqdef{}& \cinl(\melt_1:\monoid_1) \mid \cinr(\melt_2:\monoid_2) \mid \bot \\
   \mval_n \eqdef{}& \setComp{\cinl(\melt_1)\!}{\!\melt_1 \in \mval'_n}
@@ -91,7 +91,7 @@ Crucially, the second rule allows us to \emph{swap} the ``side'' of the sum that
 \subsection{Finite partial function}
 \label{sec:fpfnm}
 
-Given some countable $K$ and some CMRA $\monoid$, the set of finite partial functions $K \fpfn \monoid$ is equipped with a COFE and CMRA structure by lifting everything pointwise.
+Given some infinite countable $K$ and some CMRA $\monoid$, the set of finite partial functions $K \fpfn \monoid$ is equipped with a COFE and CMRA structure by lifting everything pointwise.
 
 We obtain the following frame-preserving updates:
 \begin{mathpar}
@@ -104,36 +104,35 @@ We obtain the following frame-preserving updates:
   {\emptyset \mupd \setComp{[\gname \mapsto \melt]}{\gname \in K}}
 
   \inferH{fpfn-update}
-  {\melt \mupd \meltsB}
+  {\melt \mupd_\monoid \meltsB}
   {f[i \mapsto \melt] \mupd \setComp{ f[i \mapsto \meltB]}{\meltB \in \meltsB}}
 \end{mathpar}
-Remember that $\mval$ is the set of elements of a CMRA that are valid at \emph{all} step-indices.
+Above, $\mval$ refers to the validity of $\monoid$.
 
 $K \fpfn (-)$ is a locally non-expansive functor from $\CMRAs$ to $\CMRAs$.
 
 \subsection{Agreement}
 
 Given some COFE $\cofe$, we define $\agm(\cofe)$ as follows:
-\newcommand{\aginjc}{\mathrm{c}} % the "c" field of an agreement element
-\newcommand{\aginjV}{\mathrm{V}} % the "V" field of an agreement element
 \begin{align*}
-  \agm(\cofe) \eqdef{}& \record{\aginjc : \mathbb{N} \to \cofe , \aginjV : \SProp} \\
-  & \text{quotiented by} \\
-  \melt \equiv \meltB \eqdef{}& \melt.\aginjV = \meltB.\aginjV \land \All n. n \in \melt.\aginjV \Ra \melt.\aginjc(n) \nequiv{n} \meltB.\aginjc(n) \\
-  \melt \nequiv{n} \meltB \eqdef{}& (\All m \leq n. m \in \melt.\aginjV \Lra m \in \meltB.\aginjV) \land (\All m \leq n. m \in \melt.\aginjV \Ra \melt.\aginjc(m) \nequiv{m} \meltB.\aginjc(m)) \\
-  \mval_n \eqdef{}& \setComp{\melt \in \agm(\cofe)}{ n \in \melt.\aginjV \land \All m \leq n. \melt.\aginjc(n) \nequiv{m} \melt.\aginjc(m) } \\
+  \agm(\cofe) \eqdef{}& \set{(c, V) \in (\mathbb{N} \to \cofe) \times \SProp}/\ {\sim} \\[-0.2em]
+  \textnormal{where }& \melt \sim \meltB \eqdef{} \melt.V = \meltB.V \land 
+    \All n. n \in \melt.V \Ra \melt.c(n) \nequiv{n} \meltB.c(n)  \\
+%    \All n \in {\melt.V}.\, \melt.x \nequiv{n} \meltB.x \\
+  \melt \nequiv{n} \meltB \eqdef{}& (\All m \leq n. m \in \melt.V \Lra m \in \meltB.V) \land (\All m \leq n. m \in \melt.V \Ra \melt.c(m) \nequiv{m} \meltB.c(m)) \\
+  \mval_n \eqdef{}& \setComp{\melt \in \agm(\cofe)}{ n \in \melt.V \land \All m \leq n. \melt.c(n) \nequiv{m} \melt.c(m) } \\
   \mcore\melt \eqdef{}& \melt \\
-  \melt \mtimes \meltB \eqdef{}& (\melt.\aginjc, \setComp{n}{n \in \melt.\aginjV \land n \in \meltB.\aginjV \land \melt \nequiv{n} \meltB })
+  \melt \mtimes \meltB \eqdef{}& \left(\melt.c, \setComp{n}{n \in \melt.V \land n \in \meltB.V \land \melt \nequiv{n} \meltB }\right)
 \end{align*}
-Note that the carrier $\agm(\cofe)$ is a \emph{record} consisting of the two fields $\aginjc$ and $\aginjV$.
+%Note that the carrier $\agm(\cofe)$ is a \emph{record} consisting of the two fields $c$ and $V$.
 
 $\agm(-)$ is a locally non-expansive functor from $\COFEs$ to $\CMRAs$.
 
-You can think of the $\aginjc$ as a \emph{chain} of elements of $\cofe$ that has to converge only for $n \in \aginjV$ steps.
+You can think of the $c$ as a \emph{chain} of elements of $\cofe$ that has to converge only for $n \in V$ steps.
 The reason we store a chain, rather than a single element, is that $\agm(\cofe)$ needs to be a COFE itself, so we need to be able to give a limit for every chain of $\agm(\cofe)$.
-However, given such a chain, we cannot constructively define its limit: Clearly, the $\aginjV$ of the limit is the limit of the $\aginjV$ of the chain.
+However, given such a chain, we cannot constructively define its limit: Clearly, the $V$ of the limit is the limit of the $V$ of the chain.
 But what to pick for the actual data, for the element of $\cofe$?
-Only if $\aginjV = \mathbb{N}$ we have a chain of $\cofe$ that we can take a limit of; if the $\aginjV$ is smaller, the chain ``cancels'', \ie stops converging as we reach indices $n \notin \aginjV$.
+Only if $V = \mathbb{N}$ we have a chain of $\cofe$ that we can take a limit of; if the $V$ is smaller, the chain ``cancels'', \ie stops converging as we reach indices $n \notin V$.
 To mitigate this, we apply the usual construction to close a set; we go from elements of $\cofe$ to chains of $\cofe$.
 
 We define an injection $\aginj$ into $\agm(\cofe)$ as follows:
@@ -150,7 +149,7 @@ There are no interesting frame-preserving updates for $\agm(\cofe)$, but we can
 
 \subsection{Exclusive CMRA}
 
-Given a cofe $\cofe$, we define a CMRA $\exm(\cofe)$ such that at most one $x \in \cofe$ can be owned:
+Given a COFE $\cofe$, we define a CMRA $\exm(\cofe)$ such that at most one $x \in \cofe$ can be owned:
 \begin{align*}
   \exm(\cofe) \eqdef{}& \exinj(\cofe) + \bot \\
   \mval_n \eqdef{}& \setComp{\melt\in\exm(\cofe)}{\melt \neq \bot}
@@ -160,12 +159,12 @@ All cases of composition go to $\bot$.
   \mcore{\exinj(x)} \eqdef{}& \mnocore &
   \mcore{\bot} \eqdef{}& \bot
 \end{align*}
+Remember that $\mnocore$ is the ``dummy'' element in $\maybe\monoid$ indicating (in this case) that $\exinj(x)$ has no core.
+
 The step-indexed equivalence is inductively defined as follows:
 \begin{mathpar}
   \infer{x \nequiv{n} y}{\exinj(x) \nequiv{n} \exinj(y)}
 
-  \axiom{\munit \nequiv{n} \munit}
-
   \axiom{\bot \nequiv{n} \bot}
 \end{mathpar}
 $\exm(-)$ is a locally non-expansive functor from $\COFEs$ to $\CMRAs$.
@@ -376,7 +375,7 @@ We obtain the following frame-preserving update:
 \subsection{STS with tokens}
 \label{sec:stsmon}
 
-Given a state-transition system~(STS, \ie a directed graph) $(\STSS, {\stsstep} \subseteq \STSS \times \STSS)$, a set of tokens $\STST$, and a labeling $\STSL: \STSS \ra \wp(\STST)$ of \emph{protocol-owned} tokens for each state, we construct a monoid modeling an authoritative current state and permitting transitions given a \emph{bound} on the current state and a set of \emph{locally-owned} tokens.
+Given a state-transition system~(STS, \ie a directed graph) $(\STSS, {\stsstep} \subseteq \STSS \times \STSS)$, a set of tokens $\STST$, and a labeling $\STSL: \STSS \ra \wp(\STST)$ of \emph{protocol-owned} tokens for each state, we construct an RA modeling an authoritative current state and permitting transitions given a \emph{bound} on the current state and a set of \emph{locally-owned} tokens.
 
 The construction follows the idea of STSs as described in CaReSL \cite{caresl}.
 We first lift the transition relation to $\STSS \times \wp(\STST)$ (implementing a \emph{law of token conservation}) and define a stepping relation for the \emph{frame} of a given token set:
@@ -387,7 +386,7 @@ We first lift the transition relation to $\STSS \times \wp(\STST)$ (implementing
 
 We further define \emph{closed} sets of states (given a particular set of tokens) as well as the \emph{closure} of a set:
 \begin{align*}
-\STSclsd(S, T) \eqdef{}& \All s \in S. \STSL(s) \disj T \land \All s'. s \stsfstep{T} s' \Ra s' \in S \\
+\STSclsd(S, T) \eqdef{}& \All s \in S. \STSL(s) \disj T \land \left(\All s'. s \stsfstep{T} s' \Ra s' \in S\right) \\
 \upclose(S, T) \eqdef{}& \setComp{ s' \in \STSS}{\Exists s \in S. s \stsftrans{T} s' }
 \end{align*}
 
diff --git a/docs/derived.tex b/docs/derived.tex
index 8fa3c08cdc3127bcf5ba551db4bcdd4e762f1d4f..a1ed61190f44277591b06974d9bb413e8dc69ccb 100644
--- a/docs/derived.tex
+++ b/docs/derived.tex
@@ -1,6 +1,6 @@
 \section{Derived proof rules and other constructions}
 
-We will below abuse notation, using the \emph{term} meta-variables like $\val$ to range over (bound) \emph{variables} of the corresponding type..
+We will below abuse notation, using the \emph{term} meta-variables like $\val$ to range over (bound) \emph{variables} of the corresponding type.
 We omit type annotations in binders and equality, when the type is clear from context.
 We assume that the signature $\Sig$ embeds all the meta-level concepts we use, and their properties, into the logic.
 (The Coq formalization is a \emph{shallow embedding} of the logic, so we have direct access to all meta-level notions within the logic anyways.)
@@ -16,7 +16,7 @@ We collect here some important and frequently used derived proof rules.
   {\prop * \Exists\var.\propB \provesIff \Exists\var. \prop * \propB}
 
   \infer{}
-  {\prop * \Exists\var.\propB \proves \Exists\var. \prop * \propB}
+  {\prop * \All\var.\propB \proves \All\var. \prop * \propB}
 
   \infer{}
   {\always(\prop*\propB) \provesIff \always\prop * \always\propB}
@@ -221,9 +221,9 @@ We can derive some specialized forms of the lifting axioms for the operational s
 \begin{mathparpagebreakable}
   \infer[wp-lift-atomic-step]
   {\atomic(\expr_1) \and
-   \red(\expr_1, \state_1) \and
-   \All \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \pred(\expr_2,\state_2,\expr_\f)}
-  {\later\ownPhys{\state_1} * \later\All \val_2, \state_2, \expr_\f. \pred(\ofval(\val), \state_2, \expr_\f) \land \ownPhys{\state_2} \wand \prop[\val_2/\var] * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE} \proves \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}}
+   \red(\expr_1, \state_1)}
+  { {\begin{inbox}~~\later\ownPhys{\state_1} * \later\All \val_2, \state_2, \expr_\f. (\expr_1,\state_1 \step \ofval(\val),\state_2,\expr_\f)  \land \ownPhys{\state_2} \wand \prop[\val_2/\var] * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE} {}\\ \proves  \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}
+  \end{inbox}} }
 
   \infer[wp-lift-atomic-det-step]
   {\atomic(\expr_1) \and
@@ -238,50 +238,12 @@ We can derive some specialized forms of the lifting axioms for the operational s
   {\later ( \wpre{\expr_2}[\mask_1]{\Ret\var.\prop} * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE}) \proves \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}}
 \end{mathparpagebreakable}
 
-Furthermore, we derive some forms that directly involve view shifts and Hoare triples.
-\begin{mathparpagebreakable}
-  \infer[ht-lift-step]
-  {\mask_2 \subseteq \mask_1 \and
-   \toval(\expr_1) = \bot \and
-   \red(\expr_1, \state_1) \and
-   \All \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \pred(\expr_2,\state_2,\expr_\f) \\\\
-   \prop \vs[\mask_1][\mask_2] \later\ownPhys{\state_1} * \later\prop' \and
-   \All \expr_2, \state_2, \expr_\f. \pred(\expr_2, \state_2, \expr_\f) * \ownPhys{\state_2} * \prop' \vs[\mask_2][\mask_1] \propB_1 * \propB_2 \\\\
-   \All \expr_2, \state_2, \expr_\f. \hoare{\propB_1}{\expr_2}{\Ret\val.\propC}[\mask_1] \and
-   \All \expr_2, \state_2, \expr_\f. \hoare{\propB_2}{\expr_\f}{\Ret\any. \TRUE}[\top]}
-  { \hoare\prop{\expr_1}{\Ret\val.\propC}[\mask_1] }
-
-  \infer[ht-lift-atomic-step]
-  {\atomic(\expr_1) \and
-   \red(\expr_1, \state_1) \and
-   \All \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \pred(\expr_2,\state_2,\expr_\f) \\\\
-   \prop \vs[\mask_1][\mask_2] \later\ownPhys{\state_1} * \later\prop' \and
-   \All \expr_2, \state_2, \expr_\f. \hoare{\pred(\expr_2,\state_2,\expr_\f) * \prop}{\expr_\f}{\Ret\any. \TRUE}[\top]}
-  { \hoare{\later\ownPhys{\state_1} * \later\prop}{\expr_1}{\Ret\val.\Exists \state_2, \expr_\f. \ownPhys{\state_2} * \pred(\ofval(\expr_2),\state_2,\expr_\f)}[\mask_1] }
-
-  \infer[ht-lift-pure-step]
-  {\toval(\expr_1) = \bot \and
-   \All\state_1. \red(\expr_1, \state_1) \and
-   \All \state_1, \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \state_1 = \state_2 \land \pred(\expr_2,\expr_\f) \\\\
-   \All \expr_2, \expr_\f. \hoare{\pred(\expr_2,\expr_\f) * \prop}{\expr_2}{\Ret\val.\propB}[\mask_1] \and
-   \All \expr_2, \expr_\f. \hoare{\pred(\expr_2,\expr_\f) * \prop'}{\expr_\f}{\Ret\any. \TRUE}[\top]}
-  { \hoare{\later(\prop*\prop')}{\expr_1}{\Ret\val.\propB}[\mask_1] }
-
-  \infer[ht-lift-pure-det-step]
-  {\toval(\expr_1) = \bot \and
-   \All\state_1. \red(\expr_1, \state_1) \and
-   \All \state_1, \expr_2', \state_2, \expr_\f'. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \state_1 = \state_2 \land \expr_2 = \expr_2' \land \expr_\f = \expr_\f' \\\\
-   \hoare{\prop}{\expr_2}{\Ret\val.\propB}[\mask_1] \and
-   \hoare{\prop'}{\expr_\f}{\Ret\any. \TRUE}[\top]}
-  { \hoare{\later(\prop*\prop')}{\expr_1}{\Ret\val.\propB}[\mask_1] }
-\end{mathparpagebreakable}
-
 \subsection{Global functor and ghost ownership}
 
-Hereinafter we assume the global CMRA functor (served up as a parameter to Iris) is obtained from a family of functors $(F_i)_{i \in I}$ for some finite $I$ by picking
-\[ F(\cofe) \eqdef \prod_{i \in I} \textlog{GhName} \fpfn F_i(\cofe) \]
+Hereinafter we assume the global CMRA functor (served up as a parameter to Iris) is obtained from a family of functors $(\iFunc_i)_{i \in I}$ for some finite $I$ by picking
+\[ \iFunc(\cofe) \eqdef \prod_{i \in I} \textlog{GhName} \fpfn \iFunc_i(\cofe) \]
 We don't care so much about what concretely $\textlog{GhName}$ is, as long as it is countable and infinite.
-With $M_i \eqdef F_i(\iProp)$, we write $\ownGhost{\gname}{\melt : M_i}$ (or just $\ownGhost{\gname}{\melt}$ if $M_i$ is clear from the context) for $\ownGGhost{[i \mapsto [\gname \mapsto \melt]]}$.
+With $M_i \eqdef \iFunc_i(\iProp)$, we write $\ownGhost{\gname}{\melt : M_i}$ (or just $\ownGhost{\gname}{\melt}$ if $M_i$ is clear from the context) for $\ownGGhost{[i \mapsto [\gname \mapsto \melt]]}$.
 In other words, $\ownGhost{\gname}{\melt : M_i}$ asserts that in the current state of monoid $M_i$, the ``ghost location'' $\gname$ is allocated and we own piece $\melt$.
 
 From~\ruleref{pvs-update}, \ruleref{vs-update} and the frame-preserving updates in~\Sref{sec:prodm} and~\Sref{sec:fpfnm}, we have the following derived rules.
@@ -311,7 +273,7 @@ From~\ruleref{pvs-update}, \ruleref{vs-update} and the frame-preserving updates
 
 \subsection{Invariant identifier namespaces}
 
-Let $\namesp \ni \textlog{InvNamesp} \eqdef \textlog{list}(\textlog{InvName})$ be the type of \emph{namespaces} for invariant names.
+Let $\namesp \in \textlog{InvNamesp} \eqdef \textlog{list}(\textlog{InvName})$ be the type of \emph{namespaces} for invariant names.
 Notice that there is an injection $\textlog{namesp\_inj}: \textlog{InvNamesp} \ra \textlog{InvName}$.
 Whenever needed (in particular, for masks at view shifts and Hoare triples), we coerce $\namesp$ to its suffix-closure: \[\namecl\namesp \eqdef \setComp{\iname}{\Exists \namesp'. \iname = \textlog{namesp\_inj}(\namesp' \dplus \namesp)}\]
 We use the notation $\namesp.\iname$ for the namespace $[\iname] \dplus \namesp$.
diff --git a/docs/logic.tex b/docs/logic.tex
index 94b627f16a121040314d985cb3ea23dab7e99a6e..a2c4e4fd400371e6646609516e4c782fa507a66a 100644
--- a/docs/logic.tex
+++ b/docs/logic.tex
@@ -7,7 +7,7 @@ A \emph{language} $\Lang$ consists of a set \textdom{Expr} of \emph{expressions}
 \end{mathpar}
 \item There exists a \emph{primitive reduction relation} \[(-,- \step -,-,-) \subseteq \textdom{Expr} \times \textdom{State} \times \textdom{Expr} \times \textdom{State} \times (\textdom{Expr} \uplus \set{\bot})\]
   We will write $\expr_1, \state_1 \step \expr_2, \state_2$ for $\expr_1, \state_1 \step \expr_2, \state_2, \bot$. \\
-  A reduction $\expr_1, \state_1 \step \expr_2, \state_2, \expr_\f$ indicates that, when $\expr_1$ reduces to $\expr$, a \emph{new thread} $\expr_\f$ is forked off.
+  A reduction $\expr_1, \state_1 \step \expr_2, \state_2, \expr_\f$ indicates that, when $\expr_1$ reduces to $\expr_2$, a \emph{new thread} $\expr_\f$ is forked off.
 \item All values are stuck:
 \[ \expr, \_ \step  \_, \_, \_ \Ra \toval(\expr) = \bot \]
 \end{itemize}
@@ -40,7 +40,7 @@ For any language $\Lang$, we define the corresponding thread-pool semantics.
 
 \paragraph{Machine syntax}
 \[
-	\tpool \in \textdom{ThreadPool} \eqdef \bigcup_n \textdom{Exp}^n
+	\tpool \in \textdom{ThreadPool} \eqdef \bigcup_n \textdom{Expr}^n
 \]
 
 \judgment[Machine reduction]{\cfg{\tpool}{\state} \step
@@ -48,16 +48,17 @@ For any language $\Lang$, we define the corresponding thread-pool semantics.
 \begin{mathpar}
 \infer
   {\expr_1, \state_1 \step \expr_2, \state_2, \expr_\f \and \expr_\f \neq \bot}
-  {\cfg{\tpool \dplus [\expr_1] \dplus \tpool'}{\state} \step
-     \cfg{\tpool \dplus [\expr_2] \dplus \tpool' \dplus [\expr_\f]}{\state'}}
+  {\cfg{\tpool \dplus [\expr_1] \dplus \tpool'}{\state_1} \step
+     \cfg{\tpool \dplus [\expr_2] \dplus \tpool' \dplus [\expr_\f]}{\state_2}}
 \and\infer
   {\expr_1, \state_1 \step \expr_2, \state_2}
-  {\cfg{\tpool \dplus [\expr_1] \dplus \tpool'}{\state} \step
-     \cfg{\tpool \dplus [\expr_2] \dplus \tpool'}{\state'}}
+  {\cfg{\tpool \dplus [\expr_1] \dplus \tpool'}{\state_1} \step
+     \cfg{\tpool \dplus [\expr_2] \dplus \tpool'}{\state_2}}
 \end{mathpar}
 
 \clearpage
 \section{Logic}
+\label{sec:logic}
 
 To instantiate Iris, you need to define the following parameters:
 \begin{itemize}
@@ -269,7 +270,7 @@ In writing $\vctx, x:\type$, we presuppose that $x$ is not already declared in $
 	\infer{\vctx \proves \wtt{\melt}{\textlog{M}}}
 		{\vctx \proves \wtt{\ownGGhost{\melt}}{\Prop}}
 \and
-	\infer{\vctx \proves \wtt{\melt}{\textlog{M}}}
+	\infer{\vctx \proves \wtt{\melt}{\type} \and \text{$\type$ is a CMRA}}
 		{\vctx \proves \wtt{\mval(\melt)}{\Prop}}
 \and
 	\infer{\vctx \proves \wtt{\state}{\textlog{State}}}
@@ -382,15 +383,17 @@ This is entirely standard.
   {\vctx\mid\pfctx \proves \exists \var: \type.\; \prop \\
    \vctx,\var : \type\mid\pfctx , \prop \proves \propB}
   {\vctx\mid\pfctx \proves \propB}
-\and
-\infer[$\lambda$]
-  {}
-  {\pfctx \proves (\Lam\var: \type. \prop)(\term) =_{\type\to\type'} \prop[\term/\var]}
-\and
-\infer[$\mu$]
-  {}
-  {\pfctx \proves \mu\var: \type. \prop =_{\type} \prop[\mu\var: \type. \prop/\var]}
+% \and
+% \infer[$\lambda$]
+%   {}
+%   {\pfctx \proves (\Lam\var: \type. \prop)(\term) =_{\type\to\type'} \prop[\term/\var]}
+% \and
+% \infer[$\mu$]
+%   {}
+%   {\pfctx \proves \mu\var: \type. \prop =_{\type} \prop[\mu\var: \type. \prop/\var]}
 \end{mathparpagebreakable}
+Furthermore, we have the usual $\eta$ and $\beta$ laws for projections, $\lambda$ and $\mu$.
+
 
 \paragraph{Laws of (affine) bunched implications.}
 \begin{mathpar}
@@ -414,7 +417,7 @@ This is entirely standard.
 \begin{mathpar}
 \begin{array}{rMcMl}
 \ownGGhost{\melt} * \ownGGhost{\meltB} &\provesIff&  \ownGGhost{\melt \mtimes \meltB} \\
-\ownGGhost{\melt} &\provesIff& \mval(\melt) \\
+\ownGGhost{\melt} &\proves& \mval(\melt) \\
 \TRUE &\proves&  \ownGGhost{\munit}
 \end{array}
 \and
@@ -449,6 +452,7 @@ This is entirely standard.
   \later{(\prop * \propB)} &\provesIff& \later\prop * \later\propB
 \end{array}
 \end{mathpar}
+A type $\type$ being \emph{inhabited} means that $ \proves \wtt{\term}{\type}$ is derivable for some $\term$.
 
 \begin{mathpar}
 \infer
@@ -460,7 +464,7 @@ This is entirely standard.
 {\timeless{\ownGGhost\melt}}
 
 \infer
-{\text{$\melt$ is a discrete COFE element}}
+{\text{$\melt$ is an element of a discrete CMRA}}
 {\timeless{\mval(\melt)}}
 
 \infer{}
@@ -494,8 +498,8 @@ This is entirely standard.
   {\always{\prop} \proves \prop}
 \and
 \begin{array}[c]{rMcMl}
-  \always{(\prop * \propB)} &\proves& \always{(\prop \land \propB)} \\
-  \always{\prop} * \propB &\proves& \always{\prop} \land \propB \\
+  \always{(\prop \land \propB)} &\proves& \always{(\prop * \propB)} \\
+  \always{\prop} \land \propB &\proves& \always{\prop} * \propB \\
   \always{\later\prop} &\provesIff& \later\always{\prop} \\
 \end{array}
 \and
@@ -589,19 +593,18 @@ This is entirely standard.
 \begin{mathpar}
   \infer[wp-lift-step]
   {\mask_2 \subseteq \mask_1 \and
-   \toval(\expr_1) = \bot \and
-   \red(\expr_1, \state_1) \and
-   \All \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \pred(\expr_2,\state_2,\expr_\f)}
+   \toval(\expr_1) = \bot}
   { {\begin{inbox} % for some crazy reason, LaTeX is actually sensitive to the space between the "{ {" here and the "} }" below...
-        ~~\pvs[\mask_1][\mask_2] \later\ownPhys{\state_1} * \later\All \expr_2, \state_2, \expr_\f. \pred(\expr_2, \state_2, \expr_\f) \land {}\\\qquad\qquad\qquad\qquad\qquad \ownPhys{\state_2} \wand \pvs[\mask_2][\mask_1] \wpre{\expr_2}[\mask_1]{\Ret\var.\prop} * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE} {}\\\proves \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}
+        ~~\pvs[\mask_1][\mask_2] \Exists \state_1. \red(\expr_1,\state_1) \land \later\ownPhys{\state_1} * {}\\\qquad\qquad\qquad \later\All \expr_2, \state_2, \expr_\f. \left( (\expr_1, \state_1 \step \expr_2, \state_2, \expr_\f) \land \ownPhys{\state_2} \right) \wand \pvs[\mask_2][\mask_1] \wpre{\expr_2}[\mask_1]{\Ret\var.\prop} * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE}  {}\\\proves \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}
       \end{inbox}} }
-
+\\\\
   \infer[wp-lift-pure-step]
   {\toval(\expr_1) = \bot \and
    \All \state_1. \red(\expr_1, \state_1) \and
-   \All \state_1, \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \state_1 = \state_2 \land \pred(\expr_2,\expr_\f)}
-  {\later\All \expr_2, \expr_\f. \pred(\expr_2, \expr_\f)  \Ra \wpre{\expr_2}[\mask_1]{\Ret\var.\prop} * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE} \proves \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}}
+   \All \state_1, \expr_2, \state_2, \expr_\f. \expr_1,\state_1 \step \expr_2,\state_2,\expr_\f \Ra \state_1 = \state_2 }
+  {\later\All \state, \expr_2, \expr_\f. (\expr_1,\state \step \expr_2, \state,\expr_\f)  \Ra \wpre{\expr_2}[\mask_1]{\Ret\var.\prop} * \wpre{\expr_\f}[\top]{\Ret\any.\TRUE} \proves \wpre{\expr_1}[\mask_1]{\Ret\var.\prop}}
 \end{mathpar}
+Notice that primitive view shifts cover everything to their right, \ie $\pvs \prop * \propB \eqdef \pvs (\prop * \propB)$.
 
 Here we define $\wpre{\expr_\f}[\mask]{\Ret\var.\prop} \eqdef \TRUE$ if $\expr_\f = \bot$ (remember that our stepping relation can, but does not have to, define a forked-off expression).
 
diff --git a/docs/model.tex b/docs/model.tex
index f6ce778183b83d9714185250e1a298940cf0736e..9546c8d1110d57c2e6e2d746d3b9a9cf6776e173 100644
--- a/docs/model.tex
+++ b/docs/model.tex
@@ -26,7 +26,7 @@ We introduce an additional logical connective $\ownM\melt$, which will later be
 	\Sem{\vctx \proves \prop \Ra \propB : \Prop}_\gamma &\eqdef
 	\Lam \melt. \setComp{n}{\begin{aligned}
             \All m, \meltB.& m \leq n \land \melt \mincl \meltB \land \meltB \in \mval_m \Ra {} \\
-            & m \in \Sem{\vctx \proves \prop : \Prop}_\gamma(\melt) \Ra {}\\& m \in \Sem{\vctx \proves \propB : \Prop}_\gamma(\melt)\end{aligned}}\\
+            & m \in \Sem{\vctx \proves \prop : \Prop}_\gamma(\meltB) \Ra {}\\& m \in \Sem{\vctx \proves \propB : \Prop}_\gamma(\meltB)\end{aligned}}\\
 	\Sem{\vctx \proves \All x : \type. \prop : \Prop}_\gamma &\eqdef
 	\Lam \melt. \setComp{n}{ \All v \in \Sem{\type}. n \in \Sem{\vctx, x : \type \proves \prop : \Prop}_{\gamma[x \mapsto v]}(\melt) } \\
 	\Sem{\vctx \proves \Exists x : \type. \prop : \Prop}_\gamma &\eqdef
@@ -40,8 +40,8 @@ We introduce an additional logical connective $\ownM\melt$, which will later be
 	\Lam \melt. \setComp{n}{\begin{aligned}
             \All m, \meltB.& m \leq n \land  \melt\mtimes\meltB \in \mval_m \Ra {} \\
             & m \in \Sem{\vctx \proves \prop : \Prop}_\gamma(\meltB) \Ra {}\\& m \in \Sem{\vctx \proves \propB : \Prop}_\gamma(\melt\mtimes\meltB)\end{aligned}} \\
-        \Sem{\vctx \proves \ownM{\melt} : \Prop}_\gamma &\eqdef \Lam\meltB. \setComp{n}{\melt \mincl[n] \meltB}  \\
-        \Sem{\vctx \proves \mval(\melt) : \Prop}_\gamma &\eqdef \Lam\any. \setComp{n}{\melt \in \mval_n} \\
+        \Sem{\vctx \proves \ownM{\melt} : \Prop}_\gamma &\eqdef \Lam\meltB. \setComp{n}{\Sem{\vctx \proves \melt : \textlog{M}} \mincl[n] \meltB}  \\
+        \Sem{\vctx \proves \mval(\melt) : \Prop}_\gamma &\eqdef \Lam\any. \setComp{n}{\Sem{\vctx \proves \melt : \type} \in \mval_n} \\
 \end{align*}
 
 For every definition, we have to show all the side-conditions: The maps have to be non-expansive and monotone.
@@ -52,15 +52,18 @@ For every definition, we have to show all the side-conditions: The maps have to
 The first complicated task in building a model of full Iris is defining the semantic model of $\Prop$.
 We start by defining the functor that assembles the CMRAs we need to the global resource CMRA:
 \begin{align*}
-  \textdom{ResF}(\cofe^\op, \cofe) \eqdef{}& \record{\wld: \mathbb{N} \fpfn \agm(\latert \cofe), \pres: \exm(\textdom{State}), \ghostRes: \iFunc(\cofe^\op, \cofe)}
+  \textdom{ResF}(\cofe^\op, \cofe) \eqdef{}& \record{\wld: \mathbb{N} \fpfn \agm(\latert \cofe), \pres: \maybe{\exm(\textdom{State})}, \ghostRes: \iFunc(\cofe^\op, \cofe)}
 \end{align*}
-Remember that $\iFunc$ is the user-chosen bifunctor from $\COFEs$ to $\CMRAs$.
+Above, $\maybe\monoid$ is the monoid obtained by adding a unit to $\monoid$.
+(It's not a coincidence that we used the same notation for the range of the core; it's the same type either way: $\monoid + 1$.)
+Remember that $\iFunc$ is the user-chosen bifunctor from $\COFEs$ to $\CMRAs$ (see~\Sref{sec:logic}).
 $\textdom{ResF}(\cofe^\op, \cofe)$ is a CMRA by lifting the individual CMRAs pointwise.
-Furthermore, if $F$ is locally contractive, then so is $\textdom{ResF}$.
+Furthermore, since $\Sigma$ is locally contractive, so is $\textdom{ResF}$.
 
 Now we can write down the recursive domain equation:
 \[ \iPreProp \cong \UPred(\textdom{ResF}(\iPreProp, \iPreProp)) \]
-$\iPreProp$ is a COFE, which exists by America and Rutten's theorem~\cite{America-Rutten:JCSS89,birkedal:metric-space}.
+$\iPreProp$ is a COFE defined as the fixed-point of a locally contractive bifunctor.
+This fixed-point exists and is unique by America and Rutten's theorem~\cite{America-Rutten:JCSS89,birkedal:metric-space}.
 We do not need to consider how the object is constructed. 
 We only need the isomorphism, given by
 \begin{align*}
@@ -105,9 +108,9 @@ $\textdom{wp}$ is defined as the fixed-point of a contractive function.
 \begin{align*}
   \textdom{pre-wp}(\textdom{wp})(\mask, \expr, \pred) &\eqdef \Lam\rs. \setComp{n}{\begin{aligned}
         \All &\rs_\f, m, \mask_\f, \state. 0 \leq m < n \land \mask \disj \mask_\f \land m+1 \in \wsat\state{\mask \cup \mask_\f}{\rs \mtimes \rs_\f} \Ra {}\\
-        &(\All\val. \toval(\expr) = \val \Ra \Exists \rsB. m+1 \in \pred(\rsB) \land m+1 \in \wsat\state{\mask \cup \mask_\f}{\rsB \mtimes \rs_\f}) \land {}\\
+        &(\All\val. \toval(\expr) = \val \Ra \Exists \rsB. m+1 \in \pred(\val)(\rsB) \land m+1 \in \wsat\state{\mask \cup \mask_\f}{\rsB \mtimes \rs_\f}) \land {}\\
         &(\toval(\expr) = \bot \land 0 < m \Ra \red(\expr, \state) \land \All \expr_2, \state_2, \expr_\f. \expr,\state \step \expr_2,\state_2,\expr_\f \Ra {}\\
-        &\qquad \Exists \rsB_1, \rsB_2. m \in \wsat\state{\mask \cup \mask_\f}{\rsB \mtimes \rs_\f} \land  m \in \textdom{wp}(\mask, \expr_2, \pred)(\rsB_1) \land {}&\\
+        &\qquad \Exists \rsB_1, \rsB_2. m \in \wsat\state{\mask \cup \mask_\f}{\rsB_1 \mtimes \rsB_2 \mtimes \rs_\f} \land  m \in \textdom{wp}(\mask, \expr_2, \pred)(\rsB_1) \land {}&\\
         &\qquad\qquad (\expr_\f = \bot \lor m \in \textdom{wp}(\top, \expr_\f, \Lam\any.\Lam\any.\mathbb{N})(\rsB_2))
     \end{aligned}} \\
   \textdom{wp}_\mask(\expr, \pred) &\eqdef \mathit{fix}(\textdom{pre-wp})(\mask, \expr, \pred)
@@ -150,7 +153,7 @@ The remaining domains are interpreted as follows:
 \Sem{\type \to \type'} &\eqdef& \Sem{\type} \nfn \Sem{\type} \\
 \end{array}
 \]
-For the remaining base types $\type$ defined by the signature $\Sig$, we pick an object $X_\type$ in $\cal U$ and define
+For the remaining base types $\type$ defined by the signature $\Sig$, we pick an object $X_\type$ in $\COFEs$ and define
 \[
 \Sem{\type} \eqdef X_\type
 \]
@@ -185,10 +188,10 @@ $\rho(x)\in\Sem{\vctx(x)}$.
 \paragraph{Logical entailment.}
 We can now define \emph{semantic} logical entailment.
 
-\typedsection{Interpretation of entailment}{\Sem{\vctx \mid \pfctx \proves \prop} : 2}
+\typedsection{Interpretation of entailment}{\Sem{\vctx \mid \pfctx \proves \prop} : \mProp}
 
 \[
-\Sem{\vctx \mid \pfctx \proves \propB} \eqdef
+\Sem{\vctx \mid \pfctx \proves \prop} \eqdef
 \begin{aligned}[t]
 \MoveEqLeft
 \forall n \in \mathbb{N}.\;
diff --git a/docs/setup.tex b/docs/setup.tex
index 8c5b92196c1351dd6338fe89b83b3394d4d77514..29030f00923474307b3dac7679f250a01263353e 100644
--- a/docs/setup.tex
+++ b/docs/setup.tex
@@ -14,7 +14,7 @@
 \usepackage{amssymb}
 \usepackage{stmaryrd}
 
-\usepackage{\basedir mathpartir}
+\usepackage{mathpartir}
 
 \usepackage{\basedir pftools}
 \usepackage{\basedir iris}
diff --git a/heap_lang/adequacy.v b/heap_lang/adequacy.v
new file mode 100644
index 0000000000000000000000000000000000000000..89990989d5a81b261b795a6498c9819173e489dd
--- /dev/null
+++ b/heap_lang/adequacy.v
@@ -0,0 +1,18 @@
+From iris.program_logic Require Export weakestpre adequacy.
+From iris.heap_lang Require Export heap.
+From iris.program_logic Require Import auth ownership.
+From iris.heap_lang Require Import proofmode notation.
+From iris.proofmode Require Import tactics weakestpre.
+
+Definition heapΣ : gFunctors := #[authΣ heapUR; irisΣ heap_lang].
+
+Definition heap_adequacy Σ `{irisPreG heap_lang Σ, authG Σ heapUR} e σ φ :
+  (∀ `{heapG Σ}, heap_ctx ⊢ WP e {{ v, ■ φ v }}) →
+  adequate e σ φ.
+Proof.
+  intros Hwp; eapply (wp_adequacy Σ); iIntros (?) "Hσ".
+  iVs (auth_alloc (ownP ∘ of_heap) heapN _ (to_heap σ) with "[Hσ]") as (γ) "[??]".
+  - auto using to_heap_valid.
+  - rewrite /= (from_to_heap σ); auto.
+  - iApply (Hwp (HeapG _ _ _ γ)). by rewrite /heap_ctx.
+Qed.
\ No newline at end of file
diff --git a/heap_lang/derived.v b/heap_lang/derived.v
index 7344167581e707950ee0f3e8f0c7e4f6f7edd8ee..cd65d6692e343290e2165cd6d0b448a70f0daba3 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 Φ :
@@ -63,12 +63,13 @@ Proof.
   destruct (bool_decide_reflect (n1 < n2)); by eauto with omega.
 Qed.
 
-Lemma wp_eq E (n1 n2 : Z) P Φ :
-  (n1 = n2 → P ⊢ ▷ |={E}=> Φ (LitV (LitBool true))) →
-  (n1 ≠ n2 → P ⊢ ▷ |={E}=> Φ (LitV (LitBool false))) →
-  P ⊢ WP BinOp EqOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
+Lemma wp_eq E e1 e2 v1 v2 P Φ :
+  to_val e1 = Some v1 → to_val e2 = Some v2 →
+  (v1 = v2 → P ⊢ ▷ |={E}=> Φ (LitV (LitBool true))) →
+  (v1 ≠ v2 → P ⊢ ▷ |={E}=> Φ (LitV (LitBool false))) →
+  P ⊢ WP BinOp EqOp e1 e2 @ E {{ Φ }}.
 Proof.
   intros. rewrite -wp_bin_op //; [].
-  destruct (bool_decide_reflect (n1 = n2)); by eauto with omega.
+  destruct (bool_decide_reflect (v1 = v2)); by eauto.
 Qed.
 End derived.
diff --git a/heap_lang/heap.v b/heap_lang/heap.v
index 65de4c44875bb15a58ad31bfb256a225b92706c6..597066ded0a206bd929ae9412eb29cc9c6ced61e 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,28 +13,27 @@ 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. *)
-Definition heapGF : gFunctor := authGF heapUR.
-
 Definition to_heap : state → heapUR := fmap (λ v, (1%Qp, DecAgree v)).
 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 +43,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 +51,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.
 
@@ -102,25 +99,6 @@ Section heap.
   Qed.
   Hint Resolve heap_store_valid.
 
-  (** Allocation *)
-  Lemma heap_alloc E σ :
-    authG heap_lang Σ heapUR → nclose heapN ⊆ 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 heap_mapsto_eq /heap_mapsto_def /heap_name.
-    induction σ as [|l v σ Hl IH] using map_ind.
-    { rewrite big_sepM_empty; apply True_intro. }
-    rewrite to_heap_insert big_sepM_insert //.
-    rewrite (insert_singleton_op (to_heap σ));
-      last by rewrite lookup_fmap Hl; auto.
-    by rewrite auth_own_op IH.
-  Qed.
-
   Context `{heapG Σ}.
 
   (** General properties of mapsto *)
@@ -140,30 +118,38 @@ Section heap.
     rewrite heap_mapsto_eq -auth_own_op op_singleton pair_op dec_agree_ne //.
     apply (anti_symm (⊢)); last by apply pure_elim_l.
     rewrite auth_own_valid gmap_validI (forall_elim l) lookup_singleton.
-    rewrite option_validI prod_validI frac_validI discrete_valid.
+    rewrite option_validI prod_validI !discrete_valid /=.
     by apply pure_elim_r.
   Qed.
 
-  Lemma heap_mapsto_op_split l q v : l ↦{q} v ⊣⊢ (l ↦{q/2} v ★ l ↦{q/2} v).
-  Proof. by rewrite heap_mapsto_op_eq Qp_div_2. Qed.
+  Lemma heap_mapsto_op_1 l q1 q2 v1 v2 :
+    l ↦{q1} v1 ★ l ↦{q2} v2 ⊢ v1 = v2 ∧ l ↦{q1+q2} v1.
+  Proof. by rewrite heap_mapsto_op. Qed.
+
+  Lemma heap_mapsto_op_half l q v1 v2 :
+    l ↦{q/2} v1 ★ l ↦{q/2} v2 ⊣⊢ v1 = v2 ∧ l ↦{q} v1.
+  Proof. by rewrite heap_mapsto_op Qp_div_2. Qed.
+
+  Lemma heap_mapsto_op_half_1 l q v1 v2 :
+    l ↦{q/2} v1 ★ l ↦{q/2} v2 ⊢ v1 = v2 ∧ l ↦{q} v1.
+  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] !==>".
+    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 Φ :
@@ -171,14 +157,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 !==>". 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 Φ :
@@ -186,15 +173,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 !==>".
+    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 Φ :
@@ -204,12 +194,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 !==>". 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 Φ :
@@ -219,12 +210,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 !==>".
+    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/lang.v b/heap_lang/lang.v
index 881e486cc2d0b59fc0da2ec654792352496b8ab6..8f6bac10b0424eb691f7303a195af81432b429a2 100644
--- a/heap_lang/lang.v
+++ b/heap_lang/lang.v
@@ -59,24 +59,6 @@ Inductive expr :=
   | CAS (e0 : expr) (e1 : expr) (e2 : expr).
 
 Bind Scope expr_scope with expr.
-Delimit Scope expr_scope with E.
-Arguments Rec _ _ _%E.
-Arguments App _%E _%E.
-Arguments Lit _.
-Arguments UnOp _ _%E.
-Arguments BinOp _ _%E _%E.
-Arguments If _%E _%E _%E.
-Arguments Pair _%E _%E.
-Arguments Fst _%E.
-Arguments Snd _%E.
-Arguments InjL _%E.
-Arguments InjR _%E.
-Arguments Case _%E _%E _%E.
-Arguments Fork _%E.
-Arguments Alloc _%E.
-Arguments Load _%E.
-Arguments Store _%E _%E.
-Arguments CAS _%E _%E _%E.
 
 Fixpoint is_closed (X : list string) (e : expr) : bool :=
   match e with
@@ -105,10 +87,6 @@ Inductive val :=
   | InjRV (v : val).
 
 Bind Scope val_scope with val.
-Delimit Scope val_scope with V.
-Arguments PairV _%V _%V.
-Arguments InjLV _%V.
-Arguments InjRV _%V.
 
 Fixpoint of_val (v : val) : expr :=
   match v with
@@ -133,6 +111,40 @@ Fixpoint to_val (e : expr) : option val :=
 (** The state: heaps of vals. *)
 Definition state := gmap loc val.
 
+(** Equality and other typeclass stuff *)
+Lemma to_of_val v : to_val (of_val v) = Some v.
+Proof.
+  by induction v; simplify_option_eq; repeat f_equal; try apply (proof_irrel _).
+Qed.
+
+Lemma of_to_val e v : to_val e = Some v → of_val v = e.
+Proof.
+  revert v; induction e; intros v ?; simplify_option_eq; auto with f_equal.
+Qed.
+
+Instance of_val_inj : Inj (=) (=) of_val.
+Proof. by intros ?? Hv; apply (inj Some); rewrite -!to_of_val Hv. Qed.
+
+Instance base_lit_dec_eq (l1 l2 : base_lit) : Decision (l1 = l2).
+Proof. solve_decision. Defined.
+Instance un_op_dec_eq (op1 op2 : un_op) : Decision (op1 = op2).
+Proof. solve_decision. Defined.
+Instance bin_op_dec_eq (op1 op2 : bin_op) : Decision (op1 = op2).
+Proof. solve_decision. Defined.
+Instance expr_dec_eq (e1 e2 : expr) : Decision (e1 = e2).
+Proof. solve_decision. Defined.
+Instance val_dec_eq (v1 v2 : val) : Decision (v1 = v2).
+Proof.
+ refine (cast_if (decide (of_val v1 = of_val v2))); abstract naive_solver.
+Defined.
+
+Instance expr_inhabited : Inhabited expr := populate (Lit LitUnit).
+Instance val_inhabited : Inhabited val := populate (LitV LitUnit).
+
+Canonical Structure stateC := leibnizC state.
+Canonical Structure valC := leibnizC val.
+Canonical Structure exprC := leibnizC expr.
+
 (** Evaluation contexts *)
 Inductive ectx_item :=
   | AppLCtx (e2 : expr)
@@ -208,85 +220,74 @@ Definition subst' (mx : binder) (es : expr) : expr → expr :=
   match mx with BNamed x => subst x es | BAnon => id end.
 
 (** The stepping relation *)
-Definition un_op_eval (op : un_op) (l : base_lit) : option base_lit :=
-  match op, l with
-  | NegOp, LitBool b => Some (LitBool (negb b))
-  | MinusUnOp, LitInt n => Some (LitInt (- n))
+Definition un_op_eval (op : un_op) (v : val) : option val :=
+  match op, v with
+  | NegOp, LitV (LitBool b) => Some $ LitV $ LitBool (negb b)
+  | MinusUnOp, LitV (LitInt n) => Some $ LitV $ LitInt (- n)
   | _, _ => None
   end.
 
-Definition bin_op_eval (op : bin_op) (l1 l2 : base_lit) : option base_lit :=
-  match op, l1, l2 with
-  | PlusOp, LitInt n1, LitInt n2 => Some $ LitInt (n1 + n2)
-  | MinusOp, LitInt n1, LitInt n2 => Some $ LitInt (n1 - n2)
-  | LeOp, LitInt n1, LitInt n2 => Some $ LitBool $ bool_decide (n1 ≤ n2)
-  | LtOp, LitInt n1, LitInt n2 => Some $ LitBool $ bool_decide (n1 < n2)
-  | EqOp, LitInt n1, LitInt n2 => Some $ LitBool $ bool_decide (n1 = n2)
+Definition bin_op_eval (op : bin_op) (v1 v2 : val) : option val :=
+  match op, v1, v2 with
+  | PlusOp, LitV (LitInt n1), LitV (LitInt n2) => Some $ LitV $ LitInt (n1 + n2)
+  | MinusOp, LitV (LitInt n1), LitV (LitInt n2) => Some $ LitV $ LitInt (n1 - n2)
+  | LeOp, LitV (LitInt n1), LitV (LitInt n2) => Some $ LitV $ LitBool $ bool_decide (n1 ≤ n2)
+  | LtOp, LitV (LitInt n1), LitV (LitInt n2) => Some $ LitV $ LitBool $ bool_decide (n1 < n2)
+  | EqOp, v1, v2 => Some $ LitV $ LitBool $ bool_decide (v1 = v2)
   | _, _, _ => None
   end.
 
-Inductive head_step : expr → state → expr → state → option (expr) → Prop :=
+Inductive head_step : expr → state → expr → state → list (expr) → Prop :=
   | BetaS f x e1 e2 v2 e' σ :
      to_val e2 = Some v2 →
      Closed (f :b: x :b: []) e1 →
      e' = subst' x (of_val v2) (subst' f (Rec f x e1) e1) →
-     head_step (App (Rec f x e1) e2) σ e' σ None
-  | UnOpS op l l' σ :
-     un_op_eval op l = Some l' → 
-     head_step (UnOp op (Lit l)) σ (Lit l') σ None
-  | BinOpS op l1 l2 l' σ :
-     bin_op_eval op l1 l2 = Some l' → 
-     head_step (BinOp op (Lit l1) (Lit l2)) σ (Lit l') σ None
+     head_step (App (Rec f x e1) e2) σ e' σ []
+  | UnOpS op e v v' σ :
+     to_val e = Some v →
+     un_op_eval op v = Some v' → 
+     head_step (UnOp op e) σ (of_val v') σ []
+  | BinOpS op e1 e2 v1 v2 v' σ :
+     to_val e1 = Some v1 → to_val e2 = Some v2 →
+     bin_op_eval op v1 v2 = Some v' → 
+     head_step (BinOp op e1 e2) σ (of_val v') σ []
   | IfTrueS e1 e2 σ :
-     head_step (If (Lit $ LitBool true) e1 e2) σ e1 σ None
+     head_step (If (Lit $ LitBool true) e1 e2) σ e1 σ []
   | IfFalseS e1 e2 σ :
-     head_step (If (Lit $ LitBool false) e1 e2) σ e2 σ None
+     head_step (If (Lit $ LitBool false) e1 e2) σ e2 σ []
   | FstS e1 v1 e2 v2 σ :
      to_val e1 = Some v1 → to_val e2 = Some v2 →
-     head_step (Fst (Pair e1 e2)) σ e1 σ None
+     head_step (Fst (Pair e1 e2)) σ e1 σ []
   | SndS e1 v1 e2 v2 σ :
      to_val e1 = Some v1 → to_val e2 = Some v2 →
-     head_step (Snd (Pair e1 e2)) σ e2 σ None
+     head_step (Snd (Pair e1 e2)) σ e2 σ []
   | CaseLS e0 v0 e1 e2 σ :
      to_val e0 = Some v0 →
-     head_step (Case (InjL e0) e1 e2) σ (App e1 e0) σ None
+     head_step (Case (InjL e0) e1 e2) σ (App e1 e0) σ []
   | CaseRS e0 v0 e1 e2 σ :
      to_val e0 = Some v0 →
-     head_step (Case (InjR e0) e1 e2) σ (App e2 e0) σ None
+     head_step (Case (InjR e0) e1 e2) σ (App e2 e0) σ []
   | ForkS e σ:
-     head_step (Fork e) σ (Lit LitUnit) σ (Some e)
+     head_step (Fork e) σ (Lit LitUnit) σ [e]
   | AllocS e v σ l :
      to_val e = Some v → σ !! l = None →
-     head_step (Alloc e) σ (Lit $ LitLoc l) (<[l:=v]>σ) None
+     head_step (Alloc e) σ (Lit $ LitLoc l) (<[l:=v]>σ) []
   | LoadS l v σ :
      σ !! l = Some v →
-     head_step (Load (Lit $ LitLoc l)) σ (of_val v) σ None
+     head_step (Load (Lit $ LitLoc l)) σ (of_val v) σ []
   | StoreS l e v σ :
      to_val e = Some v → is_Some (σ !! l) →
-     head_step (Store (Lit $ LitLoc l) e) σ (Lit LitUnit) (<[l:=v]>σ) None
+     head_step (Store (Lit $ LitLoc l) e) σ (Lit LitUnit) (<[l:=v]>σ) []
   | CasFailS l e1 v1 e2 v2 vl σ :
      to_val e1 = Some v1 → to_val e2 = Some v2 →
      σ !! l = Some vl → vl ≠ v1 →
-     head_step (CAS (Lit $ LitLoc l) e1 e2) σ (Lit $ LitBool false) σ None
+     head_step (CAS (Lit $ LitLoc l) e1 e2) σ (Lit $ LitBool false) σ []
   | CasSucS l e1 v1 e2 v2 σ :
      to_val e1 = Some v1 → to_val e2 = Some v2 →
      σ !! l = Some v1 →
-     head_step (CAS (Lit $ LitLoc l) e1 e2) σ (Lit $ LitBool true) (<[l:=v2]>σ) None.
+     head_step (CAS (Lit $ LitLoc l) e1 e2) σ (Lit $ LitBool true) (<[l:=v2]>σ) [].
 
 (** Basic properties about the language *)
-Lemma to_of_val v : to_val (of_val v) = Some v.
-Proof.
-  by induction v; simplify_option_eq; repeat f_equal; try apply (proof_irrel _).
-Qed.
-
-Lemma of_to_val e v : to_val e = Some v → of_val v = e.
-Proof.
-  revert v; induction e; intros v ?; simplify_option_eq; auto with f_equal.
-Qed.
-
-Instance of_val_inj : Inj (=) (=) of_val.
-Proof. by intros ?? Hv; apply (inj Some); rewrite -!to_of_val Hv. Qed.
-
 Instance fill_item_inj Ki : Inj (=) (=) (fill_item Ki).
 Proof. destruct Ki; intros ???; simplify_eq/=; auto with f_equal. Qed.
 
@@ -294,11 +295,11 @@ Lemma fill_item_val Ki e :
   is_Some (to_val (fill_item Ki e)) → is_Some (to_val e).
 Proof. intros [v ?]. destruct Ki; simplify_option_eq; eauto. Qed.
 
-Lemma val_stuck e1 σ1 e2 σ2 ef : head_step e1 σ1 e2 σ2 ef → to_val e1 = None.
+Lemma val_stuck e1 σ1 e2 σ2 efs : head_step e1 σ1 e2 σ2 efs → to_val e1 = None.
 Proof. destruct 1; naive_solver. Qed.
 
-Lemma head_ctx_step_val Ki e σ1 e2 σ2 ef :
-  head_step (fill_item Ki e) σ1 e2 σ2 ef → is_Some (to_val e).
+Lemma head_ctx_step_val Ki e σ1 e2 σ2 efs :
+  head_step (fill_item Ki e) σ1 e2 σ2 efs → is_Some (to_val e).
 Proof. destruct Ki; inversion_clear 1; simplify_option_eq; by eauto. Qed.
 
 Lemma fill_item_no_val_inj Ki1 Ki2 e1 e2 :
@@ -313,7 +314,7 @@ Qed.
 
 Lemma alloc_fresh e v σ :
   let l := fresh (dom _ σ) in
-  to_val e = Some v → head_step (Alloc e) σ (Lit (LitLoc l)) (<[l:=v]>σ) None.
+  to_val e = Some v → head_step (Alloc e) σ (Lit (LitLoc l)) (<[l:=v]>σ) [].
 Proof. by intros; apply AllocS, (not_elem_of_dom (D:=gset _)), is_fresh. Qed.
 
 (** Closed expressions *)
@@ -334,27 +335,6 @@ Proof. intros. apply is_closed_subst with []; set_solver. Qed.
 
 Lemma is_closed_of_val X v : is_closed X (of_val v).
 Proof. apply is_closed_weaken_nil. induction v; simpl; auto. Qed.
-
-(** Equality and other typeclass stuff *)
-Instance base_lit_dec_eq (l1 l2 : base_lit) : Decision (l1 = l2).
-Proof. solve_decision. Defined.
-Instance un_op_dec_eq (op1 op2 : un_op) : Decision (op1 = op2).
-Proof. solve_decision. Defined.
-Instance bin_op_dec_eq (op1 op2 : bin_op) : Decision (op1 = op2).
-Proof. solve_decision. Defined.
-Instance expr_dec_eq (e1 e2 : expr) : Decision (e1 = e2).
-Proof. solve_decision. Defined.
-Instance val_dec_eq (v1 v2 : val) : Decision (v1 = v2).
-Proof.
- refine (cast_if (decide (of_val v1 = of_val v2))); abstract naive_solver.
-Defined.
-
-Instance expr_inhabited : Inhabited (expr) := populate (Lit LitUnit).
-Instance val_inhabited : Inhabited val := populate (LitV LitUnit).
-
-Canonical Structure stateC := leibnizC state.
-Canonical Structure valC := leibnizC val.
-Canonical Structure exprC := leibnizC expr.
 End heap_lang.
 
 (** Language *)
diff --git a/heap_lang/lib/assert.v b/heap_lang/lib/assert.v
index 3cc335765624ebfc9f5598ccb39d62b5ec3d1ae9..bae4e2c03d608dda2f547c1ed3e7e6f9e34a37db 100644
--- a/heap_lang/lib/assert.v
+++ b/heap_lang/lib/assert.v
@@ -1,3 +1,5 @@
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
 From iris.proofmode Require Import tactics.
 From iris.heap_lang Require Import proofmode notation.
 
@@ -7,7 +9,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 879ac67ca45340116ee364ed291821299a213aab..b427bb5b36f6322cdc4a199b011508b9e2798e62 100644
--- a/heap_lang/lib/barrier/proof.v
+++ b/heap_lang/lib/barrier/proof.v
@@ -1,57 +1,55 @@
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
+From iris.heap_lang.lib.barrier Require Export barrier.
 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. *)
+(** The CMRAs/functors 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 Σ.
-Proof. destruct H as (?&?&?). split; apply _. Qed.
+Definition barrierΣ : gFunctors := #[stsΣ sts; savedPropΣ idCF].
+
+Instance subG_barrierΣ {Σ} : subG barrierΣ Σ → barrierG Σ.
+Proof. intros [? [? _]%subG_inv]%subG_inv. 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,76 +91,74 @@ 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Φ".
-  iPvs (saved_prop_alloc (F:=idCF) _ P) as (γ) "#?".
-  iPvs (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]")
+  rewrite /newbarrier /=. wp_seq. wp_alloc l as "Hl".
+  iApply ("HΦ" 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. }
   iAssert (barrier_ctx γ' l P)%I as "#?".
   { rewrite /barrier_ctx. by repeat iSplit. }
   iAssert (sts_ownS γ' (i_states γ) {[Change γ]}
-    ★ sts_ownS γ' low_states {[Send]})%I with "|==>[-]" as "[Hr Hs]".
+    ★ 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.
+  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".
+  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γ".
-    iAssert (sts_ownS γ (i_states i) {[Change i]})%I with "|==>[Hγ]" as "Hγ".
+  iLöb as "IH". wp_rec. wp_bind (! _)%E.
+  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 +166,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.
@@ -199,13 +194,11 @@ Proof.
   rewrite /recv.
   iIntros "HP HP1"; iDestruct "HP1" as (γ P Q i) "(#Hctx&Hγ&Hi&HP1)".
   iExists γ, P, Q, i. iFrame "Hctx Hγ Hi".
-  iIntros "> HQ". by iApply "HP"; iApply "HP1".
+  iIntros "!> HQ". by iApply "HP"; iApply "HP1".
 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..93f70b4078af4a97b0a109e68a7989417723745f 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 }}) ∧
@@ -22,9 +20,9 @@ Proof.
   intros HN.
   exists (λ l, CofeMor (recv N l)), (λ l, CofeMor (send N l)).
   split_and?; simpl.
-  - iIntros (P) "#? ! _". iApply (newbarrier_spec _ P); eauto.
-  - iIntros (l P) "! [Hl HP]". by iApply signal_spec; iFrame "Hl HP".
-  - iIntros (l P) "! Hl". iApply wait_spec; iFrame "Hl"; eauto.
+  - iIntros (P) "#? !# _". iApply (newbarrier_spec _ P); eauto.
+  - iIntros (l P) "!# [Hl HP]". by iApply signal_spec; iFrame "Hl HP".
+  - iIntros (l P) "!# Hl". iApply wait_spec; iFrame "Hl"; eauto.
   - intros; by apply recv_split.
   - apply recv_weaken.
 Qed.
diff --git a/heap_lang/lib/counter.v b/heap_lang/lib/counter.v
index bd7c6a9754e3cf785a0ad743a3286ca27ec50d95..f318d1b6b13382341ad04e3419b5b773bffdd3c8 100644
--- a/heap_lang/lib/counter.v
+++ b/heap_lang/lib/counter.v
@@ -1,9 +1,8 @@
-(* Monotone counter, but using an explicit CMRA instead of auth *)
-From iris.program_logic Require Export global_functor.
+From iris.program_logic Require Export weakestpre.
+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 +13,18 @@ Definition read : val := λ: "l", !"l".
 Global Opaque newcounter inc get.
 
 (** The CMRA we need. *)
-Class counterG Σ := CounterG { counter_tokG :> authG heap_lang Σ mnatUR }.
-Definition counterGF : gFunctorList := [authGF mnatUR].
-Instance inGF_counterG `{H : inGFs heap_lang Σ counterGF} : counterG Σ.
-Proof. destruct H; split; apply _. Qed.
+Class counterG Σ := CounterG { counter_tokG :> authG Σ mnatUR }.
+Definition counterΣ : gFunctors := #[authΣ mnatUR].
+
+Instance subG_counterΣ {Σ} : subG counterΣ Σ → counterG Σ.
+Proof. intros [? _]%subG_inv. 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 +32,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]")
+  iIntros (?) "[#Hh HΦ]". rewrite /newcounter /=. wp_seq. wp_alloc l as "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.
+  wp_bind (! _)%E.
+  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_bind (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.
+  rewrite /read /=. wp_let.
+  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..0b764bb1ae480cc818132982ccc78cfc95589d17 100644
--- a/heap_lang/lib/lock.v
+++ b/heap_lang/lib/lock.v
@@ -1,81 +1,38 @@
-From iris.program_logic Require Export global_functor.
-From iris.proofmode Require Import invariants ghost_ownership.
-From iris.heap_lang Require Import proofmode notation.
-Import uPred.
+From iris.heap_lang Require Import heap notation.
+
+Structure lock Σ `{!heapG Σ} := Lock {
+  (* -- operations -- *)
+  newlock : val;
+  acquire : val;
+  release : val;
+  (* -- predicates -- *)
+  (* name is used to associate locked with is_lock *)
+  name : Type;
+  is_lock (N: namespace) (γ: name) (lock: val) (R: iProp Σ) : iProp Σ;
+  locked (γ: name) : iProp Σ;
+  (* -- general properties -- *)
+  is_lock_ne N γ lk n: Proper (dist n ==> dist n) (is_lock N γ lk);
+  is_lock_persistent N γ lk R : PersistentP (is_lock N γ lk R);
+  locked_timeless γ : TimelessP (locked γ);
+  locked_exclusive γ : locked γ ★ locked γ ⊢ False;
+  (* -- operation specs -- *)
+  newlock_spec N (R : iProp Σ) Φ :
+    heapN ⊥ N →
+    heap_ctx ★ R ★ (∀ l γ, is_lock N γ l R -★ Φ l) ⊢ WP newlock #() {{ Φ }};
+  acquire_spec N γ lk R (Φ : val → iProp Σ) :
+    is_lock N γ lk R ★ (locked γ -★ R -★ Φ #()) ⊢ WP acquire lk {{ Φ }};
+  release_spec N γ lk R (Φ : val → iProp Σ) :
+    is_lock N γ lk R ★ locked γ ★ R ★ Φ #() ⊢ WP release lk {{ Φ }}
+}.
+
+Arguments newlock {_ _} _.
+Arguments acquire {_ _} _.
+Arguments release {_ _} _.
+Arguments is_lock {_ _} _ _ _ _ _.
+Arguments locked {_ _} _ _.
+
+Existing Instances is_lock_ne is_lock_persistent locked_timeless.
+
+Instance is_lock_proper Σ `{!heapG Σ} (L: lock Σ) N lk R:
+  Proper ((≡) ==> (≡)) (is_lock L N lk R) := ne_proper _.
 
-Definition newlock : val := λ: <>, ref #false.
-Definition acquire : val :=
-  rec: "acquire" "l" :=
-    if: CAS "l" #false #true then #() else "acquire" "l".
-Definition release : val := λ: "l", "l" <- #false.
-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) }.
-Definition lockGF : gFunctorList := [GFunctor (constRF (exclR unitC))].
-Instance inGF_lockG `{H : inGFs heap_lang Σ 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 :=
-  (∃ b : bool, l ↦ #b ★ if b then True else own γ (Excl ()) ★ R)%I.
-
-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 :=
-  (∃ γ, heapN ⊥ N ∧ heap_ctx ∧
-        inv N (lock_inv γ l R) ∧ own γ (Excl ()))%I.
-
-Global Instance lock_inv_ne n γ l : Proper (dist n ==> dist n) (lock_inv γ l).
-Proof. solve_proper. Qed.
-Global Instance is_lock_ne n l : Proper (dist n ==> dist n) (is_lock l).
-Proof. solve_proper. Qed.
-Global Instance locked_ne n l : Proper (dist n ==> dist n) (locked l).
-Proof. solve_proper. Qed.
-
-(** The main proofs. *)
-Global Instance is_lock_persistent l R : PersistentP (is_lock l R).
-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) Φ :
-  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.
-  { iIntros ">". iExists false. by iFrame. }
-  iPvsIntro. iApply "HΦ". iExists γ; eauto.
-Qed.
-
-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.
-Qed.
-
-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.
-Qed.
-End proof.
diff --git a/heap_lang/lib/par.v b/heap_lang/lib/par.v
index 1d9cca643ee1a94959fa85bfcb199cfad0ffd5d4..1a91d7e7576e97668309ad7ef1add0f4c2c502a4 100644
--- a/heap_lang/lib/par.v
+++ b/heap_lang/lib/par.v
@@ -15,25 +15,24 @@ 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 _).
+  iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _).
   iApply wp_wand_l; iFrame "Hf2"; iIntros (v) "H2". wp_let.
   wp_apply join_spec; iFrame "Hl". iIntros (w) "H1".
   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 221ca9b806b725b91727bb981ab7a9d573687ede..ce437f6ca59ab00fdcc75726fc64d343da63542d 100644
--- a/heap_lang/lib/spawn.v
+++ b/heap_lang/lib/spawn.v
@@ -1,7 +1,8 @@
-From iris.program_logic Require Export global_functor.
-From iris.proofmode Require Import invariants ghost_ownership.
+From iris.program_logic Require Export weakestpre.
+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",
@@ -15,27 +16,23 @@ Definition join : val :=
     end.
 Global Opaque spawn join.
 
-(** The CMRA we need. *)
+(** The CMRA & functor we need. *)
 (* Not bundling heapG, as it may be shared with other users. *)
-Class spawnG Σ := SpawnG {
-  spawn_tokG :> inG heap_lang Σ (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 Σ.
-Proof. destruct H as (?&?). split. apply: inGF_inG. Qed.
+Class spawnG Σ := SpawnG { spawn_tokG :> inG Σ (exclR unitC) }.
+Definition spawnΣ : gFunctors := #[GFunctor (constRF (exclR unitC))].
+
+Instance subG_spawnΣ {Σ} : subG spawnΣ Σ → spawnG Σ.
+Proof. intros [?%subG_inG _]%subG_inv. split; apply _. 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,38 +46,36 @@ 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)
   ⊢ WP spawn e {{ Φ }}.
 Proof.
-  iIntros (<-%of_to_val ?) "(#Hh & Hf & HΦ)". rewrite /spawn.
+  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.
-  - 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.
+  wp_apply wp_fork; simpl. iSplitR "Hf".
+  - iVsIntro. wp_seq. iVsIntro. iApply "HΦ". rewrite /join_handle. eauto.
+  - wp_bind (f _). iApply wp_wand_l. iFrame "Hf"; iIntros (v) "Hv".
+    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_bind (! _)%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".
-    + iCombine "Hγ" "Hγ'" as "Hγ". iDestruct (@own_valid with "Hγ") as %[].
+    + 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/spin_lock.v b/heap_lang/lib/spin_lock.v
new file mode 100644
index 0000000000000000000000000000000000000000..870b48e8dae7befba469d621abf28f1005e83dba
--- /dev/null
+++ b/heap_lang/lib/spin_lock.v
@@ -0,0 +1,85 @@
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
+From iris.proofmode Require Import invariants tactics.
+From iris.heap_lang Require Import proofmode notation.
+From iris.algebra Require Import excl.
+From iris.heap_lang.lib Require Import lock.
+
+Definition newlock : val := λ: <>, ref #false.
+Definition acquire : val :=
+  rec: "acquire" "l" :=
+    if: CAS "l" #false #true then #() else "acquire" "l".
+Definition release : val := λ: "l", "l" <- #false.
+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 Σ (exclR unitC) }.
+Definition lockΣ : gFunctors := #[GFunctor (constRF (exclR unitC))].
+
+Instance subG_lockΣ {Σ} : subG lockΣ Σ → lockG Σ.
+Proof. intros [?%subG_inG _]%subG_inv. split; apply _. Qed.
+
+Section proof.
+  Context `{!heapG Σ, !lockG Σ} (N : namespace).
+
+  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 (γ : gname) (lk : val) (R : iProp Σ) : iProp Σ :=
+    (∃ l: loc, heapN ⊥ N ∧ heap_ctx ∧ lk = #l ∧ inv N (lock_inv γ l R))%I.
+
+  Definition locked (γ : gname): iProp Σ := own γ (Excl ()).
+
+  Lemma locked_exclusive (γ : gname) : locked γ ★ locked γ ⊢ False.
+  Proof. rewrite /locked -own_op own_valid. by iIntros (?). Qed.
+
+  Global Instance lock_inv_ne n γ l : Proper (dist n ==> dist n) (lock_inv γ l).
+  Proof. solve_proper. Qed.
+  Global Instance is_lock_ne n l : Proper (dist n ==> dist n) (is_lock γ l).
+  Proof. solve_proper. Qed.
+
+  (** The main proofs. *)
+  Global Instance is_lock_persistent γ l R : PersistentP (is_lock γ l R).
+  Proof. apply _. Qed.
+  Global Instance locked_timeless γ : TimelessP (locked γ).
+  Proof. apply _. Qed.
+
+  Lemma newlock_spec (R : iProp Σ) Φ :
+    heapN ⊥ N →
+    heap_ctx ★ R ★ (∀ lk γ, is_lock γ lk R -★ Φ lk) ⊢ WP newlock #() {{ Φ }}.
+  Proof.
+    iIntros (?) "(#Hh & HR & HΦ)". rewrite /newlock /=.
+    wp_seq. wp_alloc l as "Hl".
+    iVs (own_alloc (Excl ())) as (γ) "Hγ"; first done.
+    iVs (inv_alloc N _ (lock_inv γ l R) with "[-HΦ]") as "#?".
+    { iIntros "!>". iExists false. by iFrame. }
+    iVsIntro. iApply "HΦ". iExists l. eauto.
+  Qed.
+
+  Lemma acquire_spec γ lk R (Φ : val → iProp Σ) :
+    is_lock γ lk R ★ (locked γ -★ R -★ Φ #()) ⊢ WP acquire lk {{ Φ }}.
+  Proof.
+    iIntros "[Hl HΦ]". iDestruct "Hl" as (l) "(% & #? & % & #?)". subst.
+    iLöb as "IH". wp_rec. wp_bind (CAS _ _ _)%E.
+    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"). by iFrame.
+  Qed.
+
+  Lemma release_spec γ lk R (Φ : val → iProp Σ) :
+    is_lock γ lk R ★ locked γ ★ R ★ Φ #() ⊢ WP release lk {{ Φ }}.
+  Proof.
+    iIntros "(Hlock & Hlocked & HR & HΦ)".
+    iDestruct "Hlock" as (l) "(% & #? & % & #?)". subst.
+    rewrite /release /=. wp_let. iInv N as (b) "[Hl _]" "Hclose".
+    wp_store. iFrame "HΦ". iApply "Hclose". iNext. iExists false. by iFrame.
+  Qed.
+End proof.
+
+Definition spin_lock `{!heapG Σ, !lockG Σ} : lock Σ :=
+  {| lock.locked_exclusive := locked_exclusive; lock.newlock_spec := newlock_spec;
+     lock.acquire_spec := acquire_spec; lock.release_spec := release_spec |}.
diff --git a/heap_lang/lib/ticket_lock.v b/heap_lang/lib/ticket_lock.v
new file mode 100644
index 0000000000000000000000000000000000000000..9611cd2bb46752cc3eb0ed72aeda13a73dc42539
--- /dev/null
+++ b/heap_lang/lib/ticket_lock.v
@@ -0,0 +1,194 @@
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
+From iris.program_logic Require Import auth.
+From iris.proofmode Require Import invariants.
+From iris.heap_lang Require Import proofmode notation.
+From iris.algebra Require Import gset.
+From iris.heap_lang.lib Require Export lock.
+Import uPred.
+
+Definition wait_loop: val :=
+  rec: "wait_loop" "x" "lock" :=
+    let: "o" := !(Fst "lock") in
+    if: "x" = "o"
+      then #() (* my turn *)
+      else "wait_loop" "x" "lock".
+
+Definition newlock : val := λ: <>, ((* owner *) ref #0, (* next *) ref #0).
+
+Definition acquire : val :=
+  rec: "acquire" "lock" :=
+    let: "n" := !(Snd "lock") in
+    if: CAS (Snd "lock") "n" ("n" + #1)
+      then wait_loop "n" "lock"
+      else "acquire" "lock".
+
+Definition release : val :=
+  rec: "release" "lock" :=
+    let: "o" := !(Fst "lock") in
+    if: CAS (Fst "lock") "o" ("o" + #1)
+      then #()
+      else "release" "lock".
+
+Global Opaque newlock acquire release wait_loop.
+
+(** The CMRAs we need. *)
+Class tlockG Σ := TlockG {
+   tlock_G :> authG Σ (gset_disjUR nat);
+   tlock_exclG  :> inG Σ (exclR unitC)
+}.
+Definition tlockΣ : gFunctors :=
+  #[authΣ (gset_disjUR nat); GFunctor (constRF (exclR unitC))].
+
+Instance subG_tlockΣ {Σ} : subG tlockΣ Σ → tlockG Σ.
+Proof. intros [? [?%subG_inG _]%subG_inv]%subG_inv. split; apply _. Qed.
+
+Section proof.
+  Context `{!heapG Σ, !tlockG Σ} (N : namespace).
+
+  Definition tickets_inv (n: nat) (gs: gset_disjUR nat) : iProp Σ :=
+    (gs = GSet (seq_set 0 n))%I.
+
+  Definition lock_inv (γ1 γ2: gname) (lo ln: loc) (R : iProp Σ) : iProp Σ :=
+    (∃ (o n: nat),
+       lo ↦ #o ★ ln ↦ #n ★
+       auth_inv γ1 (tickets_inv n) ★
+       ((own γ2 (Excl ()) ★ R) ∨ auth_own γ1 (GSet {[ o ]})))%I.
+
+  Definition is_lock (γs: gname * gname) (l: val) (R: iProp Σ) : iProp Σ :=
+    (∃ (lo ln: loc),
+       heapN ⊥ N ∧ heap_ctx ∧
+       l = (#lo, #ln)%V ∧ inv N (lock_inv (fst γs) (snd γs) lo ln R))%I.
+
+  Definition issued (γs: gname * gname) (l : val) (x: nat) (R : iProp Σ) : iProp Σ :=
+    (∃ (lo ln: loc),
+       heapN ⊥ N ∧ heap_ctx ∧
+       l = (#lo, #ln)%V ∧ inv N (lock_inv (fst γs) (snd γs) lo ln R) ∧
+       auth_own (fst γs) (GSet {[ x ]}))%I.
+
+  Definition locked (γs: gname * gname) : iProp Σ := own (snd γs) (Excl ())%I.
+
+  Global Instance lock_inv_ne n γ1 γ2 lo ln :
+    Proper (dist n ==> dist n) (lock_inv γ1 γ2 lo ln).
+  Proof. solve_proper. Qed.
+  Global Instance is_lock_ne γs n l : Proper (dist n ==> dist n) (is_lock γs l).
+  Proof. solve_proper. Qed.
+  Global Instance is_lock_persistent γs l R : PersistentP (is_lock γs l R).
+  Proof. apply _. Qed.
+  Global Instance locked_timeless γs : TimelessP (locked γs).
+  Proof. apply _. Qed.
+
+  Lemma locked_exclusive (γs: gname * gname) : (locked γs ★ locked γs ⊢ False)%I.
+  Proof. rewrite /locked -own_op own_valid. by iIntros (?). Qed.
+
+  Lemma newlock_spec (R : iProp Σ) Φ :
+    heapN ⊥ N →
+    heap_ctx ★ R ★ (∀ lk γs, is_lock γs lk R -★ Φ lk) ⊢ WP newlock #() {{ Φ }}.
+  Proof.
+    iIntros (HN) "(#Hh & HR & HΦ)". rewrite /newlock /=.
+    wp_seq. wp_alloc lo as "Hlo". wp_alloc ln as "Hln".
+    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 lo ln R) with "[-HΦ]").
+    - iNext. rewrite /lock_inv.
+      iExists 0%nat, 0%nat.
+      iFrame.
+      iSplitL "Hγ1".
+      + rewrite /auth_inv.
+        iExists (GSet ∅).
+        by iFrame.
+      + iLeft. by iFrame.
+    - iVsIntro.
+      iApply ("HΦ" $! (#lo, #ln)%V (γ1, γ2)).
+      iExists lo, ln.
+      iSplit; by eauto.
+  Qed.
+
+  Lemma wait_loop_spec γs l x R (Φ : val → iProp Σ) :
+    issued γs l x R ★ (locked γs -★ R -★ Φ #()) ⊢ WP wait_loop #x l {{ Φ }}.
+  Proof.
+    iIntros "[Hl HΦ]". iDestruct "Hl" as (lo ln) "(% & #? & % & #? & Ht)".
+    iLöb as "IH". wp_rec. subst. wp_let. wp_proj. wp_bind (! _)%E.
+    iInv N as (o n) "[Hlo [Hln Ha]]" "Hclose".
+    wp_load. destruct (decide (x = o)) as [->|Hneq].
+    - iDestruct "Ha" as "[Hainv [[Ho HR] | Haown]]".
+      + iVs ("Hclose" with "[Hlo Hln Hainv Ht]").
+        { iNext. iExists o, n. iFrame. eauto. }
+        iVsIntro. wp_let. wp_op=>[_|[]] //.
+        wp_if. iVsIntro.
+        iApply ("HΦ" with "[-HR] HR"). eauto.
+      + iExFalso. iCombine "Ht" "Haown" as "Haown".
+        iDestruct (auth_own_valid with "Haown") as % ?%gset_disj_valid_op.
+        set_solver.
+    - iVs ("Hclose" with "[Hlo Hln Ha]").
+      { iNext. iExists o, n. by iFrame. }
+      iVsIntro. wp_let. wp_op=>[[/Nat2Z.inj //]|?].
+      wp_if. iApply ("IH" with "Ht"). by iExact "HΦ".
+  Qed.
+
+  Lemma acquire_spec γs l R (Φ : val → iProp Σ) :
+    is_lock γs l R ★ (locked γs -★ R -★ Φ #()) ⊢ WP acquire l {{ Φ }}.
+  Proof.
+    iIntros "[Hl HΦ]". iDestruct "Hl" as (lo ln) "(% & #? & % & #?)".
+    iLöb as "IH". wp_rec. wp_bind (! _)%E. subst. wp_proj.
+    iInv N as (o n) "[Hlo [Hln Ha]]" "Hclose".
+    wp_load. iVs ("Hclose" with "[Hlo Hln Ha]").
+    { iNext. iExists o, n. by iFrame. }
+    iVsIntro. wp_let. wp_proj. wp_op.
+    wp_bind (CAS _ _ _).
+    iInv N as (o' n') "[Hlo' [Hln' [Hainv Haown]]]" "Hclose".
+    destruct (decide (#n' = #n))%V as [[= ->%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 "[Hlo' Hln' 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.
+      iApply (wait_loop_spec γs (#lo, #ln)).
+      iSplitR "HΦ"; last by auto.
+      rewrite /issued /auth_own; eauto 10.
+    - wp_cas_fail.
+      iVs ("Hclose" with "[Hlo' Hln' Hainv Haown]").
+      { iNext. iExists o', n'. by iFrame. }
+      iVsIntro. wp_if. by iApply "IH".
+  Qed.
+
+  Lemma release_spec γs l R (Φ : val → iProp Σ):
+    is_lock γs l R ★ locked γs ★ R ★ Φ #() ⊢ WP release l {{ Φ }}.
+  Proof.
+    iIntros "(Hl & Hγ & HR & HΦ)". iDestruct "Hl" as (lo ln) "(% & #? & % & #?)".
+    iLöb as "IH". wp_rec. subst. wp_proj. wp_bind (! _)%E.
+    iInv N as (o n) "[Hlo [Hln Hr]]" "Hclose".
+    wp_load. iVs ("Hclose" with "[Hlo Hln Hr]").
+    { iNext. iExists o, n. by iFrame. }
+    iVsIntro. wp_let. wp_bind (CAS _ _ _ ).
+    wp_proj. wp_op.
+    iInv N as (o' n') "[Hlo' [Hln' Hr]]" "Hclose".
+    destruct (decide (#o' = #o))%V as [[= ->%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 "[Hlo' Hln' 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 "[Hlo' Hln' Hr]").
+      { iNext. iExists o', n'. by iFrame. }
+      iVsIntro. wp_if. by iApply ("IH" with "Hγ HR").
+  Qed.
+End proof.
+
+Typeclasses Opaque is_lock issued locked.
+
+Definition ticket_lock `{!heapG Σ, !tlockG Σ} : lock Σ :=
+  {| lock.locked_exclusive := locked_exclusive; lock.newlock_spec := newlock_spec;
+     lock.acquire_spec := acquire_spec; lock.release_spec := release_spec |}.
diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v
index 24eb4d3a6c6fe8c986870c08711745eea10f24e1..3c8b7494f90c50cc69872b9dc4441b901fb2c511 100644
--- a/heap_lang/lifting.v
+++ b/heap_lang/lifting.v
@@ -3,21 +3,44 @@ From iris.program_logic Require Import ownership ectx_lifting. (* for ownP *)
 From iris.heap_lang Require Export lang.
 From iris.heap_lang Require Import tactics.
 From iris.proofmode Require Import weakestpre.
+From iris.prelude Require Import fin_maps.
 Import uPred.
-Local Hint Extern 0 (head_reducible _ _) => do_head_step eauto 2.
+
+(** The tactic [inv_head_step] performs inversion on hypotheses of the shape
+[head_step]. The tactic will discharge head-reductions starting from values, and
+simplifies hypothesis related to conversions from and to values, and finite map
+operations. This tactic is slightly ad-hoc and tuned for proving our lifting
+lemmas. *)
+Ltac inv_head_step :=
+  repeat match goal with
+  | _ => progress simplify_map_eq/= (* simplify memory stuff *)
+  | H : to_val _ = Some _ |- _ => apply of_to_val in H
+  | H : head_step ?e _ _ _ _ |- _ =>
+     try (is_var e; fail 1); (* inversion yields many goals if [e] is a variable
+     and can thus better be avoided. *)
+     inversion H; subst; clear H
+  end.
+
+Local Hint Extern 0 (atomic _) => solve_atomic.
+Local Hint Extern 0 (head_reducible _ _) => eexists _, _, _; simpl.
+
+Local Hint Constructors head_step.
+Local Hint Resolve alloc_fresh.
+Local Hint Resolve to_of_val.
 
 Section lifting.
-Context {Σ : iFunctor}.
-Implicit Types P Q : iProp heap_lang Σ.
-Implicit Types Φ : val → iProp heap_lang Σ.
-Implicit Types ef : option expr.
+Context `{irisG heap_lang Σ}.
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val → iProp Σ.
+Implicit Types efs : list expr.
+Implicit Types σ : state.
 
 (** Bind. This bundles some arguments that wp_ectx_bind leaves as indices. *)
 Lemma wp_bind {E e} K Φ :
   WP e @ E {{ v, WP fill K (of_val v) @ E {{ Φ }} }} ⊢ WP fill K e @ E {{ Φ }}.
 Proof. exact: wp_ectx_bind. Qed.
 
-Lemma wp_bindi {E e} Ki Φ :
+Lemma wp_bind_ctxi {E e} Ki Φ :
   WP e @ E {{ v, WP fill_item Ki (of_val v) @ E {{ Φ }} }} ⊢
      WP fill_item Ki e @ E {{ Φ }}.
 Proof. exact: weakestpre.wp_bind. Qed.
@@ -28,18 +51,18 @@ 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.
+  subst v2. iSplit. iApply "HΦ"; by iSplit. by iApply big_sepL_nil.
 Qed.
 
 Lemma wp_load_pst E σ l v Φ :
   σ !! l = Some v →
   ▷ ownP σ ★ ▷ (ownP σ ={E}=★ Φ v) ⊢ WP Load (Lit (LitLoc l)) @ E {{ Φ }}.
 Proof.
-  intros. rewrite -(wp_lift_atomic_det_head_step σ v σ None) ?right_id //;
-    last (by intros; inv_head_step; eauto using to_of_val). solve_atomic.
+  intros. rewrite -(wp_lift_atomic_det_head_step' σ v σ); eauto 10.
+  intros; inv_head_step; eauto 10.
 Qed.
 
 Lemma wp_store_pst E σ l v v' Φ :
@@ -47,9 +70,8 @@ Lemma wp_store_pst E σ l v v' Φ :
   ▷ ownP σ ★ ▷ (ownP (<[l:=v]>σ) ={E}=★ Φ (LitV LitUnit))
   ⊢ WP Store (Lit (LitLoc l)) (of_val v) @ E {{ Φ }}.
 Proof.
-  intros ?.
-  rewrite-(wp_lift_atomic_det_head_step σ (LitV LitUnit) (<[l:=v]>σ) None)
-    ?right_id //; last (by intros; inv_head_step; eauto). solve_atomic.
+  intros. rewrite-(wp_lift_atomic_det_head_step' σ (LitV LitUnit) (<[l:=v]>σ)); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_cas_fail_pst E σ l v1 v2 v' Φ :
@@ -57,9 +79,8 @@ Lemma wp_cas_fail_pst E σ l v1 v2 v' Φ :
   ▷ ownP σ ★ ▷ (ownP σ ={E}=★ Φ (LitV $ LitBool false))
   ⊢ WP CAS (Lit (LitLoc l)) (of_val v1) (of_val v2) @ E {{ Φ }}.
 Proof.
-  intros ??.
-  rewrite -(wp_lift_atomic_det_head_step σ (LitV $ LitBool false) σ None)
-    ?right_id //; last (by intros; inv_head_step; eauto). solve_atomic.
+  intros. rewrite -(wp_lift_atomic_det_head_step' σ (LitV $ LitBool false) σ); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_cas_suc_pst E σ l v1 v2 Φ :
@@ -67,18 +88,18 @@ Lemma wp_cas_suc_pst E σ l v1 v2 Φ :
   ▷ ownP σ ★ ▷ (ownP (<[l:=v2]>σ) ={E}=★ Φ (LitV $ LitBool true))
   ⊢ WP CAS (Lit (LitLoc l)) (of_val v1) (of_val v2) @ E {{ Φ }}.
 Proof.
-  intros ?. rewrite -(wp_lift_atomic_det_head_step σ (LitV $ LitBool true)
-    (<[l:=v2]>σ) None) ?right_id //; last (by intros; inv_head_step; eauto).
-  solve_atomic.
+  intros. rewrite -(wp_lift_atomic_det_head_step' σ (LitV $ LitBool true)
+    (<[l:=v2]>σ)); eauto 10.
+  intros; inv_head_step; eauto.
 Qed.
 
 (** Base axioms for core primitives of the language: Stateless reductions *)
 Lemma wp_fork E e Φ :
   ▷ (|={E}=> Φ (LitV LitUnit)) ★ ▷ WP e {{ _, True }} ⊢ WP Fork e @ E {{ Φ }}.
 Proof.
-  rewrite -(wp_lift_pure_det_head_step (Fork e) (Lit LitUnit) (Some e)) //=;
-    last by intros; inv_head_step; eauto.
-  rewrite later_sep -(wp_value_pvs _ _ (Lit _)) //.
+  rewrite -(wp_lift_pure_det_head_step (Fork e) (Lit LitUnit) [e]) //=; eauto.
+  - by rewrite later_sep -(wp_value_pvs _ _ (Lit _)) // big_sepL_singleton.
+  - intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_rec E f x erec e1 e2 Φ :
@@ -87,70 +108,78 @@ Lemma wp_rec E f x erec e1 e2 Φ :
   Closed (f :b: x :b: []) erec →
   ▷ WP subst' x e2 (subst' f e1 erec) @ E {{ Φ }} ⊢ WP App e1 e2 @ E {{ Φ }}.
 Proof.
-  intros -> [v2 ?] ?. rewrite -(wp_lift_pure_det_head_step (App _ _)
-    (subst' x e2 (subst' f (Rec f x erec) erec)) None) //= ?right_id;
-    intros; inv_head_step; eauto.
+  intros -> [v2 ?] ?. rewrite -(wp_lift_pure_det_head_step' (App _ _)
+    (subst' x e2 (subst' f (Rec f x erec) erec))); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
-Lemma wp_un_op E op l l' Φ :
-  un_op_eval op l = Some l' →
-  ▷ (|={E}=> Φ (LitV l')) ⊢ WP UnOp op (Lit l) @ E {{ Φ }}.
+Lemma wp_un_op E op e v v' Φ :
+  to_val e = Some v →
+  un_op_eval op v = Some v' →
+  ▷ (|={E}=> Φ v') ⊢ WP UnOp op e @ E {{ Φ }}.
 Proof.
-  intros. rewrite -(wp_lift_pure_det_head_step (UnOp op _) (Lit l') None)
-    ?right_id -?wp_value_pvs //; intros; inv_head_step; eauto.
+  intros. rewrite -(wp_lift_pure_det_head_step' (UnOp op _) (of_val v'))
+    -?wp_value_pvs'; eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
-Lemma wp_bin_op E op l1 l2 l' Φ :
-  bin_op_eval op l1 l2 = Some l' →
-  ▷ (|={E}=> Φ (LitV l')) ⊢ WP BinOp op (Lit l1) (Lit l2) @ E {{ Φ }}.
+Lemma wp_bin_op E op e1 e2 v1 v2 v' Φ :
+  to_val e1 = Some v1 → to_val e2 = Some v2 →
+  bin_op_eval op v1 v2 = Some v' →
+  ▷ (|={E}=> Φ v') ⊢ WP BinOp op e1 e2 @ E {{ Φ }}.
 Proof.
-  intros Heval. rewrite -(wp_lift_pure_det_head_step (BinOp op _ _) (Lit l') None)
-    ?right_id -?wp_value_pvs //; intros; inv_head_step; eauto.
+  intros. rewrite -(wp_lift_pure_det_head_step' (BinOp op _ _) (of_val v'))
+    -?wp_value_pvs'; eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_if_true E e1 e2 Φ :
   ▷ WP e1 @ E {{ Φ }} ⊢ WP If (Lit (LitBool true)) e1 e2 @ E {{ Φ }}.
 Proof.
-  rewrite -(wp_lift_pure_det_head_step (If _ _ _) e1 None)
-    ?right_id //; intros; inv_head_step; eauto.
+  rewrite -(wp_lift_pure_det_head_step' (If _ _ _) e1); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_if_false E e1 e2 Φ :
   ▷ WP e2 @ E {{ Φ }} ⊢ WP If (Lit (LitBool false)) e1 e2 @ E {{ Φ }}.
 Proof.
-  rewrite -(wp_lift_pure_det_head_step (If _ _ _) e2 None)
-    ?right_id //; intros; inv_head_step; eauto.
+  rewrite -(wp_lift_pure_det_head_step' (If _ _ _) e2); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_fst E e1 v1 e2 Φ :
   to_val e1 = Some v1 → is_Some (to_val e2) →
   ▷ (|={E}=> Φ v1) ⊢ WP Fst (Pair e1 e2) @ E {{ Φ }}.
 Proof.
-  intros ? [v2 ?]. rewrite -(wp_lift_pure_det_head_step (Fst _) e1 None)
-    ?right_id -?wp_value_pvs //; intros; inv_head_step; eauto.
+  intros ? [v2 ?].
+  rewrite -(wp_lift_pure_det_head_step' (Fst _) e1) -?wp_value_pvs; eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_snd E e1 e2 v2 Φ :
   is_Some (to_val e1) → to_val e2 = Some v2 →
   ▷ (|={E}=> Φ v2) ⊢ WP Snd (Pair e1 e2) @ E {{ Φ }}.
 Proof.
-  intros [v1 ?] ?. rewrite -(wp_lift_pure_det_head_step (Snd _) e2 None)
-    ?right_id -?wp_value_pvs //; intros; inv_head_step; eauto.
+  intros [v1 ?] ?.
+  rewrite -(wp_lift_pure_det_head_step' (Snd _) e2) -?wp_value_pvs; eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_case_inl E e0 e1 e2 Φ :
   is_Some (to_val e0) →
   ▷ WP App e1 e0 @ E {{ Φ }} ⊢ WP Case (InjL e0) e1 e2 @ E {{ Φ }}.
 Proof.
-  intros [v0 ?]. rewrite -(wp_lift_pure_det_head_step (Case _ _ _)
-    (App e1 e0) None) ?right_id //; intros; inv_head_step; eauto.
+  intros [v0 ?].
+  rewrite -(wp_lift_pure_det_head_step' (Case _ _ _) (App e1 e0)); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 
 Lemma wp_case_inr E e0 e1 e2 Φ :
   is_Some (to_val e0) →
   ▷ WP App e2 e0 @ E {{ Φ }} ⊢ WP Case (InjR e0) e1 e2 @ E {{ Φ }}.
 Proof.
-  intros [v0 ?]. rewrite -(wp_lift_pure_det_head_step (Case _ _ _)
-    (App e2 e0) None) ?right_id //; intros; inv_head_step; eauto.
+  intros [v0 ?].
+  rewrite -(wp_lift_pure_det_head_step' (Case _ _ _) (App e2 e0)); eauto.
+  intros; inv_head_step; eauto.
 Qed.
 End lifting.
diff --git a/heap_lang/notation.v b/heap_lang/notation.v
index 05e173432ec7b604ec1ba7fbc5f8a765c6edd02b..600cca0a3679c6dfd00b3b5579c168ffe2669c08 100644
--- a/heap_lang/notation.v
+++ b/heap_lang/notation.v
@@ -1,22 +1,6 @@
 From iris.heap_lang Require Export derived.
 Export heap_lang.
 
-Arguments wp {_ _} _ _%E _.
-
-Notation "'WP' e @ E {{ Φ } }" := (wp E e%E Φ)
-  (at level 20, e, Φ at level 200,
-   format "'WP'  e  @  E  {{  Φ  } }") : uPred_scope.
-Notation "'WP' e {{ Φ } }" := (wp ⊤ e%E Φ)
-  (at level 20, e, Φ at level 200,
-   format "'WP'  e  {{  Φ  } }") : uPred_scope.
-
-Notation "'WP' e @ E {{ v , Q } }" := (wp E e%E (λ v, Q))
-  (at level 20, e, Q at level 200,
-   format "'WP'  e  @  E  {{  v ,  Q  } }") : uPred_scope.
-Notation "'WP' e {{ v , Q } }" := (wp ⊤ e%E (λ v, Q))
-  (at level 20, e, Q at level 200,
-   format "'WP'  e  {{  v ,  Q  } }") : uPred_scope.
-
 Coercion LitInt : Z >-> base_lit.
 Coercion LitBool : bool >-> base_lit.
 Coercion LitLoc : loc >-> base_lit.
diff --git a/heap_lang/proofmode.v b/heap_lang/proofmode.v
index 68f480a12d057f378f9e49a0453b5bbbcfc9eae6..1a0a66e79a9fafeab5e9c49694b3407742e6219d 100644
--- a/heap_lang/proofmode.v
+++ b/heap_lang/proofmode.v
@@ -7,13 +7,13 @@ 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).
-Proof. by rewrite /IntoSep heap_mapsto_op_split. Qed.
+Global Instance into_and_mapsto l q v :
+  IntoAnd false (l ↦{q} v) (l ↦{q/2} v) (l ↦{q/2} v).
+Proof. by rewrite /IntoAnd heap_mapsto_op_eq Qp_div_2. Qed.
 
 Lemma tac_wp_alloc Δ Δ' E j e v Φ :
   to_val e = Some v →
@@ -89,7 +89,7 @@ End heap.
 Tactic Notation "wp_apply" open_constr(lem) :=
   lazymatch goal with
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
-    wp_bind K; iApply lem; try iNext)
+    wp_bind_core K; iApply lem; try iNext)
   | _ => fail "wp_apply: not a 'wp'"
   end.
 
@@ -98,7 +98,7 @@ Tactic Notation "wp_alloc" ident(l) "as" constr(H) :=
   | |- _ ⊢ wp ?E ?e ?Q =>
     first
       [reshape_expr e ltac:(fun K e' =>
-         match eval hnf in e' with Alloc _ => wp_bind K end)
+         match eval hnf in e' with Alloc _ => wp_bind_core K end)
       |fail 1 "wp_alloc: cannot find 'Alloc' in" e];
     eapply tac_wp_alloc with _ H _;
       [let e' := match goal with |- to_val ?e' = _ => e' end in
@@ -121,7 +121,7 @@ Tactic Notation "wp_load" :=
   | |- _ ⊢ wp ?E ?e ?Q =>
     first
       [reshape_expr e ltac:(fun K e' =>
-         match eval hnf in e' with Load _ => wp_bind K end)
+         match eval hnf in e' with Load _ => wp_bind_core K end)
       |fail 1 "wp_load: cannot find 'Load' in" e];
     eapply tac_wp_load;
       [iAssumption || fail "wp_load: cannot find heap_ctx"
@@ -138,7 +138,7 @@ Tactic Notation "wp_store" :=
   | |- _ ⊢ wp ?E ?e ?Q =>
     first
       [reshape_expr e ltac:(fun K e' =>
-         match eval hnf in e' with Store _ _ => wp_bind K end)
+         match eval hnf in e' with Store _ _ => wp_bind_core K end)
       |fail 1 "wp_store: cannot find 'Store' in" e];
     eapply tac_wp_store;
       [let e' := match goal with |- to_val ?e' = _ => e' end in
@@ -158,7 +158,7 @@ Tactic Notation "wp_cas_fail" :=
   | |- _ ⊢ wp ?E ?e ?Q =>
     first
       [reshape_expr e ltac:(fun K e' =>
-         match eval hnf in e' with CAS _ _ _ => wp_bind K end)
+         match eval hnf in e' with CAS _ _ _ => wp_bind_core K end)
       |fail 1 "wp_cas_fail: cannot find 'CAS' in" e];
     eapply tac_wp_cas_fail;
       [let e' := match goal with |- to_val ?e' = _ => e' end in
@@ -180,7 +180,7 @@ Tactic Notation "wp_cas_suc" :=
   | |- _ ⊢ wp ?E ?e ?Q =>
     first
       [reshape_expr e ltac:(fun K e' =>
-         match eval hnf in e' with CAS _ _ _ => wp_bind K end)
+         match eval hnf in e' with CAS _ _ _ => wp_bind_core K end)
       |fail 1 "wp_cas_suc: cannot find 'CAS' in" e];
     eapply tac_wp_cas_suc;
       [let e' := match goal with |- to_val ?e' = _ => e' end in
diff --git a/heap_lang/tactics.v b/heap_lang/tactics.v
index da060db51554e3dd23928e7676e1ddd7cda84891..dfb46f6c6a4b10f3b3a92f3714eb292c3ffa7c7d 100644
--- a/heap_lang/tactics.v
+++ b/heap_lang/tactics.v
@@ -1,5 +1,4 @@
 From iris.heap_lang Require Export lang.
-From iris.prelude Require Import fin_maps.
 Import heap_lang.
 
 (** We define an alternative representation of expressions in which the
@@ -179,6 +178,7 @@ Definition atomic (e : expr) :=
   | Store e1 e2 => bool_decide (is_Some (to_val e1) ∧ is_Some (to_val e2))
   | CAS e0 e1 e2 =>
      bool_decide (is_Some (to_val e0) ∧ is_Some (to_val e1) ∧ is_Some (to_val e2))
+  | Fork _ => true
   (* Make "skip" atomic *)
   | App (Rec _ _ (Lit _)) (Lit _) => true
   | _ => false
@@ -228,7 +228,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 :=
@@ -243,22 +245,6 @@ Ltac simpl_subst :=
 Arguments W.to_expr : simpl never.
 Arguments subst : simpl never.
 
-(** The tactic [inv_head_step] performs inversion on hypotheses of the
-shape [head_step]. The tactic will discharge head-reductions starting
-from values, and simplifies hypothesis related to conversions from and
-to values, and finite map operations. This tactic is slightly ad-hoc
-and tuned for proving our lifting lemmas. *)
-Ltac inv_head_step :=
-  repeat match goal with
-  | _ => progress simplify_map_eq/= (* simplify memory stuff *)
-  | H : to_val _ = Some _ |- _ => apply of_to_val in H
-  | H : context [to_val (of_val _)] |- _ => rewrite to_of_val in H
-  | H : head_step ?e _ _ _ _ |- _ =>
-     try (is_var e; fail 1); (* inversion yields many goals if [e] is a variable
-     and can thus better be avoided. *)
-     inversion H; subst; clear H
-  end.
-
 (** The tactic [reshape_expr e tac] decomposes the expression [e] into an
 evaluation context [K] and a subexpression [e']. It calls the tactic [tac K e']
 for each possible decomposition until [tac] succeeds. *)
@@ -301,16 +287,3 @@ Ltac reshape_expr e tac :=
      | go (CasMCtx v0 e2 :: K) e1 ])
   | CAS ?e0 ?e1 ?e2 => go (CasLCtx e1 e2 :: K) e0
   end in go (@nil ectx_item) e.
-
-(** The tactic [do_head_step tac] solves goals of the shape [head_reducible] and
-[head_step] by performing a reduction step and uses [tac] to solve any
-side-conditions generated by individual steps. *)
-Tactic Notation "do_head_step" tactic3(tac) :=
-  try match goal with |- head_reducible _ _ => eexists _, _, _ end;
-  simpl;
-  match goal with
-  | |- head_step ?e1 ?σ1 ?e2 ?σ2 ?ef =>
-     first [apply alloc_fresh|econstructor];
-       (* solve [to_val] side-conditions *)
-       first [rewrite ?to_of_val; reflexivity|simpl_subst; tac; fast_done]
-  end.
diff --git a/heap_lang/wp_tactics.v b/heap_lang/wp_tactics.v
index 361b83a3a61a055095cf72f16f8c079a77a0a698..c4bed529737b1e10df0596f7ca7495331f08f75d 100644
--- a/heap_lang/wp_tactics.v
+++ b/heap_lang/wp_tactics.v
@@ -1,9 +1,8 @@
-From iris.algebra Require Export upred_tactics.
 From iris.heap_lang Require Export tactics derived.
 Import uPred.
 
 (** wp-specific helper tactics *)
-Ltac wp_bind K :=
+Ltac wp_bind_core K :=
   lazymatch eval hnf in K with
   | [] => idtac
   | _ => etrans; [|fast_by apply (wp_bind K)]; simpl
@@ -25,7 +24,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,13 +48,13 @@ 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" :=
   lazymatch goal with
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
-    wp_bind K; wp_value_head) || fail "wp_value: cannot find value in" e
+    wp_bind_core K; wp_value_head) || fail "wp_value: cannot find value in" e
   | _ => fail "wp_value: not a wp"
   end.
 
@@ -61,7 +64,7 @@ Tactic Notation "wp_rec" :=
     match eval hnf in e' with App ?e1 _ =>
 (* hnf does not reduce through an of_val *)
 (*      match eval hnf in e1 with Rec _ _ _ => *)
-    wp_bind K; etrans; [|eapply wp_rec; wp_done]; simpl_subst; wp_finish
+    wp_bind_core K; etrans; [|eapply wp_rec; wp_done]; simpl_subst; wp_finish
 (*      end *) end) || fail "wp_rec: cannot find 'Rec' in" e
   | _ => fail "wp_rec: not a 'wp'"
   end.
@@ -71,7 +74,7 @@ Tactic Notation "wp_lam" :=
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match eval hnf in e' with App ?e1 _ =>
 (*    match eval hnf in e1 with Rec BAnon _ _ => *)
-    wp_bind K; etrans; [|eapply wp_lam; wp_done]; simpl_subst; wp_finish
+    wp_bind_core K; etrans; [|eapply wp_lam; wp_done]; simpl_subst; wp_finish
 (*    end *) end) || fail "wp_lam: cannot find 'Lam' in" e
   | _ => fail "wp_lam: not a 'wp'"
   end.
@@ -82,14 +85,17 @@ Tactic Notation "wp_seq" := wp_let.
 Tactic Notation "wp_op" :=
   lazymatch goal with
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
-    match eval hnf in e' with
-    | BinOp LtOp _ _ => wp_bind K; apply wp_lt; wp_finish
-    | BinOp LeOp _ _ => wp_bind K; apply wp_le; wp_finish
-    | BinOp EqOp _ _ => wp_bind K; apply wp_eq; wp_finish
+    lazymatch eval hnf in e' with
+    | BinOp LtOp _ _ => wp_bind_core K; apply wp_lt; wp_finish
+    | BinOp LeOp _ _ => wp_bind_core K; apply wp_le; wp_finish
+    | BinOp EqOp _ _ =>
+       wp_bind_core K; eapply wp_eq; [wp_done|wp_done|wp_finish|wp_finish]
     | BinOp _ _ _ =>
-       wp_bind K; etrans; [|eapply wp_bin_op; try fast_done]; wp_finish
+       wp_bind_core K; etrans;
+         [|eapply wp_bin_op; [wp_done|wp_done|try fast_done]]; wp_finish
     | UnOp _ _ =>
-       wp_bind K; etrans; [|eapply wp_un_op; try fast_done]; wp_finish
+       wp_bind_core K; etrans;
+         [|eapply wp_un_op; [wp_done|try fast_done]]; wp_finish
     end) || fail "wp_op: cannot find 'BinOp' or 'UnOp' in" e
   | _ => fail "wp_op: not a 'wp'"
   end.
@@ -98,8 +104,8 @@ Tactic Notation "wp_proj" :=
   lazymatch goal with
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match eval hnf in e' with
-    | Fst _ => wp_bind K; etrans; [|eapply wp_fst; wp_done]; wp_finish
-    | Snd _ => wp_bind K; etrans; [|eapply wp_snd; wp_done]; wp_finish
+    | Fst _ => wp_bind_core K; etrans; [|eapply wp_fst; wp_done]; wp_finish
+    | Snd _ => wp_bind_core K; etrans; [|eapply wp_snd; wp_done]; wp_finish
     end) || fail "wp_proj: cannot find 'Fst' or 'Snd' in" e
   | _ => fail "wp_proj: not a 'wp'"
   end.
@@ -109,7 +115,7 @@ Tactic Notation "wp_if" :=
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match eval hnf in e' with
     | If _ _ _ =>
-      wp_bind K;
+      wp_bind_core K;
       etrans; [|eapply wp_if_true || eapply wp_if_false]; wp_finish
     end) || fail "wp_if: cannot find 'If' in" e
   | _ => fail "wp_if: not a 'wp'"
@@ -120,18 +126,18 @@ Tactic Notation "wp_match" :=
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match eval hnf in e' with
     | Case _ _ _ =>
-      wp_bind K;
+      wp_bind_core K;
       etrans; [|first[eapply wp_match_inl; wp_done|eapply wp_match_inr; wp_done]];
       simpl_subst; wp_finish
     end) || fail "wp_match: cannot find 'Match' in" e
   | _ => fail "wp_match: not a 'wp'"
   end.
 
-Tactic Notation "wp_focus" open_constr(efoc) :=
+Tactic Notation "wp_bind" open_constr(efoc) :=
   lazymatch goal with
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match e' with
-    | efoc => unify e' efoc; wp_bind K
-    end) || fail "wp_focus: cannot find" efoc "in" e
-  | _ => fail "wp_focus: not a 'wp'"
+    | efoc => unify e' efoc; wp_bind_core K
+    end) || fail "wp_bind: cannot find" efoc "in" e
+  | _ => fail "wp_bind: not a 'wp'"
   end.
diff --git a/prelude/base.v b/prelude/base.v
index b046cf8e814a59f69444b04f98611140c266b996..1d35448e8c2d753375a6e5891afe7f009a2b0178 100644
--- a/prelude/base.v
+++ b/prelude/base.v
@@ -8,7 +8,9 @@ Global Generalizable All Variables.
 Global Set Automatic Coercions Import.
 Global Set Asymmetric Patterns.
 Global Unset Transparent Obligations.
-From Coq Require Export Morphisms RelationClasses List Bool Utf8 Program Setoid.
+From Coq Require Export Morphisms RelationClasses List Bool Utf8 Setoid.
+Export ListNotations.
+From Coq.Program Require Export Basics Syntax.
 Obligation Tactic := idtac.
 
 (** Throughout this development we use [C_scope] for all general purpose
@@ -567,6 +569,8 @@ intersection [(∩)], and difference [(∖)], the singleton [{[_]}], the subset
 Class Empty A := empty: A.
 Notation "∅" := empty : C_scope.
 
+Instance empty_inhabited `(Empty A) : Inhabited A := populate ∅.
+
 Class Top A := top : A.
 Notation "⊤" := top : C_scope.
 
diff --git a/prelude/coPset.v b/prelude/coPset.v
index 9e37b07d05b7479608af74e3588599e5d1ecc0b7..85aec1c9cb567e4cecc8ea2c2d95bc762e69fe3c 100644
--- a/prelude/coPset.v
+++ b/prelude/coPset.v
@@ -315,9 +315,22 @@ Proof.
   apply coPset_finite_spec; destruct X as [[t ?]]; apply of_Pset_raw_finite.
 Qed.
 
-(** * Conversion from gsets of positives *)
+(** * Conversion to and from gsets of positives *)
+Lemma to_gset_wf (m : Pmap ()) : gmap_wf (K:=positive) m.
+Proof. done. Qed.
+Definition to_gset (X : coPset) : gset positive :=
+  let 'Mapset m := to_Pset X in
+  Mapset (GMap m (bool_decide_pack _ (to_gset_wf m))).
+
 Definition of_gset (X : gset positive) : coPset :=
   let 'Mapset (GMap (PMap t Ht) _) := X in of_Pset_raw t ↾ of_Pset_wf _ Ht.
+
+Lemma elem_of_to_gset X i : set_finite X → i ∈ to_gset X ↔ i ∈ X.
+Proof.
+  intros ?. rewrite <-elem_of_to_Pset by done.
+  unfold to_gset. by destruct (to_Pset X).
+Qed.
+
 Lemma elem_of_of_gset X i : i ∈ of_gset X ↔ i ∈ X.
 Proof. destruct X as [[[t ?]]]; apply elem_of_of_Pset_raw. Qed.
 Lemma of_gset_finite X : set_finite (of_gset X).
diff --git a/prelude/collections.v b/prelude/collections.v
index 7762ce91a3398701a4cb186635d3d664a4afa9f2..ef7383d529cf25477858fe854a4fb7bbf46bfc2f 100644
--- a/prelude/collections.v
+++ b/prelude/collections.v
@@ -390,7 +390,7 @@ Section simple_collection.
     split.
     - induction Xs; simpl; intros HXs; [by apply elem_of_empty in HXs|].
       setoid_rewrite elem_of_cons. apply elem_of_union in HXs. naive_solver.
-    - intros [X []]. induction 1; simpl; [by apply elem_of_union_l |].
+    - intros [X [Hx]]. induction Hx; simpl; [by apply elem_of_union_l |].
       intros. apply elem_of_union_r; auto.
   Qed.
 
@@ -948,3 +948,38 @@ Section more_finite.
     intros x ?; destruct (decide (x ∈ Y)); rewrite elem_of_app; set_solver.
   Qed.
 End more_finite.
+
+(** Sets of sequences of natural numbers *)
+(* The set [seq_seq start len] of natural numbers contains the sequence
+[start, start + 1, ..., start + (len-1)]. *)
+Fixpoint seq_set `{Singleton nat C, Union C, Empty C} (start len : nat) : C :=
+  match len with
+  | O => ∅
+  | S len' => {[ start ]} ∪ seq_set (S start) len'
+  end.
+
+Section seq_set.
+  Context `{SimpleCollection nat C}.
+  Implicit Types start len x : nat.
+
+  Lemma elem_of_seq_set start len x :
+    x ∈ seq_set start len ↔ start ≤ x < start + len.
+  Proof.
+    revert start. induction len as [|len IH]; intros start; simpl.
+    - rewrite elem_of_empty. omega.
+    - rewrite elem_of_union, elem_of_singleton, IH. omega.
+  Qed.
+
+  Lemma seq_set_S_disjoint start len : {[ start + len ]} ⊥ seq_set start len.
+  Proof. intros x. rewrite elem_of_singleton, elem_of_seq_set. omega. Qed.
+
+  Lemma seq_set_S_union start len :
+    seq_set start (C:=C) (S len) ≡ {[ start + len ]} ∪ seq_set start len.
+  Proof.
+    intros x. rewrite elem_of_union, elem_of_singleton, !elem_of_seq_set. omega.
+  Qed.
+
+  Lemma seq_set_S_union_L `{!LeibnizEquiv C} start len :
+    seq_set start (S len) = {[ start + len ]} ∪ seq_set start len.
+  Proof. unfold_leibniz. apply seq_set_S_union. Qed.
+End seq_set.
diff --git a/prelude/gmap.v b/prelude/gmap.v
index bf74baaafa5a49f93f2384d51016fa85080e2e59..e87f1434ea69bbd17ad1c2c75390bb0af1fe2447 100644
--- a/prelude/gmap.v
+++ b/prelude/gmap.v
@@ -33,6 +33,7 @@ Defined.
 Instance gmap_lookup `{Countable K} {A} : Lookup K A (gmap K A) := λ i m,
   let (m,_) := m in m !! encode i.
 Instance gmap_empty `{Countable K} {A} : Empty (gmap K A) := GMap ∅ I.
+Global Opaque gmap_empty.
 Lemma gmap_partial_alter_wf `{Countable K} {A} (f : option A → option A) m i :
   gmap_wf m → gmap_wf (partial_alter f (encode i) m).
 Proof.
diff --git a/prelude/hlist.v b/prelude/hlist.v
index 8c888c732290a9b90917fe6d68b26ffc56ded8c1..386cc42066800d3ca7986b4521f49102dc2db8a0 100644
--- a/prelude/hlist.v
+++ b/prelude/hlist.v
@@ -1,4 +1,5 @@
 From iris.prelude Require Import tactics.
+Local Set Universe Polymorphism.
 
 (* Not using [list Type] in order to avoid universe inconsistencies *)
 Inductive tlist := tnil : tlist | tcons : Type → tlist → tlist.
diff --git a/prelude/list.v b/prelude/list.v
index 59b26331813bf82dc7cd243e88f8721e7b9fd82e..e349d2fd441acdc86234de1b6ca3e7999c68aa4c 100644
--- a/prelude/list.v
+++ b/prelude/list.v
@@ -196,6 +196,8 @@ Definition imap_go {A B} (f : nat → A → B) : nat → list A → list B :=
   fix go (n : nat) (l : list A) :=
   match l with [] => [] | x :: l => f n x :: go (S n) l end.
 Definition imap {A B} (f : nat → A → B) : list A → list B := imap_go f 0.
+Arguments imap : simpl never.
+
 Definition zipped_map {A B} (f : list A → list A → A → B) :
   list A → list A → list B := fix go l k :=
   match k with [] => [] | x :: k => f l k x :: go (x :: l) k end.
@@ -926,7 +928,7 @@ Proof. by destruct n. Qed.
 Lemma drop_length l n : length (drop n l) = length l - n.
 Proof. revert n. by induction l; intros [|i]; f_equal/=. Qed.
 Lemma drop_ge l n : length l ≤ n → drop n l = [].
-Proof. revert n. induction l; intros [|??]; simpl in *; auto with lia. Qed.
+Proof. revert n. induction l; intros [|?]; simpl in *; auto with lia. Qed.
 Lemma drop_all l : drop (length l) l = [].
 Proof. by apply drop_ge. Qed.
 Lemma drop_drop l n1 n2 : drop n1 (drop n2 l) = drop (n2 + n1) l.
@@ -1266,20 +1268,31 @@ Proof.
 Qed.
 
 (** ** Properties of the [imap] function *)
-Lemma imap_cons {B} (f : nat → A → B) x l :
-  imap f (x :: l) = f 0 x :: imap (f ∘ S) l.
+Lemma imap_nil {B} (f : nat → A → B) : imap f [] = [].
+Proof. done. Qed.
+Lemma imap_app {B} (f : nat → A → B) l1 l2 :
+  imap f (l1 ++ l2) = imap f l1 ++ imap (λ n, f (length l1 + n)) l2.
 Proof.
-  unfold imap. simpl. f_equal. generalize 0.
-  induction l; intros n; simpl; repeat (auto||f_equal).
+  unfold imap. generalize 0. revert l2.
+  induction l1 as [|x l1 IH]; intros l2 n; f_equal/=; auto.
+  rewrite IH. f_equal. clear. revert n.
+  induction l2; simpl; auto with f_equal lia.
 Qed.
+Lemma imap_cons {B} (f : nat → A → B) x l :
+  imap f (x :: l) = f 0 x :: imap (f ∘ S) l.
+Proof. apply (imap_app _ [_]). Qed.
+
 Lemma imap_ext {B} (f g : nat → A → B) l :
-  (∀ i x, f i x = g i x) →
-  imap f l = imap g l.
+  (∀ i x, l !! i = Some x → f i x = g i x) → imap f l = imap g l.
 Proof.
-  unfold imap. intro EQ. generalize 0.
-  induction l; simpl; intros n; f_equal; auto.
+  revert f g; induction l as [|x l IH]; intros f g Hfg; auto.
+  rewrite !imap_cons; f_equal; eauto.
 Qed.
 
+Lemma imap_fmap {B C} (f : nat → B → C) (g : A → B) l :
+  imap f (g <$> l) = imap (λ n, f n ∘ g) l.
+Proof. unfold imap. generalize 0. induction l; csimpl; auto with f_equal. Qed.
+
 (** ** Properties of the [mask] function *)
 Lemma mask_nil f βs : mask f βs (@nil A) = [].
 Proof. by destruct βs. Qed.
@@ -2828,7 +2841,7 @@ Section fmap.
     (∀ x, f x = y) → f <$> l = replicate (length l) y.
   Proof. intros; induction l; f_equal/=; auto. Qed.
   Lemma list_lookup_fmap l i : (f <$> l) !! i = f <$> (l !! i).
-  Proof. revert i. induction l; by intros [|]. Qed.
+  Proof. revert i. induction l; intros [|n]; by try revert n. Qed.
   Lemma list_lookup_fmap_inv l i x :
     (f <$> l) !! i = Some x → ∃ y, x = f y ∧ l !! i = Some y.
   Proof.
diff --git a/prelude/numbers.v b/prelude/numbers.v
index 0dad1337861dfe92739d36d21f2e48c4549a8aad..6fc3ed36e8abd2e03e2a5eca67c5093adf31a545 100644
--- a/prelude/numbers.v
+++ b/prelude/numbers.v
@@ -3,7 +3,7 @@
 (** This file collects some trivial facts on the Coq types [nat] and [N] for
 natural numbers, and the type [Z] for integers. It also declares some useful
 notations. *)
-From Coq Require Export Eqdep PArith NArith ZArith NPeano.
+From Coq Require Export EqdepFacts PArith NArith ZArith NPeano.
 From Coq Require Import QArith Qcanon.
 From iris.prelude Require Export base decidable option.
 Open Scope nat_scope.
@@ -82,7 +82,7 @@ Proof. intros. destruct (Nat_mul_split_l n x2 x1 y2 y1); auto with lia. Qed.
 Notation lcm := Nat.lcm.
 Notation divide := Nat.divide.
 Notation "( x | y )" := (divide x y) : nat_scope.
-Instance divide_dec x y : Decision (x | y).
+Instance Nat_divide_dec x y : Decision (x | y).
 Proof.
   refine (cast_if (decide (lcm x y = y))); by rewrite Nat.divide_lcm_iff.
 Defined.
@@ -94,6 +94,11 @@ Hint Extern 0 (_ | _) => reflexivity.
 Lemma Nat_divide_ne_0 x y : (x | y) → y ≠ 0 → x ≠ 0.
 Proof. intros Hxy Hy ->. by apply Hy, Nat.divide_0_l. Qed.
 
+Lemma Nat_iter_S {A} n (f: A → A) x : Nat.iter (S n) f x = f (Nat.iter n f x).
+Proof. done. Qed.
+Lemma Nat_iter_S_r {A} n (f: A → A) x : Nat.iter (S n) f x = Nat.iter n f (f x).
+Proof. induction n; f_equal/=; auto. Qed.
+
 (** * Notations and properties of [positive] *)
 Open Scope positive_scope.
 
@@ -226,16 +231,19 @@ Infix "`rem`" := Z.rem (at level 35) : Z_scope.
 Infix "≪" := Z.shiftl (at level 35) : Z_scope.
 Infix "≫" := Z.shiftr (at level 35) : Z_scope.
 
-Instance: Inj (=) (=) Zpos.
+Instance Zpos_inj : Inj (=) (=) Zpos.
 Proof. by injection 1. Qed.
-Instance: Inj (=) (=) Zneg.
+Instance Zneg_inj : Inj (=) (=) Zneg.
 Proof. by injection 1. Qed.
 
+Instance Z_of_nat_inj : Inj (=) (=) Z.of_nat.
+Proof. intros n1 n2. apply Nat2Z.inj. Qed.
+
 Instance Z_eq_dec: ∀ x y : Z, Decision (x = y) := Z.eq_dec.
 Instance Z_le_dec: ∀ x y : Z, Decision (x ≤ y) := Z_le_dec.
 Instance Z_lt_dec: ∀ x y : Z, Decision (x < y) := Z_lt_dec.
 Instance Z_inhabited: Inhabited Z := populate 1.
-Instance: PartialOrder (≤).
+Instance Z_le_order : PartialOrder (≤).
 Proof.
   repeat split; red. apply Z.le_refl. apply Z.le_trans. apply Z.le_antisymm.
 Qed.
diff --git a/prelude/pretty.v b/prelude/pretty.v
index 7637aedf0ec050d3b7d4ce6fbbcc07cfafd96625..1514dbc829a77081313670ed35028d5fa64596f8 100644
--- a/prelude/pretty.v
+++ b/prelude/pretty.v
@@ -32,10 +32,13 @@ Lemma pretty_N_go_step x s :
   = pretty_N_go (x `div` 10) (String (pretty_N_char (x `mod` 10)) s).
 Proof.
   unfold pretty_N_go; intros; destruct (wf_guard 32 N.lt_wf_0 x).
-  unfold pretty_N_go_help; fold pretty_N_go_help.
+  destruct wf_guard. (* this makes coqchk happy. *)
+  unfold pretty_N_go_help at 1; fold pretty_N_go_help.
   by destruct (decide (0 < x)%N); auto using pretty_N_go_help_irrel.
 Qed.
 Instance pretty_N : Pretty N := λ x, pretty_N_go x ""%string.
+Lemma pretty_N_unfold x : pretty x = pretty_N_go x "".
+Proof. done. Qed.
 Instance pretty_N_inj : Inj (@eq N) (=) pretty.
 Proof.
   assert (∀ x y, x < 10 → y < 10 →
@@ -43,7 +46,8 @@ Proof.
   { compute; intros. by repeat (discriminate || case_match). }
   cut (∀ x y s s', pretty_N_go x s = pretty_N_go y s' →
     String.length s = String.length s' → x = y ∧ s = s').
-  { intros help x y ?. eapply help; eauto. }
+  { intros help x y Hp.
+    eapply (help x y "" ""); [by rewrite <-!pretty_N_unfold|done]. }
   assert (∀ x s, ¬String.length (pretty_N_go x s) < String.length s) as help.
   { setoid_rewrite <-Nat.le_ngt.
     intros x; induction (N.lt_wf_0 x) as [x _ IH]; intros s.
diff --git a/prelude/proof_irrel.v b/prelude/proof_irrel.v
index 4afb6c5d13258b604db923c01217d47097853dbe..892a9213c3d6ce39de3c94fad263708e53436d8e 100644
--- a/prelude/proof_irrel.v
+++ b/prelude/proof_irrel.v
@@ -1,14 +1,13 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects facts on proof irrelevant types/propositions. *)
-From Coq Require Import Eqdep_dec.
 From iris.prelude Require Export base.
 
 Hint Extern 200 (ProofIrrel _) => progress (lazy beta) : typeclass_instances.
 
-Instance: ProofIrrel True.
+Instance True_pi: ProofIrrel True.
 Proof. intros [] []; reflexivity. Qed.
-Instance: ProofIrrel False.
+Instance False_pi: ProofIrrel False.
 Proof. intros []. Qed.
 Instance and_pi (A B : Prop) :
   ProofIrrel A → ProofIrrel B → ProofIrrel (A ∧ B).
@@ -16,11 +15,18 @@ Proof. intros ?? [??] [??]. f_equal; trivial. Qed.
 Instance prod_pi (A B : Type) :
   ProofIrrel A → ProofIrrel B → ProofIrrel (A * B).
 Proof. intros ?? [??] [??]. f_equal; trivial. Qed.
-Instance eq_pi {A} `{∀ x y : A, Decision (x = y)} (x y : A) :
+Instance eq_pi {A} (x : A) `{∀ z, Decision (x = z)} (y : A) :
   ProofIrrel (x = y).
 Proof.
-  intros ??. apply eq_proofs_unicity.
-  intros x' y'. destruct (decide (x' = y')); tauto.
+  set (f z (H : x = z) :=
+    match decide (x = z) return x = z with
+    | left H => H | right H' => False_rect _ (H' H)
+    end).
+  assert (∀ z (H : x = z),
+    eq_trans (eq_sym (f x (eq_refl x))) (f z H) = H) as help.
+  { intros ? []. destruct (f x eq_refl); tauto. }
+  intros p q. rewrite <-(help _ p), <-(help _ q).
+  unfold f at 2 4. destruct (decide _). reflexivity. exfalso; tauto.
 Qed.
 Instance Is_true_pi (b : bool) : ProofIrrel (Is_true b).
 Proof. destruct b; simpl; apply _. Qed.
diff --git a/prelude/tactics.v b/prelude/tactics.v
index 13e74464c22841fe004c0f096cc4276338d7ca91..dd46cd809fa5c9a1c4b616179d0a8015cd1fc005 100644
--- a/prelude/tactics.v
+++ b/prelude/tactics.v
@@ -3,7 +3,7 @@
 (** This file collects general purpose tactics that are used throughout
 the development. *)
 From Coq Require Import Omega.
-From Coq Require Export Psatz.
+From Coq Require Export Lia.
 From iris.prelude Require Export decidable.
 
 Lemma f_equal_dep {A B} (f g : ∀ x : A, B x) x : f = g → f x = g x.
@@ -407,6 +407,13 @@ Tactic Notation "feed" "destruct" constr(H) :=
 Tactic Notation "feed" "destruct" constr(H) "as" simple_intropattern(IP) :=
   feed (fun p => let H':=fresh in pose proof p as H'; destruct H' as IP) H.
 
+(** The block definitions are taken from [Coq.Program.Equality] and can be used
+by tactics to separate their goal from hypotheses they generalize over. *)
+Definition block {A : Type} (a : A) := a.
+
+Ltac block_goal := match goal with [ |- ?T ] => change (block T) end.
+Ltac unblock_goal := unfold block in *.
+
 
 (** The following tactic can be used to add support for patterns to tactic notation:
 It will search for the first subterm of the goal matching [pat], and then call [tac]
diff --git a/prelude/vector.v b/prelude/vector.v
index 26f78e9451e419fb4f04d05a9365e2f3addbbe12..9253a797f7b162ac3f5681259f1284d682564851 100644
--- a/prelude/vector.v
+++ b/prelude/vector.v
@@ -69,12 +69,13 @@ Ltac inv_fin i :=
     revert dependent i; match goal with |- ∀ i, @?P i => apply (fin_S_inv P) end
   end.
 
-Instance: Inj (=) (=) (@FS n).
+Instance FS_inj: Inj (=) (=) (@FS n).
 Proof. intros n i j. apply Fin.FS_inj. Qed.
-Instance: Inj (=) (=) (@fin_to_nat n).
+Instance fin_to_nat_inj : Inj (=) (=) (@fin_to_nat n).
 Proof.
   intros n i. induction i; intros j; inv_fin j; intros; f_equal/=; auto with lia.
 Qed.
+
 Lemma fin_to_nat_lt {n} (i : fin n) : fin_to_nat i < n.
 Proof. induction i; simpl; lia. Qed.
 Lemma fin_to_of_nat n m (H : n < m) : fin_to_nat (Fin.of_nat_lt H) = n.
@@ -82,6 +83,31 @@ Proof.
   revert m H. induction n; intros [|?]; simpl; auto; intros; exfalso; lia.
 Qed.
 
+Fixpoint fin_plus_inv {n1 n2} : ∀ (P : fin (n1 + n2) → Type)
+    (H1 : ∀ i1 : fin n1, P (Fin.L n2 i1))
+    (H2 : ∀ i2, P (Fin.R n1 i2)) (i : fin (n1 + n2)), P i :=
+  match n1 with
+  | 0 => λ P H1 H2 i, H2 i
+  | S n => λ P H1 H2, fin_S_inv P (H1 0%fin) (fin_plus_inv _ (λ i, H1 (FS i)) H2)
+  end.
+
+Lemma fin_plus_inv_L {n1 n2} (P : fin (n1 + n2) → Type)
+    (H1: ∀ i1 : fin n1, P (Fin.L _ i1)) (H2: ∀ i2, P (Fin.R _ i2)) (i: fin n1) :
+  fin_plus_inv P H1 H2 (Fin.L n2 i) = H1 i.
+Proof.
+  revert P H1 H2 i.
+  induction n1 as [|n1 IH]; intros P H1 H2 i; inv_fin i; simpl; auto.
+  intros i. apply (IH (λ i, P (FS i))).
+Qed.
+
+Lemma fin_plus_inv_R {n1 n2} (P : fin (n1 + n2) → Type)
+    (H1: ∀ i1 : fin n1, P (Fin.L _ i1)) (H2: ∀ i2, P (Fin.R _ i2)) (i: fin n2) :
+  fin_plus_inv P H1 H2 (Fin.R n1 i) = H2 i.
+Proof.
+  revert P H1 H2 i; induction n1 as [|n1 IH]; intros P H1 H2 i; simpl; auto.
+  apply (IH (λ i, P (FS i))).
+Qed.
+
 (** * Vectors *)
 (** The type [vec n] represents lists of consisting of exactly [n] elements.
 Whereas the standard library declares exactly the same notations for vectors as
diff --git a/program_logic/adequacy.v b/program_logic/adequacy.v
index 2acb580466368a5c18e85b2afb83b2531a58eb96..b2480ac7ed9d2e702f47e0b82cc2640da20b790b 100644
--- a/program_logic/adequacy.v
+++ b/program_logic/adequacy.v
@@ -1,155 +1,170 @@
-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.
+
+Record adequate {Λ} (e1 : expr Λ) (σ1 : state Λ) (φ : val Λ → Prop) := {
+  adequate_result t2 σ2 v2 :
+   rtc step ([e1], σ1) (of_val v2 :: t2, σ2) → φ v2;
+  adequate_safe t2 σ2 e2 :
+   rtc step ([e1], σ1) (t2, σ2) →
+   e2 ∈ t2 → (is_Some (to_val e2) ∨ reducible e2 σ2)
+}.
+
+Theorem adequate_tp_safe {Λ} (e1 : expr Λ) t2 σ1 σ2 φ :
+  adequate e1 σ1 φ →
+  rtc step ([e1], σ1) (t2, σ2) →
+  Forall (λ e, is_Some (to_val e)) t2 ∨ ∃ t3 σ3, step (t2, σ2) (t3, σ3).
+Proof.
+  intros Had ?.
+  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 (adequate_safe e1 σ1 φ Had t2 σ2 e2) as [?|(e3&σ3&efs&?)];
+    rewrite ?eq_None_not_Some; auto.
+  { exfalso. eauto. }
+  destruct (elem_of_list_split t2 e2) as (t2'&t2''&->); auto.
+  right; exists (t2' ++ e3 :: t2'' ++ efs), σ3; econstructor; eauto.
+Qed.
 
 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 Λ Σ.
-
-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).
+Implicit Types P Q : iProp Σ.
+Implicit Types Φ : val Λ → iProp Σ.
+Implicit Types Φs : list (val Λ → iProp Σ).
+
+Notation world σ := (wsat ★ ownE ⊤ ★ ownP_auth σ)%I.
+
+Notation wptp t := ([★ list] ef ∈ t, WP ef {{ _, True }})%I.
+
+Lemma wp_step e1 σ1 e2 σ2 efs Φ :
+  prim_step e1 σ1 e2 σ2 efs →
+  world σ1 ★ WP e1 {{ Φ }} =r=> ▷ |=r=> ◇ (world σ2 ★ WP e2 {{ Φ }} ★ wptp efs).
 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.
+  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 ">(Hw & HE & _ & H)"; first by iFrame.
+  iVsIntro; iNext.
+  iVs ("H" $! e2 σ2 efs with "[%] [Hw HE]")
+    as ">($ & $ & $ & $)"; try iFrame; eauto.
 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).
+
+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.
-  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.
+  iIntros (Hstep) "(HW & He & Ht)".
+  destruct Hstep as [e1' σ1' e2' σ2' efs [|? t1'] t2' ?? Hstep]; simplify_eq/=.
+  - iExists e2', (t2' ++ efs); iSplitR; first eauto.
+    rewrite big_sepL_app. iFrame "Ht". iApply wp_step; try iFrame; eauto.
+  - iExists e, (t1' ++ e2' :: t2' ++ efs); iSplitR; first eauto.
+    rewrite !big_sepL_app !big_sepL_cons big_sepL_app.
+    iDestruct "Ht" as "($ & He' & $)"; iFrame "He".
+    iApply wp_step; try iFrame; eauto.
 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 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 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.
+  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 ">?". by iApply IH.
 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.
+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 rvs_iter_frame_l n R Q :
+  R ★ Nat.iter n (λ P, |=r=> ▷ P) Q ⊢ Nat.iter n (λ P, |=r=> ▷ P) (R ★ Q).
 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.
+  induction n as [|n IH]; simpl; [done|].
+  by rewrite rvs_frame_l {1}(later_intro R) -later_sep IH.
 Qed.
 
-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.
+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. eapply wp_adequacy_result with (E:=E); first done.
-  move:Hht. by rewrite /ht uPred.always_elim=>/uPred.impl_entails.
+  intros. rewrite wptp_steps //.
+  rewrite (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 ">(_ & _ & $)"; iFrame; auto.
 Qed.
 
-Lemma wp_adequacy_reducible E Φ e1 e2 t2 σ1 m σ2 :
-  ✓ m →
-  (ownP σ1 ★ ownG m ⊢ WP e1 @ E {{ Φ }}) →
-  rtc step ([e1], σ1) (t2, σ2) →
-  e2 ∈ t2 → (is_Some (to_val e2) ∨ reducible e2 σ2).
+Lemma wp_safe e σ Φ :
+  world σ ★ WP e {{ Φ }} =r=> ▷ ■ (is_Some (to_val e) ∨ reducible e σ).
 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.
+  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 ">(?&?&%&?)"; first by iFrame.
+  eauto 10.
 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 {{ Φ }}) →
-  rtc step ([e1], σ1) (t2, σ2) →
-  Forall (λ e, is_Some (to_val e)) t2 ∨ ∃ t3 σ3, step (t2, σ2) (t3, σ3).
+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.
-  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&?)];
-    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.
+  intros ? He2. rewrite wptp_steps //; rewrite (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". by iApply (big_sepL_elem_of with "Htp").
 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).
+Lemma wptp_invariance n e1 e2 t1 t2 σ1 σ2 I φ Φ :
+  nsteps step n (e1 :: t1, σ1) (t2, σ2) →
+  (I ={⊤,∅}=> ∃ σ', ownP σ' ∧ ■ φ σ') →
+  I ★ world σ1 ★ WP e1 {{ Φ }} ★ wptp t1 ⊢
+  Nat.iter (S (S n)) (λ P, |=r=> ▷ P) (■ φ σ2).
 Proof.
-  intros ? Hht. eapply wp_adequacy_safe with (E:=E) (Φ:=Φ); first done.
-  move:Hht. by rewrite /ht uPred.always_elim=>/uPred.impl_entails.
+  intros ? HI. rewrite wptp_steps //.
+  rewrite (Nat_iter_S_r (S n)) rvs_iter_frame_l. apply rvs_iter_mono.
+  iIntros "[HI H]".
+  iDestruct "H" as (e2' t2') "(% & (Hw&HE&Hσ) & _)"; subst.
+  rewrite pvs_eq in HI;
+    iVs (HI with "HI [Hw HE]") as "> (_ & _ & H)"; first by iFrame.
+  iDestruct "H" as (σ2') "[Hσf %]".
+  iDestruct (ownP_agree σ2 σ2' with "[#]") as %<-. by iFrame. eauto.
 Qed.
 End adequacy.
+
+Theorem wp_adequacy Σ `{irisPreG Λ Σ} e σ φ :
+  (∀ `{irisG Λ Σ}, ownP σ ⊢ WP e {{ v, ■ φ v }}) →
+  adequate e σ φ.
+Proof.
+  intros Hwp; split.
+  - intros t2 σ2 v2 [n ?]%rtc_nsteps.
+    eapply (adequacy (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
+    rewrite Nat_iter_S. iVs (iris_alloc σ) as (?) "(?&?&?&Hσ)".
+    iVsIntro. iNext. iApply wptp_result; eauto.
+    iFrame. iSplitL. by iApply Hwp. by iApply big_sepL_nil.
+  - intros t2 σ2 e2 [n ?]%rtc_nsteps ?.
+    eapply (adequacy (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
+    rewrite Nat_iter_S. iVs (iris_alloc σ) as (?) "(Hw & HE & Hσ & Hσf)".
+    iVsIntro. iNext. iApply wptp_safe; eauto.
+    iFrame "Hw HE Hσ". iSplitL. by iApply Hwp. by iApply big_sepL_nil.
+Qed.
+
+Theorem wp_invariance Σ `{irisPreG Λ Σ} e σ1 t2 σ2 I φ Φ :
+  (∀ `{irisG Λ Σ}, ownP σ1 ={⊤}=> I ★ WP e {{ Φ }}) →
+  (∀ `{irisG Λ Σ}, I ={⊤,∅}=> ∃ σ', ownP σ' ∧ ■ φ σ') →
+  rtc step ([e], σ1) (t2, σ2) →
+  φ σ2.
+Proof.
+  intros Hwp HI [n ?]%rtc_nsteps.
+  eapply (adequacy (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
+  rewrite Nat_iter_S. iVs (iris_alloc σ1) as (?) "(Hw & HE & ? & Hσ)".
+  rewrite pvs_eq in Hwp.
+  iVs (Hwp _ with "Hσ [Hw HE]") as ">(? & ? & ? & ?)"; first by iFrame.
+  iVsIntro. iNext. iApply wptp_invariance; eauto. iFrame. by iApply big_sepL_nil.
+Qed.
diff --git a/program_logic/auth.v b/program_logic/auth.v
index 2c169450819a1259c20c058da946a28da790c579..a56f54d5b2e6533e2b15122dc6cb32db4e925339 100644
--- a/program_logic/auth.v
+++ b/program_logic/auth.v
@@ -1,27 +1,26 @@
-From iris.algebra Require Export auth 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 pviewshifts.
+From iris.algebra Require Export auth.
+From iris.algebra Require Import gmap.
+From iris.proofmode Require Import invariants.
 Import uPred.
 
 (* The CMRA we need. *)
-Class authG Λ Σ (A : ucmraT) := AuthG {
-  auth_inG :> inG Λ Σ (authR A);
-  auth_timeless :> CMRADiscrete A;
+Class authG Σ (A : ucmraT) := AuthG {
+  auth_inG :> inG Σ (authR A);
+  auth_discrete :> 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.
-Proof. split; try apply _. apply: inGF_inG. Qed.
+Definition authΣ (A : ucmraT) : gFunctors := #[ GFunctor (constRF (authR A)) ].
+
+Instance subG_authΣ Σ A : subG (authΣ A) Σ → CMRADiscrete A → authG Σ A.
+Proof. intros ?%subG_inG ?. by split. 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,21 +40,25 @@ 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.
 
   Lemma auth_own_op γ a b : auth_own γ (a ⋅ b) ⊣⊢ auth_own γ a ★ auth_own γ b.
   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.
+  Proof. by rewrite /FromSep auth_own_op. Qed.
+
   Lemma auth_own_mono γ a b : a ≼ b → auth_own γ b ⊢ auth_own γ a.
   Proof. intros [? ->]. by rewrite auth_own_op sep_elim_l. Qed.
 
@@ -67,51 +70,40 @@ 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.
-
-  Context {V} (fsa : FSA Λ (globalF Σ) V) `{!FrameShiftAssertion fsaV fsa}.
+  Lemma auth_empty γ : True =r=> auth_own γ ∅.
+  Proof. by rewrite /auth_own -own_empty. Qed.
 
-  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φ]".
-    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Ψ".
+    iIntros (?) "(#? & >Hγf)". rewrite /auth_ctx /auth_own.
+    iInv N as (a') "[>Hγ Hφ]" "Hclose". iCombine "Hγ" "Hγf" as "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 e81180b7981883bb0940fa9cacbb226e59f36fc8..97a46c231fdb095347c590bfbef6efd07a22353e 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.program_logic Require Import auth saved_prop.
-From iris.proofmode Require Import tactics invariants ghost_ownership.
+From iris.program_logic Require Export pviewshifts.
+From iris.algebra Require Import auth gmap agree upred_big_op.
+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,13 +61,13 @@ 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.
-  iApply (@own_update with "Hγ"); apply prod_update; simpl; last reflexivity.
+  iApply (own_update with "Hγ"); apply prod_update; simpl; last reflexivity.
   by apply auth_update_no_frame, option_local_update, exclusive_local_update.
 Qed.
 
@@ -78,7 +76,7 @@ Lemma box_own_agree γ Q1 Q2 :
 Proof.
   rewrite /box_own_prop -own_op own_valid prod_validI /= and_elim_r.
   rewrite option_validI /= agree_validI agree_equivI later_equivI /=.
-  iIntros "#HQ >". rewrite -{2}(iProp_fold_unfold Q1).
+  iIntros "#HQ !>". rewrite -{2}(iProp_fold_unfold Q1).
   iRewrite "HQ". by rewrite iProp_fold_unfold.
 Qed.
 
@@ -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".
   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γ']")
+    as "[[>Hγ' #[HγΦ Hinv']] ?]"; first done.
+  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".
   iDestruct (big_sepM_later _ f with "Hf") as "Hf".
   iDestruct (big_sepM_delete _ f with "Hf")
-    as "[[Hγ' #[HγΦ Hinv']] ?]"; first done; iTimeless "Hγ'".
+    as "[[>Hγ' #[HγΦ Hinv']] ?]"; first done.
   iDestruct (box_own_auth_agree γ b true with "[#]")
-    as "%"; subst; first by iFrame.
+    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".
+  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 :
@@ -187,17 +184,18 @@ Lemma box_empty_all f P Q :
 Proof.
   iDestruct 1 as (Φ) "[#HeqP Hf]".
   iAssert ([★ map] γ↦b ∈ f, ▷ Φ γ ★ box_own_auth γ (◯ Excl' false) ★
-    box_own_prop γ (Φ γ) ★ inv N (slice_inv γ (Φ γ)))%I with "|==>[Hf]" as "[HΦ ?]".
+    box_own_prop γ (Φ γ) ★ inv N (slice_inv γ (Φ γ)))%I with "==>[Hf]" as "[HΦ ?]".
   { 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".
     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/cancelable_invariants.v b/program_logic/cancelable_invariants.v
new file mode 100644
index 0000000000000000000000000000000000000000..cef203f71498edb2b69073bcf85957123837da4a
--- /dev/null
+++ b/program_logic/cancelable_invariants.v
@@ -0,0 +1,71 @@
+From iris.program_logic Require Export invariants.
+From iris.algebra Require Export frac.
+From iris.proofmode Require Import invariants tactics.
+Import uPred.
+
+Class cinvG Σ := cinv_inG :> inG Σ fracR.
+
+Section defs.
+  Context `{irisG Λ Σ, cinvG Σ}.
+
+  Definition cinv_own (γ : gname) (p : frac) : iProp Σ := own γ p.
+
+  Definition cinv (N : namespace) (γ : gname) (P : iProp Σ) : iProp Σ :=
+    inv N (P ∨ cinv_own γ 1%Qp)%I.
+End defs.
+
+Instance: Params (@cinv) 6.
+Typeclasses Opaque cinv_own cinv.
+
+Section proofs.
+  Context `{irisG Λ Σ, cinvG Σ}.
+
+  Global Instance cinv_own_timeless γ p : TimelessP (cinv_own γ p).
+  Proof. rewrite /cinv_own; apply _. Qed.
+
+  Global Instance cinv_ne N γ n : Proper (dist n ==> dist n) (cinv N γ).
+  Proof. solve_proper. Qed.
+  Global Instance cinv_proper N γ : Proper ((≡) ==> (≡)) (cinv N γ).
+  Proof. apply (ne_proper _). Qed.
+
+  Global Instance cinv_persistent N γ P : PersistentP (cinv N γ P).
+  Proof. rewrite /cinv; apply _. Qed.
+
+  Lemma cinv_own_op γ q1 q2 :
+    cinv_own γ q1 ★ cinv_own γ q2 ⊣⊢ cinv_own γ (q1 + q2).
+  Proof. by rewrite /cinv_own own_op. Qed.
+
+  Lemma cinv_own_half γ q : cinv_own γ (q/2) ★ cinv_own γ (q/2) ⊣⊢ cinv_own γ q.
+  Proof. by rewrite cinv_own_op Qp_div_2. Qed.
+
+  Lemma cinv_own_valid γ q1 q2 : cinv_own γ q1 ★ cinv_own γ q2 ⊢ ✓ (q1 + q2)%Qp.
+  Proof. rewrite /cinv_own -own_op own_valid. by iIntros "% !%". Qed.
+
+  Lemma cinv_own_1_l γ q : cinv_own γ 1 ★ cinv_own γ q ⊢ False.
+  Proof. rewrite cinv_own_valid. by iIntros (?%(exclusive_l 1%Qp)). Qed.
+
+  Lemma cinv_alloc E N P : ▷ P ={E}=> ∃ γ, cinv N γ P ★ cinv_own γ 1.
+  Proof.
+    rewrite /cinv /cinv_own. iIntros "HP".
+    iVs (own_alloc 1%Qp) as (γ) "H1"; first done.
+    iVs (inv_alloc N _ (P ∨ own γ 1%Qp)%I with "[HP]"); eauto.
+  Qed.
+
+  Lemma cinv_cancel E N γ P :
+    nclose N ⊆ E → cinv N γ P ⊢ cinv_own γ 1 ={E}=★ ▷ P.
+  Proof.
+    rewrite /cinv. iIntros (?) "#Hinv Hγ".
+    iInv N as "[$|>Hγ']" "Hclose"; first iApply "Hclose"; eauto.
+    iDestruct (cinv_own_1_l with "[Hγ Hγ']") as %[]. by iFrame.
+  Qed.
+
+  Lemma cinv_open E N γ p P :
+    nclose N ⊆ E →
+    cinv N γ P ⊢ cinv_own γ p ={E,E∖N}=★ ▷ P ★ cinv_own γ p ★ (▷ P ={E∖N,E}=★ True).
+  Proof.
+    rewrite /cinv. iIntros (?) "#Hinv Hγ".
+    iInv N as "[$|>Hγ']" "Hclose".
+    - iIntros "!==> {$Hγ} HP". iApply "Hclose"; eauto.
+    - iDestruct (cinv_own_1_l with "[Hγ Hγ']") as %[]. by iFrame.
+  Qed.
+End proofs.
diff --git a/program_logic/counter_examples.v b/program_logic/counter_examples.v
new file mode 100644
index 0000000000000000000000000000000000000000..d51505fa8da76e1d533528519751a3248e0c28c8
--- /dev/null
+++ b/program_logic/counter_examples.v
@@ -0,0 +1,194 @@
+From iris.algebra Require Import upred.
+From iris.proofmode Require Import tactics.
+
+(** This proves that we need the â–· in a "Saved Proposition" construction with
+name-dependent allocation. *)
+Module savedprop. Section savedprop.
+  Context (M : ucmraT).
+  Notation iProp := (uPred M).
+  Notation "¬ P" := (□ (P → False))%I : uPred_scope.
+  Implicit Types P : iProp.
+
+  (* Saved Propositions and view shifts. *)
+  Context (sprop : Type) (saved : sprop → iProp → iProp).
+  Hypothesis sprop_persistent : ∀ i P, PersistentP (saved i P).
+  Hypothesis sprop_alloc_dep :
+    ∀ (P : sprop → iProp), True =r=> (∃ i, saved i (P i)).
+  Hypothesis sprop_agree : ∀ i P Q, saved i P ∧ saved i Q ⊢ □ (P ↔ Q).
+
+  (** A bad recursive reference: "Assertion with name [i] does not hold" *)
+  Definition A (i : sprop) : iProp := ∃ P, ¬ P ★ saved i P.
+
+  Lemma A_alloc : True =r=> ∃ i, saved i (A i).
+  Proof. by apply sprop_alloc_dep. Qed.
+
+  Lemma saved_NA i : saved i (A i) ⊢ ¬ A i.
+  Proof.
+    iIntros "#Hs !# #HA". iPoseProof "HA" as "HA'".
+    iDestruct "HA'" as (P) "[#HNP HsP]". iApply "HNP".
+    iDestruct (sprop_agree i P (A i) with "[]") as "#[_ HP]".
+    { eauto. }
+    iApply "HP". done.
+  Qed.
+
+  Lemma saved_A i : saved i (A i) ⊢ A i.
+  Proof.
+    iIntros "#Hs". iExists (A i). iFrame "#".
+    by iApply saved_NA.
+  Qed.
+
+  Lemma contradiction : False.
+  Proof.
+    apply (@uPred.adequacy M False 1); simpl.
+    iIntros "". iVs A_alloc as (i) "#H".
+    iPoseProof (saved_NA with "H") as "HN".
+    iVsIntro. iNext.
+    iApply "HN". iApply saved_A. done.
+  Qed.
+
+End savedprop. End savedprop.
+
+(** This proves that we need the â–· when opening invariants. *)
+(** We fork in [uPred M] for any M, but the proof would work in any BI. *)
+Module inv. Section inv.
+  Context (M : ucmraT).
+  Notation iProp := (uPred M).
+  Implicit Types P : iProp.
+
+  (** Assumptions *)
+  (* We have view shifts (two classes: empty/full mask) *)
+  Inductive mask := M0 | M1.
+  Context (pvs : mask → iProp → iProp).
+
+  Hypothesis pvs_intro : ∀ E P, P ⊢ pvs E P.
+  Hypothesis pvs_mono : ∀ E P Q, (P ⊢ Q) → pvs E P ⊢ pvs E Q.
+  Hypothesis pvs_pvs : ∀ E P, pvs E (pvs E P) ⊢ pvs E P.
+  Hypothesis pvs_frame_l : ∀ E P Q, P ★ pvs E Q ⊢ pvs E (P ★ Q).
+  Hypothesis pvs_mask_mono : ∀ P, pvs M0 P ⊢ pvs M1 P.
+
+  (* We have invariants *)
+  Context (name : Type) (inv : name → iProp → iProp).
+  Hypothesis inv_persistent : ∀ i P, PersistentP (inv i P).
+  Hypothesis inv_alloc : ∀ P, P ⊢ pvs M1 (∃ i, inv i P).
+  Hypothesis inv_open :
+    ∀ i P Q R, (P ★ Q ⊢ pvs M0 (P ★ R)) → (inv i P ★ Q ⊢ pvs M1 R).
+
+  (* We have tokens for a little "two-state STS": [start] -> [finish].
+     state. [start] also asserts the exact state; it is only ever owned by the
+     invariant.  [finished] is duplicable. *)
+  (* Posssible implementations of these axioms:
+     * Using the STS monoid of a two-state STS, where [start] is the
+       authoritative saying the state is exactly [start], and [finish]
+       is the "we are at least in state [finish]" typically owned by threads.
+     * Ex () +_⊥ ()
+  *)
+  Context (gname : Type).
+  Context (start finished : gname → iProp).
+
+  Hypothesis sts_alloc : True ⊢ pvs M0 (∃ γ, start γ).
+  Hypotheses start_finish : ∀ γ, start γ ⊢ pvs M0 (finished γ).
+
+  Hypothesis finished_not_start : ∀ γ, start γ ★ finished γ ⊢ False.
+
+  Hypothesis finished_dup : ∀ γ, finished γ ⊢ finished γ ★ finished γ.
+
+  (* We assume that we cannot view shift to false. *)
+  Hypothesis soundness : ¬ (True ⊢ pvs M1 False).
+
+  (** Some general lemmas and proof mode compatibility. *)
+  Lemma inv_open' i P R : inv i P ★ (P -★ pvs M0 (P ★ pvs M1 R)) ⊢ pvs M1 R.
+  Proof.
+    iIntros "(#HiP & HP)". iApply pvs_pvs. iApply inv_open; last first.
+    { iSplit; first done. iExact "HP". }
+    iIntros "(HP & HPw)". by iApply "HPw".
+  Qed.
+
+  Instance pvs_mono' E : Proper ((⊢) ==> (⊢)) (pvs E).
+  Proof. intros P Q ?. by apply pvs_mono. Qed.
+  Instance pvs_proper E : Proper ((⊣⊢) ==> (⊣⊢)) (pvs E).
+  Proof.
+    intros P Q; rewrite !uPred.equiv_spec=> -[??]; split; by apply pvs_mono.
+  Qed.
+
+  Lemma pvs_frame_r E P Q : (pvs E P ★ Q) ⊢ pvs E (P ★ Q).
+  Proof. by rewrite comm pvs_frame_l comm. Qed.
+
+  Global Instance elim_pvs_pvs E P Q : ElimVs (pvs E P) P (pvs E Q) (pvs E Q).
+  Proof. by rewrite /ElimVs pvs_frame_r uPred.wand_elim_r pvs_pvs. Qed.
+
+  Global Instance elim_pvs0_pvs1 P Q : ElimVs (pvs M0 P) P (pvs M1 Q) (pvs M1 Q).
+  Proof.
+    by rewrite /ElimVs pvs_frame_r uPred.wand_elim_r pvs_mask_mono pvs_pvs.
+  Qed.
+
+  Global Instance exists_split_pvs0 {A} E P (Φ : A → iProp) :
+    FromExist P Φ → FromExist (pvs E P) (λ a, pvs E (Φ a)).
+  Proof.
+    rewrite /FromExist=>HP. apply uPred.exist_elim=> a.
+    apply pvs_mono. by rewrite -HP -(uPred.exist_intro a).
+  Qed.
+
+  (** Now to the actual counterexample. We start with a weird form of saved propositions. *)
+  Definition saved (γ : gname) (P : iProp) : iProp :=
+    ∃ i, inv i (start γ ∨ (finished γ ★ □ P)).
+  Global Instance saved_persistent γ P : PersistentP (saved γ P) := _.
+
+  Lemma saved_alloc (P : gname → iProp) : True ⊢ pvs M1 (∃ γ, saved γ (P γ)).
+  Proof.
+    iIntros "". iVs (sts_alloc) as (γ) "Hs".
+    iVs (inv_alloc (start γ ∨ (finished γ ★ □ (P γ))) with "[Hs]") as (i) "#Hi".
+    { auto. }
+    iApply pvs_intro. by iExists γ, i.
+  Qed.
+
+  Lemma saved_cast γ P Q : saved γ P ★ saved γ Q ★ □ P ⊢ pvs M1 (□ Q).
+  Proof.
+    iIntros "(#HsP & #HsQ & #HP)". iDestruct "HsP" as (i) "HiP".
+    iApply (inv_open' i). iSplit; first done.
+    iIntros "HaP". iAssert (pvs M0 (finished γ)) with "[HaP]" as "==> Hf".
+    { iDestruct "HaP" as "[Hs | [Hf _]]".
+      - by iApply start_finish.
+      - by iApply pvs_intro. }
+    iDestruct (finished_dup with "Hf") as "[Hf Hf']".
+    iApply pvs_intro. iSplitL "Hf'"; first by eauto.
+    (* Step 2: Open the Q-invariant. *)
+    iClear "HiP". clear i. iDestruct "HsQ" as (i) "HiQ".
+    iApply (inv_open' i). iSplit; first done.
+    iIntros "[HaQ | [_ #HQ]]".
+    { iExFalso. iApply finished_not_start. by iFrame. }
+    iApply pvs_intro. iSplitL "Hf".
+    { iRight. by iFrame. }
+    by iApply pvs_intro.
+  Qed.
+
+  (** And now we tie a bad knot. *)
+  Notation "¬ P" := (□ (P -★ pvs M1 False))%I : uPred_scope.
+  Definition A i : iProp := ∃ P, ¬P ★ saved i P.
+  Global Instance A_persistent i : PersistentP (A i) := _.
+
+  Lemma A_alloc : True ⊢ pvs M1 (∃ i, saved i (A i)).
+  Proof. by apply saved_alloc. Qed.
+
+  Lemma saved_NA i : saved i (A i) ⊢ ¬A i.
+  Proof.
+    iIntros "#Hi !# #HA". iPoseProof "HA" as "HA'".
+    iDestruct "HA'" as (P) "#[HNP Hi']".
+    iVs (saved_cast i (A i) P with "[]") as "HP".
+    { eauto. }
+    by iApply "HNP".
+  Qed.
+
+  Lemma saved_A i : saved i (A i) ⊢ A i.
+  Proof.
+    iIntros "#Hi". iExists (A i). iFrame "#".
+    by iApply saved_NA.
+  Qed.
+
+  Lemma contradiction : False.
+  Proof.
+    apply soundness. iIntros "".
+    iVs A_alloc as (i) "#H".
+    iPoseProof (saved_NA with "H") as "HN".
+    iApply "HN". iApply saved_A. done.
+  Qed.
+End inv. End inv.
diff --git a/program_logic/ectx_language.v b/program_logic/ectx_language.v
index d7b58df92eac2bd6fbc9ce115420f19454bfd68b..ea0adc138f5a5dcf3bff414099fc72a3252c5477 100644
--- a/program_logic/ectx_language.v
+++ b/program_logic/ectx_language.v
@@ -11,11 +11,11 @@ Class EctxLanguage (expr val ectx state : Type) := {
   empty_ectx : ectx;
   comp_ectx : ectx → ectx → ectx;
   fill : ectx → expr → expr;
-  head_step : expr → state → expr → state → option expr → Prop;
+  head_step : expr → state → expr → state → list expr → Prop;
 
   to_of_val v : to_val (of_val v) = Some v;
   of_to_val e v : to_val e = Some v → of_val v = e;
-  val_stuck e1 σ1 e2 σ2 ef : head_step e1 σ1 e2 σ2 ef → to_val e1 = None;
+  val_stuck e1 σ1 e2 σ2 efs : head_step e1 σ1 e2 σ2 efs → to_val e1 = None;
 
   fill_empty e : fill empty_ectx e = e;
   fill_comp K1 K2 e : fill K1 (fill K2 e) = fill (comp_ectx K1 K2) e;
@@ -28,10 +28,10 @@ Class EctxLanguage (expr val ectx state : Type) := {
   ectx_positive K1 K2 :
     comp_ectx K1 K2 = empty_ectx → K1 = empty_ectx ∧ K2 = empty_ectx;
 
-  step_by_val K K' e1 e1' σ1 e2 σ2 ef :
+  step_by_val K K' e1 e1' σ1 e2 σ2 efs :
     fill K e1 = fill K' e1' →
     to_val e1 = None →
-    head_step e1' σ1 e2 σ2 ef →
+    head_step e1' σ1 e2 σ2 efs →
     exists K'', K' = comp_ectx K K'';
 }.
 
@@ -57,16 +57,20 @@ Section ectx_language.
   Implicit Types (e : expr) (K : ectx).
 
   Definition head_reducible (e : expr) (σ : state) :=
-    ∃ e' σ' ef, head_step e σ e' σ' ef.
+    ∃ e' σ' efs, head_step e σ e' σ' efs.
 
   Inductive prim_step (e1 : expr) (σ1 : state)
-      (e2 : expr) (σ2 : state) (ef : option expr) : Prop :=
+      (e2 : expr) (σ2 : state) (efs : list expr) : Prop :=
     Ectx_step K e1' e2' :
       e1 = fill K e1' → e2 = fill K e2' →
-      head_step e1' σ1 e2' σ2 ef → prim_step e1 σ1 e2 σ2 ef.
+      head_step e1' σ1 e2' σ2 efs → prim_step e1 σ1 e2 σ2 efs.
 
-  Lemma val_prim_stuck e1 σ1 e2 σ2 ef :
-    prim_step e1 σ1 e2 σ2 ef → to_val e1 = None.
+  Lemma Ectx_step' K e1 σ1 e2 σ2 efs :
+    head_step e1 σ1 e2 σ2 efs → prim_step (fill K e1) σ1 (fill K e2) σ2 efs.
+  Proof. econstructor; eauto. Qed.
+
+  Lemma val_prim_stuck e1 σ1 e2 σ2 efs :
+    prim_step e1 σ1 e2 σ2 efs → to_val e1 = None.
   Proof. intros [??? -> -> ?]; eauto using fill_not_val, val_stuck. Qed.
 
   Canonical Structure ectx_lang : language := {|
@@ -78,29 +82,29 @@ Section ectx_language.
   |}.
 
   (* Some lemmas about this language *)
-  Lemma head_prim_step e1 σ1 e2 σ2 ef :
-    head_step e1 σ1 e2 σ2 ef → prim_step e1 σ1 e2 σ2 ef.
+  Lemma head_prim_step e1 σ1 e2 σ2 efs :
+    head_step e1 σ1 e2 σ2 efs → prim_step e1 σ1 e2 σ2 efs.
   Proof. apply Ectx_step with empty_ectx; by rewrite ?fill_empty. Qed.
 
   Lemma head_prim_reducible e σ : head_reducible e σ → reducible e σ.
-  Proof. intros (e'&σ'&ef&?). eexists e', σ', ef. by apply head_prim_step. Qed.
+  Proof. intros (e'&σ'&efs&?). eexists e', σ', efs. by apply head_prim_step. Qed.
 
   Lemma ectx_language_atomic e :
-    (∀ σ e' σ' ef, head_step e σ e' σ' ef → is_Some (to_val e')) →
+    (∀ σ e' σ' efs, head_step e σ e' σ' efs → is_Some (to_val e')) →
     (∀ K e', e = fill K e' → to_val e' = None → K = empty_ectx) →
     atomic e.
   Proof.
-    intros Hatomic_step Hatomic_fill σ e' σ' ef [K e1' e2' -> -> Hstep].
+    intros Hatomic_step Hatomic_fill σ e' σ' efs [K e1' e2' -> -> Hstep].
     assert (K = empty_ectx) as -> by eauto 10 using val_stuck.
     rewrite fill_empty. eapply Hatomic_step. by rewrite fill_empty.
   Qed.
 
-  Lemma head_reducible_prim_step e1 σ1 e2 σ2 ef :
-    head_reducible e1 σ1 → prim_step e1 σ1 e2 σ2 ef →
-    head_step e1 σ1 e2 σ2 ef.
+  Lemma head_reducible_prim_step e1 σ1 e2 σ2 efs :
+    head_reducible e1 σ1 → prim_step e1 σ1 e2 σ2 efs →
+    head_step e1 σ1 e2 σ2 efs.
   Proof.
-    intros (e2''&σ2''&ef''&?) [K e1' e2' -> -> Hstep].
-    destruct (step_by_val K empty_ectx e1' (fill K e1') σ1 e2'' σ2'' ef'')
+    intros (e2''&σ2''&efs''&?) [K e1' e2' -> -> Hstep].
+    destruct (step_by_val K empty_ectx e1' (fill K e1') σ1 e2'' σ2'' efs'')
       as [K' [-> _]%symmetry%ectx_positive];
       eauto using fill_empty, fill_not_val, val_stuck.
     by rewrite !fill_empty.
@@ -114,7 +118,7 @@ Section ectx_language.
     - intros ????? [K' e1' e2' Heq1 Heq2 Hstep].
       by exists (comp_ectx K K') e1' e2'; rewrite ?Heq1 ?Heq2 ?fill_comp.
     - intros e1 σ1 e2 σ2 ? Hnval [K'' e1'' e2'' Heq1 -> Hstep].
-      destruct (step_by_val K K'' e1 e1'' σ1 e2'' σ2 ef) as [K' ->]; eauto.
+      destruct (step_by_val K K'' e1 e1'' σ1 e2'' σ2 efs) as [K' ->]; eauto.
       rewrite -fill_comp in Heq1; apply (inj (fill _)) in Heq1.
       exists (fill K' e2''); rewrite -fill_comp; split; auto.
       econstructor; eauto.
diff --git a/program_logic/ectx_lifting.v b/program_logic/ectx_lifting.v
index 48a49739df9dbe8449b2f3e08a4ecb6a0b7bf7c3..0c35491a598320365077a68331d4ecbd3164f033 100644
--- a/program_logic/ectx_lifting.v
+++ b/program_logic/ectx_lifting.v
@@ -5,37 +5,36 @@ 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 efs, ■ head_step e1 σ1 e2 σ2 efs ★ ownP σ2
+          ={∅,E}=★ WP e2 @ E {{ Φ }} ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
+  ⊢ 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.
-  iSplit; first by eauto. iFrame. iNext. iIntros (e2 σ2 ef) "[% ?]".
+  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 efs) "[% ?]".
   iApply "Hwp". by eauto.
 Qed.
 
 Lemma wp_lift_pure_head_step E Φ e1 :
   to_val e1 = None →
   (∀ σ1, head_reducible e1 σ1) →
-  (∀ σ1 e2 σ2 ef, head_step e1 σ1 e2 σ2 ef → σ1 = σ2) →
-  (▷ ∀ e2 ef σ, ■ head_step e1 σ e2 σ ef → WP e2 @ E {{ Φ }} ★ wp_fork ef)
+  (∀ σ1 e2 σ2 efs, head_step e1 σ1 e2 σ2 efs → σ1 = σ2) →
+  (▷ ∀ e2 efs σ, ■ head_step e1 σ e2 σ efs →
+    WP e2 @ E {{ Φ }} ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
   ⊢ WP e1 @ E {{ Φ }}.
 Proof.
   iIntros (???) "H". iApply wp_lift_pure_step; eauto. iNext.
@@ -45,26 +44,50 @@ Qed.
 Lemma wp_lift_atomic_head_step {E Φ} e1 σ1 :
   atomic e1 →
   head_reducible e1 σ1 →
-  ▷ ownP σ1 ★ ▷ (∀ v2 σ2 ef,
-  ■ head_step e1 σ1 (of_val v2) σ2 ef ∧ ownP σ2 -★ (|={E}=> Φ v2) ★ wp_fork ef)
+  ▷ ownP σ1 ★ ▷ (∀ v2 σ2 efs,
+  ■ head_step e1 σ1 (of_val v2) σ2 efs ∧ ownP σ2 -★
+    (|={E}=> Φ v2) ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
   ⊢ WP e1 @ E {{ Φ }}.
 Proof.
   iIntros (??) "[? H]". iApply wp_lift_atomic_step; eauto. iFrame. iNext.
   iIntros (???) "[% ?]". iApply "H". eauto.
 Qed.
 
-Lemma wp_lift_atomic_det_head_step {E Φ e1} σ1 v2 σ2 ef :
+Lemma wp_lift_atomic_det_head_step {E Φ e1} σ1 v2 σ2 efs :
   atomic e1 →
   head_reducible e1 σ1 →
-  (∀ e2' σ2' ef', head_step e1 σ1 e2' σ2' ef' →
-    σ2 = σ2' ∧ to_val e2' = Some v2 ∧ ef = ef') →
-  ▷ ownP σ1 ★ ▷ (ownP σ2 -★ (|={E}=> Φ v2) ★ wp_fork ef) ⊢ WP e1 @ E {{ Φ }}.
+  (∀ e2' σ2' efs', head_step e1 σ1 e2' σ2' efs' →
+    σ2 = σ2' ∧ to_val e2' = Some v2 ∧ efs = efs') →
+  ▷ ownP σ1 ★ ▷ (ownP σ2 -★
+    (|={E}=> Φ v2) ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
+  ⊢ WP e1 @ E {{ Φ }}.
 Proof. eauto using wp_lift_atomic_det_step. Qed.
 
-Lemma wp_lift_pure_det_head_step {E Φ} e1 e2 ef :
+Lemma wp_lift_atomic_det_head_step' {E Φ e1} σ1 v2 σ2 :
+  atomic e1 →
+  head_reducible e1 σ1 →
+  (∀ e2' σ2' efs', head_step e1 σ1 e2' σ2' efs' →
+    σ2 = σ2' ∧ to_val e2' = Some v2 ∧ [] = efs') →
+  ▷ ownP σ1 ★ ▷ (ownP σ2 ={E}=★ Φ v2) ⊢ WP e1 @ E {{ Φ }}.
+Proof.
+  intros. rewrite -(wp_lift_atomic_det_head_step σ1 v2 σ2 [])
+    ?big_sepL_nil ?right_id; eauto.
+Qed.
+
+Lemma wp_lift_pure_det_head_step {E Φ} e1 e2 efs :
   to_val e1 = None →
   (∀ σ1, head_reducible e1 σ1) →
-  (∀ σ1 e2' σ2 ef', head_step e1 σ1 e2' σ2 ef' → σ1 = σ2 ∧ e2 = e2' ∧ ef = ef')→
-  ▷ (WP e2 @ E {{ Φ }} ★ wp_fork ef) ⊢ WP e1 @ E {{ Φ }}.
+  (∀ σ1 e2' σ2 efs', head_step e1 σ1 e2' σ2 efs' → σ1 = σ2 ∧ e2 = e2' ∧ efs = efs') →
+  ▷ (WP e2 @ E {{ Φ }} ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
+  ⊢ WP e1 @ E {{ Φ }}.
 Proof. eauto using wp_lift_pure_det_step. Qed.
+
+Lemma wp_lift_pure_det_head_step' {E Φ} e1 e2 :
+  to_val e1 = None →
+  (∀ σ1, head_reducible e1 σ1) →
+  (∀ σ1 e2' σ2 efs', head_step e1 σ1 e2' σ2 efs' → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
+  ▷ WP e2 @ E {{ Φ }} ⊢ WP e1 @ E {{ Φ }}.
+Proof.
+  intros. rewrite -(wp_lift_pure_det_step e1 e2 []) ?big_sepL_nil ?right_id; eauto.
+Qed.
 End wp.
diff --git a/program_logic/ectxi_language.v b/program_logic/ectxi_language.v
index 4a8f62fd85863f80c4946907b42602a29e98fc3d..4dc3bbfd7a6316866b5fbfb2a5edbaddc1db1f3c 100644
--- a/program_logic/ectxi_language.v
+++ b/program_logic/ectxi_language.v
@@ -9,11 +9,11 @@ Class EctxiLanguage (expr val ectx_item state : Type) := {
   of_val : val → expr;
   to_val : expr → option val;
   fill_item : ectx_item → expr → expr;
-  head_step : expr → state → expr → state → option expr → Prop;
+  head_step : expr → state → expr → state → list expr → Prop;
 
   to_of_val v : to_val (of_val v) = Some v;
   of_to_val e v : to_val e = Some v → of_val v = e;
-  val_stuck e1 σ1 e2 σ2 ef : head_step e1 σ1 e2 σ2 ef → to_val e1 = None;
+  val_stuck e1 σ1 e2 σ2 efs : head_step e1 σ1 e2 σ2 efs → to_val e1 = None;
 
   fill_item_inj Ki :> Inj (=) (=) (fill_item Ki);
   fill_item_val Ki e : is_Some (to_val (fill_item Ki e)) → is_Some (to_val e);
@@ -21,8 +21,8 @@ Class EctxiLanguage (expr val ectx_item state : Type) := {
     to_val e1 = None → to_val e2 = None →
     fill_item Ki1 e1 = fill_item Ki2 e2 → Ki1 = Ki2;
 
-  head_ctx_step_val Ki e σ1 e2 σ2 ef :
-    head_step (fill_item Ki e) σ1 e2 σ2 ef → is_Some (to_val e);
+  head_ctx_step_val Ki e σ1 e2 σ2 efs :
+    head_step (fill_item Ki e) σ1 e2 σ2 efs → is_Some (to_val e);
 }.
 
 Arguments of_val {_ _ _ _ _} _.
@@ -46,6 +46,9 @@ Section ectxi_language.
 
   Definition fill (K : ectx) (e : expr) : expr := fold_right fill_item e K.
 
+  Lemma fill_app (K1 K2 : ectx) e : fill (K1 ++ K2) e = fill K1 (fill K2 e).
+  Proof. apply fold_right_app. Qed.
+
   Instance fill_inj K : Inj (=) (=) (fill K).
   Proof. red; induction K as [|Ki K IH]; naive_solver. Qed.
 
@@ -60,8 +63,8 @@ Section ectxi_language.
   (* When something does a step, and another decomposition of the same expression
   has a non-val [e] in the hole, then [K] is a left sub-context of [K'] - in
   other words, [e] also contains the reducible expression *)
-  Lemma step_by_val K K' e1 e1' σ1 e2 σ2 ef :
-    fill K e1 = fill K' e1' → to_val e1 = None → head_step e1' σ1 e2 σ2 ef →
+  Lemma step_by_val K K' e1 e1' σ1 e2 σ2 efs :
+    fill K e1 = fill K' e1' → to_val e1 = None → head_step e1' σ1 e2 σ2 efs →
     exists K'', K' = K ++ K''. (* K `prefix_of` K' *)
   Proof.
     intros Hfill Hred Hnval; revert K' Hfill.
@@ -78,7 +81,7 @@ Section ectxi_language.
        fixed fields of different records, even when I did not. *)
     Build_EctxLanguage expr val ectx state of_val to_val [] (++) fill head_step _ _ _ _ _ _ _ _ _.
   Solve Obligations with eauto using to_of_val, of_to_val, val_stuck,
-    fill_not_val, step_by_val, fold_right_app, app_eq_nil.
+    fill_not_val, fill_app, step_by_val, fold_right_app, app_eq_nil.
 
   Global Instance ectxi_lang_ctx_item Ki :
     LanguageCtx (ectx_lang expr) (fill_item Ki).
diff --git a/program_logic/ghost_ownership.v b/program_logic/ghost_ownership.v
index cd679ccdfbbe142c24487cdc4c060ef39ab91eab..1f9e4859c32cad0ccc5de5cacd6e293e14f32e02 100644
--- a/program_logic/ghost_ownership.v
+++ b/program_logic/ghost_ownership.v
@@ -1,39 +1,64 @@
-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 model.
+From iris.algebra Require Import iprod gmap.
 Import uPred.
 
-Definition own_def `{inG Λ Σ A} (γ : gname) (a : A) : iPropG Λ Σ :=
-  ownG (to_globalF γ a).
+(** The class [inG Σ A] expresses that the CMRA [A] is in the list of functors
+[Σ]. This class is similar to the [subG] class, but written down in terms of
+individual CMRAs instead of (lists of) CMRA *functors*. This additional class is
+needed because Coq is otherwise unable to solve type class constraints due to
+higher-order unification problems. *)
+Class inG (Σ : gFunctors) (A : cmraT) :=
+  InG { inG_id : gid Σ; inG_prf : A = Σ inG_id (iPreProp Σ) }.
+Arguments inG_id {_ _} _.
+
+Lemma subG_inG Σ (F : gFunctor) : subG F Σ → inG Σ (F (iPreProp Σ)).
+Proof. move=> /(_ 0%fin) /= [j ->]. by exists j. Qed.
+
+(** * Definition of the connective [own] *)
+Definition iRes_singleton `{i : inG Σ A} (γ : gname) (a : A) : iResUR Σ :=
+  iprod_singleton (inG_id i) {[ γ := cmra_transport inG_prf a ]}.
+Instance: Params (@iRes_singleton) 4.
+
+Definition own_def `{inG Σ A} (γ : gname) (a : A) : iProp Σ :=
+  uPred_ownM (iRes_singleton γ 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 *)
+(** * 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 _ γ).
+(** ** Properties of [iRes_singleton] *)
+Global Instance iRes_singleton_ne γ n :
+  Proper (dist n ==> dist n) (@iRes_singleton Σ A _ γ).
+Proof. by intros a a' Ha; apply iprod_singleton_ne; rewrite Ha. Qed.
+Lemma iRes_singleton_op γ a1 a2 :
+  iRes_singleton γ (a1 ⋅ a2) ≡ iRes_singleton γ a1 ⋅ iRes_singleton γ a2.
+Proof.
+  by rewrite /iRes_singleton iprod_op_singleton op_singleton cmra_transport_op.
+Qed.
+
+(** ** Properties of [own] *)
+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 iRes_singleton_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 /iRes_singleton.
   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.
@@ -44,54 +69,57 @@ Proof. rewrite !own_eq /own_def; apply _. Qed.
 Global Instance own_core_persistent γ a : Persistent a → PersistentP (own γ a).
 Proof. rewrite !own_eq /own_def; apply _. Qed.
 
+(** ** Allocation *)
 (* 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 = iRes_singleton γ 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'.
+(** ** Frame preserving updates *)
+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 = iRes_singleton γ 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.
 
-Section global_empty.
-Context `{i : inG Λ Σ (A:ucmraT)}.
-Implicit Types a : A.
+Arguments own_valid {_ _} [_] _ _.
+Arguments own_valid_l {_ _} [_] _ _.
+Arguments own_valid_r {_ _} [_] _ _.
+Arguments own_updateP {_ _} [_] _ _ _ _.
+Arguments own_update {_ _} [_] _ _ _ _.
 
-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
deleted file mode 100644
index 48403855e61fd192d625c7e54af8377c8c4c33c5..0000000000000000000000000000000000000000
--- a/program_logic/global_functor.v
+++ /dev/null
@@ -1,140 +0,0 @@
-From iris.algebra Require Export iprod.
-From iris.program_logic Require Export model.
-
-(** 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 {
-  inG_id : gid Σ;
-  inG_prf : A = projT2 Σ inG_id (iPreProp Λ (globalF Σ))
-}.
-Arguments inG_id {_ _ _} _.
-Hint Mode inG - - + : typeclass_instances.
-
-Definition to_globalF `{i : inG Λ Σ A} (γ : gname) (a : A) : iGst Λ (globalF Σ) :=
-  iprod_singleton (inG_id i) {[ γ := cmra_transport inG_prf a ]}.
-Instance: Params (@to_globalF) 5.
-Typeclasses Opaque to_globalF.
-
-(** * Properties of to_globalC *)
-Section to_globalF.
-Context `{i : inG Λ Σ A}.
-Implicit Types a : A.
-
-Global Instance to_globalF_ne γ n :
-  Proper (dist n ==> dist n) (@to_globalF Λ Σ 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.
-Proof.
-  by rewrite /to_globalF 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.
-
-(** When instantiating the logic, we want to just plug together the
-    requirements exported by the modules we use. To this end, we construct
-    the "gFunctors" from a "list gFunctor", and have some typeclass magic
-    to infer the inG. *)
-Module gFunctorList.
-  Inductive gFunctorList :=
-    | nil  : gFunctorList
-    | cons : gFunctor → gFunctorList → gFunctorList.
-  Coercion singleton (F : gFunctor) : gFunctorList := cons F nil.
-
-  Fixpoint fold_right {A} (f : gFunctor → A → A) (a : A) (Fs : gFunctorList) : A :=
-    match Fs with
-    | nil => a
-    | cons F Fs => f F (fold_right f a Fs)
-    end.
-End gFunctorList.
-Notation gFunctorList := gFunctorList.gFunctorList.
-
-Delimit Scope gFunctor_scope with gFunctor.
-Bind Scope gFunctor_scope with gFunctorList.
-Arguments gFunctorList.cons _ _%gFunctor.
-
-Notation "[ ]" := gFunctorList.nil (format "[ ]") : gFunctor_scope.
-Notation "[ F ]" := (gFunctorList.cons F gFunctorList.nil) : gFunctor_scope.
-Notation "[ F1 ; F2 ; .. ; Fn ]" :=
-  (gFunctorList.cons F1 (gFunctorList.cons F2 ..
-    (gFunctorList.cons Fn gFunctorList.nil) ..)) : gFunctor_scope.
-
-Module gFunctors.
-  Definition nil : gFunctors := existT 0 (fin_0_inv _).
-
-  Definition cons (F : gFunctor) (Σ : gFunctors) : gFunctors :=
-   existT (S (projT1 Σ)) (fin_S_inv _ F (projT2 Σ)).
-
-  Fixpoint app (Fs : gFunctorList) (Σ : gFunctors) : gFunctors :=
-    match Fs with
-    | gFunctorList.nil => Σ
-    | gFunctorList.cons F Fs => cons F (app Fs Σ)
-    end.
-End gFunctors.
-
-(** Cannot bind this scope with the [gFunctorG] since [gFunctorG] is a
-notation hiding a more complex type. *)
-Notation "#[ ]" := gFunctors.nil (format "#[ ]").
-Notation "#[ Fs ]" := (gFunctors.app Fs gFunctors.nil).
-Notation "#[ Fs1 ; Fs2 ; .. ; Fsn ]" :=
-  (gFunctors.app Fs1 (gFunctors.app Fs2 ..
-    (gFunctors.app Fsn gFunctors.nil) ..)).
-
-(** 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 {
-  inGF_id : gid Σ;
-  inGF_prf : F = projT2 Σ inGF_id;
-}.
-(* Avoid eager type class search: this line ensures that type class search
-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.
-
-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.
-Proof. by exists 0%fin. Qed.
-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.
-
-Instance inGFs_nil {Λ Σ} : inGFs Λ Σ [].
-Proof. exact tt. Qed.
-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..817a21e295c66caeb711b88f17ce6c2d22be4a35 100644
--- a/program_logic/hoare.v
+++ b/program_logic/hoare.v
@@ -1,79 +1,79 @@
 From iris.program_logic Require Export weakestpre viewshifts.
-From iris.proofmode Require Import weakestpre invariants.
+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 Φ)
+Notation "{{ P } } e @ E {{ Φ } }" := (ht E P e%E Φ)
   (at level 20, P, e, Φ at level 200,
    format "{{  P  } }  e  @  E  {{  Φ  } }") : uPred_scope.
-Notation "{{ P } } e {{ Φ } }" := (ht ⊤ P e Φ)
+Notation "{{ P } } e {{ Φ } }" := (ht ⊤ P e%E Φ)
   (at level 20, P, e, Φ at level 200,
    format "{{  P  } }  e  {{  Φ  } }") : uPred_scope.
-Notation "{{ P } } e @ E {{ Φ } }" := (True ⊢ ht E P e Φ)
+Notation "{{ P } } e @ E {{ Φ } }" := (True ⊢ ht E P e%E Φ)
   (at level 20, P, e, Φ at level 200,
    format "{{  P  } }  e  @  E  {{  Φ  } }") : C_scope.
-Notation "{{ P } } e {{ Φ } }" := (True ⊢ ht ⊤ P e Φ)
+Notation "{{ P } } e {{ Φ } }" := (True ⊢ ht ⊤ P e%E Φ)
   (at level 20, P, e, Φ at level 200,
    format "{{  P  } }  e  {{  Φ  } }") : C_scope.
 
-Notation "{{ P } } e @ E {{ v , Q } }" := (ht E P e (λ v, Q))
+Notation "{{ P } } e @ E {{ v , Q } }" := (ht E P e%E (λ v, Q))
   (at level 20, P, e, Q at level 200,
    format "{{  P  } }  e  @  E  {{  v ,  Q  } }") : uPred_scope.
-Notation "{{ P } } e {{ v , Q } }" := (ht ⊤ P e (λ v, Q))
+Notation "{{ P } } e {{ v , Q } }" := (ht ⊤ P e%E (λ v, Q))
   (at level 20, P, e, Q at level 200,
    format "{{  P  } }  e  {{  v ,  Q  } }") : uPred_scope.
-Notation "{{ P } } e @ E {{ v , Q } }" := (True ⊢ ht E P e (λ v, Q))
+Notation "{{ P } } e @ E {{ v , Q } }" := (True ⊢ ht E P e%E (λ v, Q))
   (at level 20, P, e, Q at level 200,
    format "{{  P  } }  e  @  E  {{  v ,  Q  } }") : C_scope.
-Notation "{{ P } } e {{ v , Q } }" := (True ⊢ ht ⊤ P e (λ v, Q))
+Notation "{{ P } } e {{ v , Q } }" := (True ⊢ ht ⊤ P e%E (λ v, Q))
   (at level 20, P, e, Q at level 200,
    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.
+Proof. iIntros (Hwp) "!# HP". by iApply Hwp. Qed.
 
-Lemma ht_val E v : {{ True : iProp Λ Σ }} of_val v @ E {{ v', v = v' }}.
-Proof. iIntros "! _". by iApply wp_value'. Qed.
+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 @ 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 →
+  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.
@@ -82,7 +82,7 @@ Lemma ht_bind `{LanguageCtx Λ K} E P Φ Φ' e :
   {{ P }} e @ E {{ Φ }} ∧ (∀ v, {{ Φ v }} K (of_val v) @ E {{ Φ' }})
   ⊢ {{ P }} K e @ E {{ Φ' }}.
 Proof.
-  iIntros "(#Hwpe&#HwpK) ! HP". iApply wp_bind.
+  iIntros "[#Hwpe #HwpK] !# HP". iApply wp_bind.
   iApply wp_wand_r; iSplitL; [by iApply "Hwpe"|].
   iIntros (v) "Hv". by iApply "HwpK".
 Qed.
@@ -90,45 +90,43 @@ 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.
 
 Lemma ht_frame_l E P Φ R e :
   {{ P }} e @ E {{ Φ }} ⊢ {{ R ★ P }} e @ E {{ v, R ★ Φ v }}.
-Proof. iIntros "#Hwp ! [$ HP]". by iApply "Hwp". Qed.
+Proof. iIntros "#Hwp !# [$ HP]". by iApply "Hwp". Qed.
 
 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.
+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 Φ :
   to_val e = None →
   {{ P }} e @ E {{ Φ }} ⊢ {{ ▷ R ★ P }} e @ E {{ v, R ★ Φ v }}.
 Proof.
-  iIntros (?) "#Hwp ! [HR HP]".
+  iIntros (?) "#Hwp !# [HR HP]".
   iApply wp_frame_step_l'; try done. iFrame "HR". by iApply "Hwp".
 Qed.
 
@@ -136,15 +134,7 @@ Lemma ht_frame_step_r' E P Φ R e :
   to_val e = None →
   {{ P }} e @ E {{ Φ }} ⊢ {{ P ★ ▷ R }} e @ E {{ v, Φ v ★ R }}.
 Proof.
-  iIntros (?) "#Hwp ! [HP HR]".
+  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 4c91e6fa64851b29ef82f4375ece0cdc8747a50d..ca670bf1bf1eec7cff02d5bfc26489169f537b35 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_except_last; 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. { 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; iApply except_last_intro.
+  iDestruct (ownI_open i P with "[Hw HE]") as "($ & $ & HD)"; first by iFrame.
+  iIntros "HP [Hw $] !==>"; iApply except_last_intro. 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.
+  iIntros "!==> {$HP} HP". iApply "Hclose"; auto.
 Qed.
 End inv.
diff --git a/program_logic/iris.v b/program_logic/iris.v
new file mode 100644
index 0000000000000000000000000000000000000000..8269a5982cd453b9cd25b11edcf28bebd834ac67
--- /dev/null
+++ b/program_logic/iris.v
@@ -0,0 +1,30 @@
+From iris.program_logic Require Export ghost_ownership language.
+From iris.prelude Require Export coPset.
+From iris.algebra Require Import gmap auth agree gset coPset upred_big_op.
+
+Class irisPreG (Λ : language) (Σ : gFunctors) : Set := IrisPreG {
+  state_inG :> inG Σ (authR (optionUR (exclR (stateC Λ))));
+  invariant_inG :> inG Σ (authR (gmapUR positive (agreeR (laterC (iPreProp Σ)))));
+  enabled_inG :> inG Σ coPset_disjR;
+  disabled_inG :> inG Σ (gset_disjR positive);
+}.
+
+Class irisG (Λ : language) (Σ : gFunctors) : Set := IrisG {
+  iris_pre_inG :> irisPreG Λ Σ;
+  state_name : gname;
+  invariant_name : gname;
+  enabled_name : gname;
+  disabled_name : gname;
+}.
+
+Definition irisΣ (Λ : language) : gFunctors :=
+  #[GFunctor (constRF (authUR (optionUR (exclR (stateC Λ)))));
+    GFunctor (authRF (gmapURF positive (agreeRF (laterCF idCF))));
+    GFunctor (constRF coPset_disjUR);
+    GFunctor (constRF (gset_disjUR positive))].
+
+Instance subG_irisΣ {Σ Λ} : subG (irisΣ Λ) Σ → irisPreG Λ Σ.
+Proof.
+  intros [?%subG_inG [?%subG_inG
+    [?%subG_inG ?%subG_inG]%subG_inv]%subG_inv]%subG_inv; by constructor.
+Qed.
\ No newline at end of file
diff --git a/program_logic/language.v b/program_logic/language.v
index 33600f6dfce705774c1a66d601f6bfcb7c46e4ae..cd33108da2ea94c753a0081d1361f5565391fca9 100644
--- a/program_logic/language.v
+++ b/program_logic/language.v
@@ -6,11 +6,15 @@ Structure language := Language {
   state : Type;
   of_val : val → expr;
   to_val : expr → option val;
-  prim_step : expr → state → expr → state → option expr → Prop;
+  prim_step : expr → state → expr → state → list expr → Prop;
   to_of_val v : to_val (of_val v) = Some v;
   of_to_val e v : to_val e = Some v → of_val v = e;
-  val_stuck e σ e' σ' ef : prim_step e σ e' σ' ef → to_val e = None
+  val_stuck e σ e' σ' efs : prim_step e σ e' σ' efs → to_val e = None
 }.
+Delimit Scope expr_scope with E.
+Delimit Scope val_scope with V.
+Bind Scope expr_scope with expr.
+Bind Scope expr_scope with val.
 Arguments of_val {_} _.
 Arguments to_val {_} _.
 Arguments prim_step {_} _ _ _ _ _.
@@ -29,31 +33,31 @@ Section language.
   Implicit Types v : val Λ.
 
   Definition reducible (e : expr Λ) (σ : state Λ) :=
-    ∃ e' σ' ef, prim_step e σ e' σ' ef.
+    ∃ e' σ' efs, prim_step e σ e' σ' efs.
   Definition atomic (e : expr Λ) : Prop :=
-    ∀ σ e' σ' ef, prim_step e σ e' σ' ef → is_Some (to_val e').
+    ∀ σ e' σ' efs, prim_step e σ e' σ' efs → is_Some (to_val e').
   Inductive step (ρ1 ρ2 : cfg Λ) : Prop :=
-    | step_atomic e1 σ1 e2 σ2 ef t1 t2 :
+    | step_atomic e1 σ1 e2 σ2 efs t1 t2 :
        ρ1 = (t1 ++ e1 :: t2, σ1) →
-       ρ2 = (t1 ++ e2 :: t2 ++ option_list ef, σ2) →
-       prim_step e1 σ1 e2 σ2 ef →
+       ρ2 = (t1 ++ e2 :: t2 ++ efs, σ2) →
+       prim_step e1 σ1 e2 σ2 efs →
        step ρ1 ρ2.
 
   Lemma of_to_val_flip v e : of_val v = e → to_val e = Some v.
   Proof. intros <-. by rewrite to_of_val. Qed.
   Lemma reducible_not_val e σ : reducible e σ → to_val e = None.
   Proof. intros (?&?&?&?); eauto using val_stuck. Qed.
-  Global Instance: Inj (=) (=) (@of_val Λ).
+  Global Instance of_val_inj : Inj (=) (=) (@of_val Λ).
   Proof. by intros v v' Hv; apply (inj Some); rewrite -!to_of_val Hv. Qed.
 End language.
 
 Class LanguageCtx (Λ : language) (K : expr Λ → expr Λ) := {
   fill_not_val e :
     to_val e = None → to_val (K e) = None;
-  fill_step e1 σ1 e2 σ2 ef :
-    prim_step e1 σ1 e2 σ2 ef →
-    prim_step (K e1) σ1 (K e2) σ2 ef;
-  fill_step_inv e1' σ1 e2 σ2 ef :
-    to_val e1' = None → prim_step (K e1') σ1 e2 σ2 ef →
-    ∃ e2', e2 = K e2' ∧ prim_step e1' σ1 e2' σ2 ef
+  fill_step e1 σ1 e2 σ2 efs :
+    prim_step e1 σ1 e2 σ2 efs →
+    prim_step (K e1) σ1 (K e2) σ2 efs;
+  fill_step_inv e1' σ1 e2 σ2 efs :
+    to_val e1' = None → prim_step (K e1') σ1 e2 σ2 efs →
+    ∃ e2', e2 = K e2' ∧ prim_step e1' σ1 e2' σ2 efs
 }.
diff --git a/program_logic/lifting.v b/program_logic/lifting.v
index 6b98cf087a3aea1de5f51eb74755aa4bd9e2e7f3..1dcd9a8b070e1170939df6d504072b4e0e9efd20 100644
--- a/program_logic/lifting.v
+++ b/program_logic/lifting.v
@@ -1,99 +1,86 @@
 From iris.program_logic Require Export weakestpre.
-From iris.program_logic Require Import wsat ownership.
+From iris.program_logic Require Import ownership.
+From iris.algebra Require Export upred_big_op.
 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 efs, ■ prim_step e1 σ1 e2 σ2 efs ★ ownP σ2
+          ={∅,E}=★ WP e2 @ E {{ Φ }} ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
+  ⊢ 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)".
+  iDestruct (ownP_agree σ1 σ1' with "[#]") as %<-; first by iFrame.
+  iVsIntro; iSplit; [done|]; iNext; iIntros (e2 σ2 efs Hstep).
+  iVs (ownP_update σ1 σ2 with "[-H]") as "[$ ?]"; first by iFrame.
+  iApply "H"; eauto.
 Qed.
 
 Lemma wp_lift_pure_step E Φ e1 :
   to_val e1 = None →
   (∀ σ1, reducible e1 σ1) →
-  (∀ σ1 e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef → σ1 = σ2) →
-  (▷ ∀ e2 ef σ, ■ prim_step e1 σ e2 σ ef → WP e2 @ E {{ Φ }} ★ wp_fork ef)
+  (∀ σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs → σ1 = σ2) →
+  (▷ ∀ e2 efs σ, ■ prim_step e1 σ e2 σ efs →
+    WP e2 @ E {{ Φ }} ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
   ⊢ 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 efs ?).
+  destruct (Hstep σ1 e2 σ2 efs); 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 →
-  ▷ ownP σ1 ★ ▷ (∀ v2 σ2 ef,
-    ■ prim_step e1 σ1 (of_val v2) σ2 ef ∧ ownP σ2 -★ (|={E}=> Φ v2) ★ wp_fork ef)
+  ▷ ownP σ1 ★ ▷ (∀ v2 σ2 efs, ■ prim_step e1 σ1 (of_val v2) σ2 efs ∧ ownP σ2 -★
+    (|={E}=> Φ v2) ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
   ⊢ 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 efs) "[% Hσ]".
+  edestruct (Hatomic σ1 e2 σ2 efs) as [v2 <-%of_to_val]; eauto.
+  iDestruct ("H" $! v2 σ2 efs 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 :
+Lemma wp_lift_atomic_det_step {E Φ e1} σ1 v2 σ2 efs :
   atomic e1 →
   reducible e1 σ1 →
-  (∀ e2' σ2' ef', prim_step e1 σ1 e2' σ2' ef' →
-                  σ2 = σ2' ∧ to_val e2' = Some v2 ∧ ef = ef') →
-  ▷ ownP σ1 ★ ▷ (ownP σ2 -★ (|={E}=> Φ v2) ★ wp_fork ef) ⊢ WP e1 @ E {{ Φ }}.
+  (∀ e2' σ2' efs', prim_step e1 σ1 e2' σ2' efs' →
+                   σ2 = σ2' ∧ to_val e2' = Some v2 ∧ efs = efs') →
+  ▷ ownP σ1 ★ ▷ (ownP σ2 -★
+    (|={E}=> Φ v2) ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
+  ⊢ WP e1 @ E {{ Φ }}.
 Proof.
   iIntros (?? Hdet) "[Hσ1 Hσ2]". iApply (wp_lift_atomic_step _ σ1); try done.
-  iFrame. iNext. iIntros (v2' σ2' ef') "[% Hσ2']".
+  iFrame. iNext. iIntros (v2' σ2' efs') "[% Hσ2']".
   edestruct Hdet as (->&->%of_to_val%(inj of_val)&->). done. by iApply "Hσ2".
 Qed.
 
-Lemma wp_lift_pure_det_step {E Φ} e1 e2 ef :
+Lemma wp_lift_pure_det_step {E Φ} e1 e2 efs :
   to_val e1 = None →
   (∀ σ1, reducible e1 σ1) →
-  (∀ σ1 e2' σ2 ef', prim_step e1 σ1 e2' σ2 ef' → σ1 = σ2 ∧ e2 = e2' ∧ ef = ef')→
-  ▷ (WP e2 @ E {{ Φ }} ★ wp_fork ef) ⊢ WP e1 @ E {{ Φ }}.
+  (∀ σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' → σ1 = σ2 ∧ e2 = e2' ∧ efs = efs')→
+  ▷ (WP e2 @ E {{ Φ }} ★ [★ list] ef ∈ efs, WP ef {{ _, True }})
+  ⊢ WP e1 @ E {{ Φ }}.
 Proof.
   iIntros (?? Hpuredet) "?". iApply (wp_lift_pure_step E); try done.
-  by intros; eapply Hpuredet. iNext. by iIntros (e' ef' σ (_&->&->)%Hpuredet).
+  by intros; eapply Hpuredet. iNext. by iIntros (e' efs' σ (_&->&->)%Hpuredet).
 Qed.
 End lifting.
diff --git a/program_logic/model.v b/program_logic/model.v
index 55cdb0ac6f870c12c7e3614ab91ac8afc1e19c41..f0199b825585af98a1bca4d4d437eac16808a5e5 100644
--- a/program_logic/model.v
+++ b/program_logic/model.v
@@ -1,61 +1,156 @@
 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
+(** In this file we construct the type [iProp] of propositions of the Iris
+logic. This is done by solving the following recursive domain equation:
+
+  iProp ≈ uPred (∀ i : gid, gname -fin-> (Σ i) iProp)
+
+where:
+
+  Σ : gFunctors  := lists of locally constractive functors
+  i : gid        := indexes addressing individual functors in [Σ]
+  γ : gname      := ghost variable names
+
+The Iris logic is parametrized by a list of locally contractive functors [Σ]
+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. *)
+
+
+(** * Locally contractive functors *)
+(** The type [gFunctor] bundles a functor from the category of COFEs to the
+category of CMRAs with a proof that it is locally contractive. *)
+Structure gFunctor := GFunctor {
+  gFunctor_F :> rFunctor;
+  gFunctor_contractive : rFunctorContractive gFunctor_F;
 }.
-Arguments IFunctor _ {_}.
-Existing Instances iFunctor_contractive.
+Arguments GFunctor _ {_}.
+Existing Instance gFunctor_contractive.
+
+(** The type [gFunctors] describes the parameters [Σ] of the Iris logic: lists
+of [gFunctor]s.
+
+Note that [gFunctors] is isomorphic to [list gFunctor], but defined in an
+alternative way to avoid universe inconsistencies with respect to the universe
+monomorphic [list] type. *)
+Definition gFunctors := { n : nat & fin n → gFunctor }.
+
+Definition gid (Σ : gFunctors) := fin (projT1 Σ).
+Definition gFunctors_lookup (Σ : gFunctors) : gid Σ → gFunctor := projT2 Σ.
+Coercion gFunctors_lookup : gFunctors >-> Funclass.
+
+Definition gname := positive.
+
+(** The resources functor [iResF Σ A := ∀ i : gid, gname -fin-> (Σ i) A]. *)
+Definition iResF (Σ : gFunctors) : urFunctor :=
+  iprodURF (λ i, gmapURF gname (Σ i)).
+
+
+(** We define functions for the empty list of functors, the singleton list of
+functors, and the append operator on lists of functors. These are used to
+compose [gFunctors] out of smaller pieces. *)
+Module gFunctors.
+  Definition nil : gFunctors := existT 0 (fin_0_inv _).
 
+  Definition singleton (F : gFunctor) : gFunctors :=
+    existT 1 (fin_S_inv (λ _, gFunctor) F (fin_0_inv _)).
+
+  Definition app (Σ1 Σ2 : gFunctors) : gFunctors :=
+    existT (projT1 Σ1 + projT1 Σ2) (fin_plus_inv _ (projT2 Σ1) (projT2 Σ2)).
+End gFunctors.
+
+Coercion gFunctors.singleton : gFunctor >-> gFunctors.
+Notation "#[ ]" := gFunctors.nil (format "#[ ]").
+Notation "#[ Σ1 ; .. ; Σn ]" :=
+  (gFunctors.app Σ1 .. (gFunctors.app Σn gFunctors.nil) ..).
+
+
+(** * Subfunctors *)
+(** In order to make proofs in the Iris logic modular, they are not done with
+respect to some concrete list of functors [Σ], but are instead parametrized by
+an arbitrary list of functors [Σ] that contains at least certain functors. For
+example, the lock library is parameterized by a functor [Σ] that should have
+the functors corresponding to the heap and the exclusive monoid to manage to
+lock invariant.
+
+The contraints to can be expressed using the type class [subG Σ1 Σ2], which
+expresses that the functors [Σ1] are contained in [Σ2]. *)
+Class subG (Σ1 Σ2 : gFunctors) := in_subG i : { j | Σ1 i = Σ2 j }.
+
+(** Avoid trigger happy type class search: this line ensures that type class
+search is only triggered if the arguments of [subG] do not contain evars. Since
+instance search for [subG] is restrained, instances should always have [subG] as
+their first parameter to avoid loops. For example, the instances [subG_authΣ]
+and [auth_discrete] otherwise create a cycle that pops up arbitrarily. *)
+Hint Mode subG + + : typeclass_instances.
+
+Lemma subG_inv Σ1 Σ2 Σ : subG (gFunctors.app Σ1 Σ2) Σ → subG Σ1 Σ * subG Σ2 Σ.
+Proof.
+  move=> H; split.
+  - move=> i; move: H=> /(_ (Fin.L _ i)) [j] /=. rewrite fin_plus_inv_L; eauto.
+  - move=> i; move: H=> /(_ (Fin.R _ i)) [j] /=. rewrite fin_plus_inv_R; eauto.
+Qed.
+
+Instance subG_refl Σ : subG Σ Σ.
+Proof. move=> i; by exists i. Qed.
+Instance subG_app_l Σ Σ1 Σ2 : subG Σ Σ1 → subG Σ (gFunctors.app Σ1 Σ2).
+Proof.
+  move=> H i; move: H=> /(_ i) [j ?].
+  exists (Fin.L _ j). by rewrite /= fin_plus_inv_L.
+Qed.
+Instance subG_app_r Σ Σ1 Σ2 : subG Σ Σ2 → subG Σ (gFunctors.app Σ1 Σ2).
+Proof.
+  move=> H i; move: H=> /(_ i) [j ?].
+  exists (Fin.R _ j). by rewrite /= fin_plus_inv_R.
+Qed.
+
+
+(** * Solution of the recursive domain equation *)
+(** We first declare a module type and then an instance of it so as to seal all of
+the construction, this way we are sure we do not use any properties of the
+construction, and also avoid Coq from blindly unfolding it. *)
 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 Λ Σ),
-  iProp_fold (iProp_unfold P) ≡ P.
-Parameter iProp_unfold_fold: ∀ {Λ Σ} (P : iPreProp Λ Σ),
-  iProp_unfold (iProp_fold P) ≡ P.
+  Parameter iPreProp : gFunctors → cofeT.
+  Definition iResUR (Σ : gFunctors) : ucmraT :=
+    iprodUR (λ i, gmapUR gname (Σ i (iPreProp Σ))).
+  Notation iProp Σ := (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 Σ),
+    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.
-Proof. apply solution_unfold_fold. Qed.
-Lemma iProp_unfold_fold {Λ Σ} (P : iPreProp Λ Σ) :
-  iProp_unfold (iProp_fold P) ≡ P.
-Proof. apply solution_fold_unfold. Qed.
+  Import cofe_solver.
+  Definition iProp_result (Σ : gFunctors) :
+    solution (uPredCF (iResF Σ)) := solver.result _.
+
+  Definition iPreProp (Σ : gFunctors) : cofeT := iProp_result Σ.
+  Definition iResUR (Σ : gFunctors) : ucmraT :=
+    iprodUR (λ i, gmapUR gname (Σ i (iPreProp Σ))).
+  Notation iProp Σ := (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.
+  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.
+(** * Properties of the solution to the recursive domain equation *)
+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/namespaces.v b/program_logic/namespaces.v
index 02e2a8dd6ec13c99ecde21d499b38c258b7fa19c..5badaaded9f02f7c05dec7d870b63a6f430e1a61 100644
--- a/program_logic/namespaces.v
+++ b/program_logic/namespaces.v
@@ -2,6 +2,10 @@ From iris.prelude Require Export countable coPset.
 From iris.algebra Require Export base.
 
 Definition namespace := list positive.
+Instance namespace_dec (N1 N2 : namespace) : Decision (N1 = N2) := _.
+Instance namespace_countable : Countable namespace := _.
+Typeclasses Opaque namespace.
+
 Definition nroot : namespace := nil.
 
 Definition ndot_def `{Countable A} (N : namespace) (x : A) : namespace :=
diff --git a/program_logic/ownership.v b/program_logic/ownership.v
index f7dfc25aad62c7a9991dbb34c527da90dde88a9a..f59c7af6fe158f124e9987a945d40f7d0ee7c92b 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..055b485a91cb39eb293d8ccb2421197c05317f01 100644
--- a/program_logic/pviewshifts.v
+++ b/program_logic/pviewshifts.v
@@ -1,290 +1,133 @@
-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) !==>".
+  iApply except_last_intro. iApply "H".
+  iIntros "[$ $] !==>". by iApply except_last_intro.
 Qed.
-Lemma pvs_timeless E P : TimelessP P → ▷ P ={E}=> P.
+Lemma except_last_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]". iTimeless "H". 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 [$ $]"; iVs "H".
+  iVsIntro. by iApply except_last_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)". 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.
+  iDestruct (ownE_op' with "[HE2 HEf]") as "[? $]"; first by iFrame.
+  iVsIntro; iApply except_last_intro. 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_except_last E1 E2 P : (|={E1,E2}=> â—‡ P) ={E1,E2}=> P.
+Proof. by rewrite {1}(pvs_intro E2 P) except_last_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..ecfe85db26e1a289c7ba5fdc941f23424ae3141a 100644
--- a/program_logic/saved_prop.v
+++ b/program_logic/saved_prop.v
@@ -1,33 +1,35 @@
-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 Λ Σ)))).
-Definition savedPropGF (F : cFunctor) : gFunctor :=
-  GFunctor (agreeRF (â–¶ F)).
-Instance inGF_savedPropG  `{inGF Λ Σ (savedPropGF F)} : savedPropG Λ Σ F.
-Proof. apply: inGF_inG. Qed.
+Class savedPropG (Σ : gFunctors) (F : cFunctor) :=
+  saved_prop_inG :> inG Σ (agreeR (laterC (F (iPreProp Σ)))).
+Definition savedPropΣ (F : cFunctor) : gFunctors :=
+  #[ GFunctor (agreeRF (â–¶ F)) ].
 
-Definition saved_prop_own `{savedPropG Λ Σ F}
-    (γ : gname) (x : F (iPropG Λ Σ)) : iPropG Λ Σ :=
+Instance subG_savedPropΣ  {Σ F} : subG (savedPropΣ F) Σ → savedPropG Σ F.
+Proof. apply subG_inG. Qed.
+
+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 :
@@ -35,8 +37,7 @@ Section saved_prop.
   Proof.
     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 7152d178f7a8b05edf7b179e188792cac0a67661..4693dc29d71f4f8da0476dcbdff5c6f20e6081bd 100644
--- a/program_logic/sts.v
+++ b/program_logic/sts.v
@@ -1,30 +1,29 @@
-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 pviewshifts.
+From iris.algebra Require Export sts.
+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.
-Proof. split; try apply _. apply: inGF_inG. Qed.
+Definition stsΣ (sts : stsT) : gFunctors := #[ GFunctor (constRF (stsR sts)) ].
+
+Instance subG_stsΣ Σ sts :
+  subG (stsΣ sts) Σ → Inhabited (sts.state sts) → stsG Σ sts.
+Proof. intros ?%subG_inG ?. by split. 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 :
@@ -47,30 +46,30 @@ Section definitions.
   Proof. apply _. Qed.
 End definitions.
 
-Typeclasses Opaque sts_own sts_ownS sts_ctx.
+Typeclasses Opaque sts_own sts_ownS sts_inv sts_ctx.
 Instance: Params (@sts_inv) 5.
 Instance: Params (@sts_ownS) 5.
 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 +78,62 @@ 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.
-    iNext. iExists s. by iFrame.
+    iVs (inv_alloc N _ (sts_inv γ φ) with "[Hφ Hγ]") as "#?"; auto.
+    rewrite /sts_inv. 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_accS E γ S T :
+    ▷ sts_inv γ φ ★ sts_ownS γ S T ={E}=> ∃ s,
+      ■ (s ∈ S) ★ ▷ φ s ★ ∀ s' T',
+      ■ sts.steps (s, T) (s', T') ★ ▷ φ s' ={E}=★ ▷ sts_inv γ φ ★ 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γ".
-    iCombine "Hγ" "Hγf" as "Hγ"; iDestruct (@own_valid with "#Hγ") as %Hvalid.
+    iIntros "[Hinv Hγf]". rewrite /sts_ownS /sts_inv /sts_own.
+    iDestruct "Hinv" as (s) "[>Hγ 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γ $]".
+    iVsIntro. 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_acc E γ s0 T :
+    ▷ sts_inv γ φ ★ sts_own γ s0 T ={E}=> ∃ s,
+      ■ sts.frame_steps T s0 s ★ ▷ φ s ★ ∀ s' T',
+      ■ sts.steps (s, T) (s', T') ★ ▷ φ s' ={E}=★ ▷ sts_inv γ φ ★ sts_own γ s' T'.
+  Proof. by apply sts_accS. Qed.
+    
+  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]". rewrite /sts_ctx. iInv N as "Hinv" "Hclose".
+    (* The following is essentially a very trivial composition of the accessors
+       [sts_acc] and [inv_open] -- but since we don't have any good support
+       for that currently, this gets more tedious than it should, with us having
+       to unpack and repack various proofs.
+       TODO: Make this mostly automatic, by supporting "opening accessors
+       around accessors". *)
+    iVs (sts_accS with "[Hinv Hγf]") as (s) "(?&?& HclSts)"; first by iFrame.
+    iVsIntro. iExists s. iFrame. iIntros (s' T') "H".
+    iVs ("HclSts" $! s' T' with "H") as "(Hinv & ?)". by iVs ("Hclose" with "Hinv").
+  Qed.
+
+  Lemma sts_open E N γ s0 T :
+    nclose N ⊆ E →
+    sts_ctx γ N φ ★ sts_own γ s0 T ={E,E∖N}=> ∃ s,
+      ■ (sts.frame_steps T s0 s) ★ ▷ φ 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/thread_local.v b/program_logic/thread_local.v
new file mode 100644
index 0000000000000000000000000000000000000000..452b897780b2b1ae34491b5dbe3bac642c7bdb61
--- /dev/null
+++ b/program_logic/thread_local.v
@@ -0,0 +1,92 @@
+From iris.algebra Require Export gmap gset coPset.
+From iris.proofmode Require Import invariants tactics.
+Import uPred.
+
+Definition tlN : namespace := nroot .@ "tl".
+Definition thread_id := gname.
+
+Class thread_localG Σ :=
+  tl_inG :> inG Σ (prodR coPset_disjR (gset_disjR positive)).
+
+Section defs.
+  Context `{irisG Λ Σ, thread_localG Σ}.
+
+  Definition tl_own (tid : thread_id) (E : coPset) : iProp Σ :=
+    own tid (CoPset E, ∅).
+
+  Definition tl_inv (tid : thread_id) (N : namespace) (P : iProp Σ) : iProp Σ :=
+    (∃ i, ■ (i ∈ nclose N) ∧
+          inv tlN (P ★ own tid (∅, GSet {[i]}) ∨ tl_own tid {[i]}))%I.
+End defs.
+
+Instance: Params (@tl_inv) 4.
+Typeclasses Opaque tl_own tl_inv.
+
+Section proofs.
+  Context `{irisG Λ Σ, thread_localG Σ}.
+
+  Global Instance tl_own_timeless tid E : TimelessP (tl_own tid E).
+  Proof. rewrite /tl_own; apply _. Qed.
+
+  Global Instance tl_inv_ne tid N n : Proper (dist n ==> dist n) (tl_inv tid N).
+  Proof. rewrite /tl_inv. solve_proper. Qed.
+  Global Instance tl_inv_proper tid N : Proper ((≡) ==> (≡)) (tl_inv tid N).
+  Proof. apply (ne_proper _). Qed.
+
+  Global Instance tl_inv_persistent tid N P : PersistentP (tl_inv tid N P).
+  Proof. rewrite /tl_inv; apply _. Qed.
+
+  Lemma tl_alloc : True =r=> ∃ tid, tl_own tid ⊤.
+  Proof. by apply own_alloc. Qed.
+
+  Lemma tl_own_disjoint tid E1 E2 : tl_own tid E1 ★ tl_own tid E2 ⊢ ■ (E1 ⊥ E2).
+  Proof.
+    rewrite /tl_own -own_op own_valid -coPset_disj_valid_op. by iIntros ([? _]).
+  Qed.
+
+  Lemma tl_own_union tid E1 E2 :
+    E1 ⊥ E2 → tl_own tid (E1 ∪ E2) ⊣⊢ tl_own tid E1 ★ tl_own tid E2.
+  Proof.
+    intros ?. by rewrite /tl_own -own_op pair_op left_id coPset_disj_union.
+  Qed.
+
+  Lemma tl_inv_alloc tid E N P : â–· P ={E}=> tl_inv tid N P.
+  Proof.
+    iIntros "HP".
+    iVs (own_empty (A:=prodUR coPset_disjUR (gset_disjUR positive)) tid) as "Hempty".
+    iVs (own_updateP with "Hempty") as ([m1 m2]) "[Hm Hown]".
+    { apply prod_updateP'. apply cmra_updateP_id, (reflexivity (R:=eq)).
+      apply (gset_alloc_empty_updateP_strong' (λ i, i ∈ nclose N)).
+      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. }
+    simpl. iDestruct "Hm" as %(<- & i & -> & ?).
+    rewrite /tl_inv.
+    iVs (inv_alloc tlN with "[-]"); last (iVsIntro; iExists i; eauto).
+    iNext. iLeft. by iFrame.
+  Qed.
+
+  Lemma tl_inv_open tid tlE E N P :
+    nclose tlN ⊆ tlE → nclose N ⊆ E →
+    tl_inv tid N P ⊢ tl_own tid E ={tlE}=★ ▷ P ★ tl_own tid (E ∖ N) ★
+                       (▷ P ★ tl_own tid (E ∖ N) ={tlE}=★ tl_own tid E).
+  Proof.
+    rewrite /tl_inv. iIntros (??) "#Htlinv Htoks".
+    iDestruct "Htlinv" as (i) "[% Hinv]".
+    rewrite {1 4}(union_difference_L (nclose N) E) //.
+    rewrite {1 5}(union_difference_L {[i]} (nclose N)) ?tl_own_union; [|set_solver..].
+    iDestruct "Htoks" as "[[Htoki $] $]".
+    iInv tlN as "[[$ >Hdis]|>Htoki2]" "Hclose".
+    - iVs ("Hclose" with "[Htoki]") as "_"; first auto.
+      iIntros "!==>[HP $]".
+      iInv tlN as "[[_ >Hdis2]|>Hitok]" "Hclose".
+      + iCombine "Hdis" "Hdis2" as "Hdis".
+        iDestruct (own_valid with "Hdis") as %[_ Hval].
+        revert Hval. rewrite gset_disj_valid_op. set_solver.
+      + iFrame. iApply "Hclose". iNext. iLeft. by iFrame.
+    - iDestruct (tl_own_disjoint tid {[i]} {[i]} with "[-]") as %?; first by iFrame.
+      set_solver.
+  Qed.
+End proofs.
diff --git a/program_logic/viewshifts.v b/program_logic/viewshifts.v
index b131eed655aada65465a81287a7533e48ce03cd7..02f910142868835e4e4fa17e90ca33621ceb3ea5 100644
--- a/program_logic/viewshifts.v
+++ b/program_logic/viewshifts.v
@@ -1,12 +1,10 @@
-From iris.program_logic Require Import ownership.
-From iris.program_logic Require Export pviewshifts invariants ghost_ownership.
+From iris.program_logic Require Export pviewshifts.
 From iris.proofmode Require Import pviewshifts invariants.
-Import uPred.
 
-Definition vs {Λ Σ} (E1 E2 : coPset) (P Q : iProp Λ Σ) : iProp Λ Σ :=
+Definition vs `{irisG Λ Σ} (E1 E2 : coPset) (P Q : iProp Σ) : iProp Σ :=
   (□ (P → |={E1,E2}=> Q))%I.
-Arguments vs {_ _} _ _ _%I _%I.
-Instance: Params (@vs) 4.
+Arguments vs {_ _ _} _ _ _%I _%I.
+Instance: Params (@vs) 5.
 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.
@@ -15,64 +13,59 @@ Notation "P ={ E }=> Q" := (P ={E,E}=> Q)%I
    format "P  ={ E }=>  Q") : uPred_scope.
 
 Section vs.
-Context {Λ : language} {Σ : iFunctor}.
-Implicit Types P Q R : iProp Λ Σ.
+Context `{irisG Λ Σ}.
+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).
+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).
+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).
+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.
+Proof. by iIntros (?) "> ?". 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.
+  (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".
+  iIntros "#[HvsP HvsQ] !# HP".
+  iVs ("HvsP" with "HP") as "HQ". 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.
+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.
+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.
+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.
+Lemma vs_mask_frame_r E1 E2 Ef P Q :
+  E1 ⊥ Ef → (P ={E1,E2}=> Q) ⊢ P ={E1 ∪ Ef,E2 ∪ Ef}=> Q.
 Proof.
-  iIntros (?) "#Hvs ! HP". iApply pvs_mask_frame; auto. by iApply "Hvs".
+  iIntros (?) "#Hvs !# HP". iApply pvs_mask_frame_r; 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".
+  iIntros (?) "#[? Hvs] !# HP". iInv N as "HR" "Hclose".
+  iVs ("Hvs" with "[HR HP]") as "[? $]"; first by iFrame.
+  by iApply "Hclose".
 Qed.
 
 Lemma vs_alloc N P : â–· P ={N}=> inv N P.
diff --git a/program_logic/weakestpre.v b/program_logic/weakestpre.v
index ab2d859b938a2cc26388a2422e7a5036c588c36a..b805f323962d873e45ae20f90e88dd1d509fb4a5 100644
--- a/program_logic/weakestpre.v
+++ b/program_logic/weakestpre.v
@@ -1,265 +1,179 @@
 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.algebra Require Import upred_big_op.
+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 efs, ■ prim_step e1 σ1 e2 σ2 efs ={∅,E}=★
+       ownP_auth σ2 ★ wp E e2 Φ ★
+       [★ list] ef ∈ efs, wp ⊤ ef (λ _, True)))%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=> efs.
+  apply wand_ne, pvs_ne, sep_ne, sep_ne; auto; first by apply Hwp.
+  apply big_sepL_ne=> ? ef. by apply Hwp.
 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 {_ _ _} _ _%E _.
+Instance: Params (@wp) 5.
 
-Notation "'WP' e @ E {{ Φ } }" := (wp E e Φ)
+Notation "'WP' e @ E {{ Φ } }" := (wp E e%E Φ)
   (at level 20, e, Φ at level 200,
    format "'WP'  e  @  E  {{  Φ  } }") : uPred_scope.
-Notation "'WP' e {{ Φ } }" := (wp ⊤ e Φ)
+Notation "'WP' e {{ Φ } }" := (wp ⊤ e%E Φ)
   (at level 20, e, Φ at level 200,
    format "'WP'  e  {{  Φ  } }") : uPred_scope.
 
-Notation "'WP' e @ E {{ v , Q } }" := (wp E e (λ v, Q))
+Notation "'WP' e @ E {{ v , Q } }" := (wp E e%E (λ v, Q))
   (at level 20, e, Q at level 200,
    format "'WP'  e  @  E  {{  v ,  Q  } }") : uPred_scope.
-Notation "'WP' e {{ v , Q } }" := (wp ⊤ e (λ v, Q))
+Notation "'WP' e {{ v , Q } }" := (wp ⊤ e%E (λ v, Q))
   (at level 20, e, Q at level 200,
    format "'WP'  e  {{  v ,  Q  } }") : uPred_scope.
 
 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.
+Lemma wp_value' E Φ v : Φ v ⊢ WP of_val v @ E {{ Φ }}.
 Proof.
-  case EQ: (to_val e).
-  - rewrite -(of_to_val _ _ EQ). constructor. rewrite pvs_eq.
-    exact: pvs_zero.
-  - constructor; first done. intros ?????. exfalso. omega.
+  iIntros "HΦ". rewrite wp_unfold /wp_pre.
+  iLeft; iExists v; rewrite to_of_val; auto.
 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_inv E Φ v : WP of_val v @ E {{ Φ }} ={E}=> Φ v.
 Proof.
-  by inversion 1 as [|??? He]; [|rewrite ?to_of_val in He]; simplify_eq.
+  rewrite wp_unfold /wp_pre to_of_val. iIntros "[H|[% _]]"; [|done].
+  by iDestruct "H" as (v') "[% ?]"; simplify_eq.
 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.
-Proof.
-  intros He; destruct 3; [by rewrite ?to_of_val in He|eauto].
+
+Lemma wp_strong_mono E1 E2 e Φ Ψ :
+  E1 ⊆ E2 → (∀ v, Φ v ={E2}=★ Ψ 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 ("HΦ" with "==>[-]"). by iApply (pvs_mask_mono E1 _). }
+  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 efs Hstep).
+  iVs ("H" $! _ σ2 efs with "[#]") as "($ & H & $)"; auto.
+  iVs "Hclose" as "_". by iApply ("IH" with "HΦ").
 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.
-Qed.
+Proof. iIntros "H". iApply (wp_strong_mono E); try iFrame; auto. 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 efs Hstep).
+  destruct (Hatomic _ _ _ _ Hstep) as [v <-%of_to_val].
+  iVs ("H" $! _ σ2 efs with "[#]") as "($ & H & $)"; auto.
+  iVs (wp_value_inv with "H") as "==> 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_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 efs Hstep).
+  iVs ("H" $! e2 σ2 efs 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 ?.
-  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.
+  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 efs Hstep).
+  destruct (fill_step_inv e σ1 e2 σ2 efs) as (e2'&->&?); auto.
+  iVs ("H" $! e2' σ2 efs 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) "?". by 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 +181,33 @@ 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.
+  iIntros "[H Hwp]". iApply (wp_strong_mono E); auto.
+  iFrame "Hwp". iIntros (?) "?". by iApply "H".
 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
new file mode 100644
index 0000000000000000000000000000000000000000..68ce0feef788c5e06f1ed49dede0cb957caa2cb4
--- /dev/null
+++ b/proofmode/class_instances.v
@@ -0,0 +1,358 @@
+From iris.proofmode Require Export classes.
+From iris.algebra Require Import upred_big_op gmap.
+Import uPred.
+
+Section classes.
+Context {M : ucmraT}.
+Implicit Types P Q R : uPred M.
+
+(* FromAssumption *)
+Global Instance from_assumption_exact p P : FromAssumption p P P.
+Proof. destruct p; by rewrite /FromAssumption /= ?always_elim. Qed.
+Global Instance from_assumption_always_l p P Q :
+  FromAssumption p P Q → FromAssumption p (□ P) Q.
+Proof. rewrite /FromAssumption=><-. by rewrite always_elim. Qed.
+Global Instance from_assumption_always_r P Q :
+  FromAssumption true P Q → FromAssumption true P (□ Q).
+Proof. rewrite /FromAssumption=><-. by rewrite always_always. Qed.
+Global Instance from_assumption_rvs p P Q :
+  FromAssumption p P Q → FromAssumption p P (|=r=> Q)%I.
+Proof. rewrite /FromAssumption=>->. apply rvs_intro. Qed.
+
+(* IntoPure *)
+Global Instance into_pure_pure φ : @IntoPure M (■ φ) φ.
+Proof. done. Qed.
+Global Instance into_pure_eq {A : cofeT} (a b : A) :
+  Timeless a → @IntoPure M (a ≡ b) (a ≡ b).
+Proof. intros. by rewrite /IntoPure timeless_eq. Qed.
+Global Instance into_pure_valid `{CMRADiscrete A} (a : A) : @IntoPure M (✓ a) (✓ a).
+Proof. by rewrite /IntoPure discrete_valid. Qed.
+
+(* FromPure *)
+Global Instance from_pure_pure φ : @FromPure M (■ φ) φ.
+Proof. intros ?. by apply pure_intro. Qed.
+Global Instance from_pure_eq {A : cofeT} (a b : A) : @FromPure M (a ≡ b) (a ≡ b).
+Proof. intros ->. apply eq_refl. Qed.
+Global Instance from_pure_valid {A : cmraT} (a : A) : @FromPure M (✓ a) (✓ a).
+Proof. intros ?. by apply valid_intro. Qed.
+Global Instance from_pure_rvs P φ : FromPure P φ → FromPure (|=r=> P) φ.
+Proof. intros ??. by rewrite -rvs_intro (from_pure P). Qed.
+
+(* IntoPersistentP *)
+Global Instance into_persistentP_always_trans P Q :
+  IntoPersistentP P Q → IntoPersistentP (□ P) Q | 0.
+Proof. rewrite /IntoPersistentP=> ->. by rewrite always_always. Qed.
+Global Instance into_persistentP_always P : IntoPersistentP (â–¡ P) P | 1.
+Proof. done. Qed.
+Global Instance into_persistentP_persistent P :
+  PersistentP P → IntoPersistentP P P | 100.
+Proof. done. Qed.
+
+(* IntoLater *)
+Global Instance into_later_default P : IntoLater P P | 1000.
+Proof. apply later_intro. Qed.
+Global Instance into_later_later P : IntoLater (â–· P) P.
+Proof. done. Qed.
+Global Instance into_later_and P1 P2 Q1 Q2 :
+  IntoLater P1 Q1 → IntoLater P2 Q2 → IntoLater (P1 ∧ P2) (Q1 ∧ Q2).
+Proof. intros ??; red. by rewrite later_and; apply and_mono. Qed.
+Global Instance into_later_or P1 P2 Q1 Q2 :
+  IntoLater P1 Q1 → IntoLater P2 Q2 → IntoLater (P1 ∨ P2) (Q1 ∨ Q2).
+Proof. intros ??; red. by rewrite later_or; apply or_mono. Qed.
+Global Instance into_later_sep P1 P2 Q1 Q2 :
+  IntoLater P1 Q1 → IntoLater P2 Q2 → IntoLater (P1 ★ P2) (Q1 ★ Q2).
+Proof. intros ??; red. by rewrite later_sep; apply sep_mono. Qed.
+
+Global Instance into_later_big_sepM `{Countable K} {A}
+    (Φ Ψ : K → A → uPred M) (m : gmap K A) :
+  (∀ x k, IntoLater (Φ k x) (Ψ k x)) →
+  IntoLater ([★ map] k ↦ x ∈ m, Φ k x) ([★ map] k ↦ x ∈ m, Ψ k x).
+Proof.
+  rewrite /IntoLater=> ?. rewrite big_sepM_later; by apply big_sepM_mono.
+Qed.
+Global Instance into_later_big_sepS `{Countable A}
+    (Φ Ψ : A → uPred M) (X : gset A) :
+  (∀ x, IntoLater (Φ x) (Ψ x)) →
+  IntoLater ([★ set] x ∈ X, Φ x) ([★ set] x ∈ X, Ψ x).
+Proof.
+  rewrite /IntoLater=> ?. rewrite big_sepS_later; by apply big_sepS_mono.
+Qed.
+
+(* FromLater *)
+Global Instance from_later_later P : FromLater (â–· P) P.
+Proof. done. Qed.
+Global Instance from_later_and P1 P2 Q1 Q2 :
+  FromLater P1 Q1 → FromLater P2 Q2 → FromLater (P1 ∧ P2) (Q1 ∧ Q2).
+Proof. intros ??; red. by rewrite later_and; apply and_mono. Qed.
+Global Instance from_later_or P1 P2 Q1 Q2 :
+  FromLater P1 Q1 → FromLater P2 Q2 → FromLater (P1 ∨ P2) (Q1 ∨ Q2).
+Proof. intros ??; red. by rewrite later_or; apply or_mono. Qed.
+Global Instance from_later_sep P1 P2 Q1 Q2 :
+  FromLater P1 Q1 → FromLater P2 Q2 → FromLater (P1 ★ P2) (Q1 ★ Q2).
+Proof. intros ??; red. by rewrite later_sep; apply sep_mono. Qed.
+
+(* IntoWand *)
+Global Instance into_wand_wand P Q Q' :
+  FromAssumption false Q Q' → IntoWand (P -★ Q) P Q'.
+Proof. by rewrite /FromAssumption /IntoWand /= => ->. Qed.
+Global Instance into_wand_impl P Q Q' :
+  FromAssumption false Q Q' → IntoWand (P → Q) P Q'.
+Proof. rewrite /FromAssumption /IntoWand /= => ->. by rewrite impl_wand. Qed.
+Global Instance into_wand_iff_l P Q : IntoWand (P ↔ Q) P Q.
+Proof. by apply and_elim_l', impl_wand. Qed.
+Global Instance into_wand_iff_r P Q : IntoWand (P ↔ Q) Q P.
+Proof. apply and_elim_r', impl_wand. Qed.
+Global Instance into_wand_always R P Q : IntoWand R P Q → IntoWand (□ R) P Q.
+Proof. rewrite /IntoWand=> ->. apply always_elim. Qed.
+Global Instance into_wand_rvs R P Q :
+  IntoWand R P Q → IntoWand R (|=r=> P) (|=r=> Q) | 100.
+Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite rvs_wand_r. Qed.
+
+(* FromAnd *)
+Global Instance from_and_and P1 P2 : FromAnd (P1 ∧ P2) P1 P2.
+Proof. done. Qed.
+Global Instance from_and_sep_persistent_l P1 P2 :
+  PersistentP P1 → FromAnd (P1 ★ P2) P1 P2 | 9.
+Proof. intros. by rewrite /FromAnd always_and_sep_l. Qed.
+Global Instance from_and_sep_persistent_r P1 P2 :
+  PersistentP P2 → FromAnd (P1 ★ P2) P1 P2 | 10.
+Proof. intros. by rewrite /FromAnd always_and_sep_r. Qed.
+Global Instance from_and_always P Q1 Q2 :
+  FromAnd P Q1 Q2 → FromAnd (□ P) (□ Q1) (□ Q2).
+Proof. rewrite /FromAnd=> <-. by rewrite always_and. Qed.
+Global Instance from_and_later P Q1 Q2 :
+  FromAnd P Q1 Q2 → FromAnd (▷ P) (▷ Q1) (▷ Q2).
+Proof. rewrite /FromAnd=> <-. by rewrite later_and. Qed.
+
+(* FromSep *)
+Global Instance from_sep_sep P1 P2 : FromSep (P1 ★ P2) P1 P2 | 100.
+Proof. done. Qed.
+Global Instance from_sep_always P Q1 Q2 :
+  FromSep P Q1 Q2 → FromSep (□ P) (□ Q1) (□ Q2).
+Proof. rewrite /FromSep=> <-. by rewrite always_sep. Qed.
+Global Instance from_sep_later P Q1 Q2 :
+  FromSep P Q1 Q2 → FromSep (▷ P) (▷ Q1) (▷ Q2).
+Proof. rewrite /FromSep=> <-. by rewrite later_sep. Qed.
+Global Instance from_sep_rvs P Q1 Q2 :
+  FromSep P Q1 Q2 → FromSep (|=r=> P) (|=r=> Q1) (|=r=> Q2).
+Proof. rewrite /FromSep=><-. apply rvs_sep. Qed.
+
+Global Instance from_sep_ownM (a b : M) :
+  FromSep (uPred_ownM (a â‹… b)) (uPred_ownM a) (uPred_ownM b) | 99.
+Proof. by rewrite /FromSep ownM_op. Qed.
+Global Instance from_sep_big_sepM
+    `{Countable K} {A} (Φ Ψ1 Ψ2 : K → A → uPred M) m :
+  (∀ k x, FromSep (Φ k x) (Ψ1 k x) (Ψ2 k x)) →
+  FromSep ([★ map] k ↦ x ∈ m, Φ k x)
+    ([★ map] k ↦ x ∈ m, Ψ1 k x) ([★ map] k ↦ x ∈ m, Ψ2 k x).
+Proof.
+  rewrite /FromSep=> ?. rewrite -big_sepM_sepM. by apply big_sepM_mono.
+Qed.
+Global Instance from_sep_big_sepS `{Countable A} (Φ Ψ1 Ψ2 : A → uPred M) X :
+  (∀ x, FromSep (Φ x) (Ψ1 x) (Ψ2 x)) →
+  FromSep ([★ set] x ∈ X, Φ x) ([★ set] x ∈ X, Ψ1 x) ([★ set] x ∈ X, Ψ2 x).
+Proof.
+  rewrite /FromSep=> ?. rewrite -big_sepS_sepS. by apply big_sepS_mono.
+Qed.
+
+(* IntoOp *)
+Global Instance into_op_op {A : cmraT} (a b : A) : IntoOp (a â‹… b) a b.
+Proof. by rewrite /IntoOp. Qed.
+Global Instance into_op_persistent {A : cmraT} (a : A) :
+  Persistent a → IntoOp a a a.
+Proof. intros; apply (persistent_dup a). Qed.
+Global Instance into_op_pair {A B : cmraT} (a b1 b2 : A) (a' b1' b2' : B) :
+  IntoOp a b1 b2 → IntoOp a' b1' b2' →
+  IntoOp (a,a') (b1,b1') (b2,b2').
+Proof. by constructor. Qed.
+Global Instance into_op_Some {A : cmraT} (a : A) b1 b2 :
+  IntoOp a b1 b2 → IntoOp (Some a) (Some b1) (Some b2).
+Proof. by constructor. Qed.
+
+(* IntoAnd *)
+Global Instance into_and_sep p P Q : IntoAnd p (P ★ Q) P Q.
+Proof. by apply mk_into_and_sep. Qed.
+Global Instance into_and_ownM p (a b1 b2 : M) :
+  IntoOp a b1 b2 →
+  IntoAnd p (uPred_ownM a) (uPred_ownM b1) (uPred_ownM b2).
+Proof. intros. apply mk_into_and_sep. by rewrite (into_op a) ownM_op. Qed.
+
+Global Instance into_and_and P Q : IntoAnd true (P ∧ Q) P Q.
+Proof. done. Qed.
+Global Instance into_and_and_persistent_l P Q :
+  PersistentP P → IntoAnd false (P ∧ Q) P Q.
+Proof. intros; by rewrite /IntoAnd /= always_and_sep_l. Qed.
+Global Instance into_and_and_persistent_r P Q :
+  PersistentP Q → IntoAnd false (P ∧ Q) P Q.
+Proof. intros; by rewrite /IntoAnd /= always_and_sep_r. Qed.
+
+Global Instance into_and_later p P Q1 Q2 :
+  IntoAnd p P Q1 Q2 → IntoAnd p (▷ P) (▷ Q1) (▷ Q2).
+Proof. rewrite /IntoAnd=>->. destruct p; by rewrite ?later_and ?later_sep. Qed.
+
+Global Instance into_and_big_sepM
+    `{Countable K} {A} (Φ Ψ1 Ψ2 : K → A → uPred M) p m :
+  (∀ k x, IntoAnd p (Φ k x) (Ψ1 k x) (Ψ2 k x)) →
+  IntoAnd p ([★ map] k ↦ x ∈ m, Φ k x)
+    ([★ map] k ↦ x ∈ m, Ψ1 k x) ([★ map] k ↦ x ∈ m, Ψ2 k x).
+Proof.
+  rewrite /IntoAnd=> HΦ. destruct p.
+  - apply and_intro; apply big_sepM_mono; auto.
+    + intros k x ?. by rewrite HΦ and_elim_l.
+    + intros k x ?. by rewrite HΦ and_elim_r.
+  - rewrite -big_sepM_sepM. apply big_sepM_mono; auto.
+Qed.
+Global Instance into_and_big_sepS `{Countable A} (Φ Ψ1 Ψ2 : A → uPred M) p X :
+  (∀ x, IntoAnd p (Φ x) (Ψ1 x) (Ψ2 x)) →
+  IntoAnd p ([★ set] x ∈ X, Φ x) ([★ set] x ∈ X, Ψ1 x) ([★ set] x ∈ X, Ψ2 x).
+Proof.
+  rewrite /IntoAnd=> HΦ. destruct p.
+  - apply and_intro; apply big_sepS_mono; auto.
+    + intros x ?. by rewrite HΦ and_elim_l.
+    + intros x ?. by rewrite HΦ and_elim_r.
+  - rewrite -big_sepS_sepS. apply big_sepS_mono; auto.
+Qed.
+
+(* Frame *)
+Global Instance frame_here R : Frame R R True.
+Proof. by rewrite /Frame right_id. Qed.
+
+Class MakeSep (P Q PQ : uPred M) := make_sep : P ★ Q ⊣⊢ PQ.
+Global Instance make_sep_true_l P : MakeSep True P P.
+Proof. by rewrite /MakeSep left_id. Qed.
+Global Instance make_sep_true_r P : MakeSep P True P.
+Proof. by rewrite /MakeSep right_id. Qed.
+Global Instance make_sep_default P Q : MakeSep P Q (P ★ Q) | 100.
+Proof. done. Qed.
+Global Instance frame_sep_l R P1 P2 Q Q' :
+  Frame R P1 Q → MakeSep Q P2 Q' → Frame R (P1 ★ P2) Q' | 9.
+Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc. Qed.
+Global Instance frame_sep_r R P1 P2 Q Q' :
+  Frame R P2 Q → MakeSep P1 Q Q' → Frame R (P1 ★ P2) Q' | 10.
+Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc (comm _ R) assoc. Qed.
+
+Class MakeAnd (P Q PQ : uPred M) := make_and : P ∧ Q ⊣⊢ PQ.
+Global Instance make_and_true_l P : MakeAnd True P P.
+Proof. by rewrite /MakeAnd left_id. Qed.
+Global Instance make_and_true_r P : MakeAnd P True P.
+Proof. by rewrite /MakeAnd right_id. Qed.
+Global Instance make_and_default P Q : MakeAnd P Q (P ∧ Q) | 100.
+Proof. done. Qed.
+Global Instance frame_and_l R P1 P2 Q Q' :
+  Frame R P1 Q → MakeAnd Q P2 Q' → Frame R (P1 ∧ P2) Q' | 9.
+Proof. rewrite /Frame /MakeAnd => <- <-; eauto 10 with I. Qed.
+Global Instance frame_and_r R P1 P2 Q Q' :
+  Frame R P2 Q → MakeAnd P1 Q Q' → Frame R (P1 ∧ P2) Q' | 10.
+Proof. rewrite /Frame /MakeAnd => <- <-; eauto 10 with I. Qed.
+
+Class MakeOr (P Q PQ : uPred M) := make_or : P ∨ Q ⊣⊢ PQ.
+Global Instance make_or_true_l P : MakeOr True P True.
+Proof. by rewrite /MakeOr left_absorb. Qed.
+Global Instance make_or_true_r P : MakeOr P True True.
+Proof. by rewrite /MakeOr right_absorb. Qed.
+Global Instance make_or_default P Q : MakeOr P Q (P ∨ Q) | 100.
+Proof. done. Qed.
+Global Instance frame_or R P1 P2 Q1 Q2 Q :
+  Frame R P1 Q1 → Frame R P2 Q2 → MakeOr Q1 Q2 Q → Frame R (P1 ∨ P2) Q.
+Proof. rewrite /Frame /MakeOr => <- <- <-. by rewrite -sep_or_l. Qed.
+
+Global Instance frame_wand R P1 P2 Q2 :
+  Frame R P2 Q2 → Frame R (P1 -★ P2) (P1 -★ Q2).
+Proof.
+  rewrite /Frame=> ?. apply wand_intro_l.
+  by rewrite assoc (comm _ P1) -assoc wand_elim_r.
+Qed.
+
+Class MakeLater (P lP : uPred M) := make_later : ▷ P ⊣⊢ lP.
+Global Instance make_later_true : MakeLater True True.
+Proof. by rewrite /MakeLater later_True. Qed.
+Global Instance make_later_default P : MakeLater P (â–· P) | 100.
+Proof. done. Qed.
+
+Global Instance frame_later R P Q Q' :
+  Frame R P Q → MakeLater Q Q' → Frame R (▷ P) Q'.
+Proof.
+  rewrite /Frame /MakeLater=><- <-. by rewrite later_sep -(later_intro R).
+Qed.
+
+Class MakeExceptLast (P Q : uPred M) := make_except_last : ◇ P ⊣⊢ Q.
+Global Instance make_except_last_True : MakeExceptLast True True.
+Proof. by rewrite /MakeExceptLast except_last_True. Qed.
+Global Instance make_except_last_default P : MakeExceptLast P (â—‡ P) | 100.
+Proof. done. Qed.
+
+Global Instance frame_except_last R P Q Q' :
+  Frame R P Q → MakeExceptLast Q Q' → Frame R (◇ P) Q'.
+Proof.
+  rewrite /Frame /MakeExceptLast=><- <-.
+  by rewrite except_last_sep -(except_last_intro R).
+Qed.
+
+Global Instance frame_exist {A} R (Φ Ψ : A → uPred M) :
+  (∀ a, Frame R (Φ a) (Ψ a)) → Frame R (∃ x, Φ x) (∃ x, Ψ x).
+Proof. rewrite /Frame=> ?. by rewrite sep_exist_l; apply exist_mono. Qed.
+Global Instance frame_forall {A} R (Φ Ψ : A → uPred M) :
+  (∀ a, Frame R (Φ a) (Ψ a)) → Frame R (∀ x, Φ x) (∀ x, Ψ x).
+Proof. rewrite /Frame=> ?. by rewrite sep_forall_l; apply forall_mono. Qed.
+
+Global Instance frame_rvs R P Q : Frame R P Q → Frame R (|=r=> P) (|=r=> Q).
+Proof. rewrite /Frame=><-. by rewrite rvs_frame_l. Qed.
+
+(* FromOr *)
+Global Instance from_or_or P1 P2 : FromOr (P1 ∨ P2) P1 P2.
+Proof. done. Qed.
+Global Instance from_or_rvs P Q1 Q2 :
+  FromOr P Q1 Q2 → FromOr (|=r=> P) (|=r=> Q1) (|=r=> Q2).
+Proof. rewrite /FromOr=><-. apply or_elim; apply rvs_mono; auto with I. Qed.
+
+(* IntoOr *)
+Global Instance into_or_or P Q : IntoOr (P ∨ Q) P Q.
+Proof. done. Qed.
+Global Instance into_or_later P Q1 Q2 :
+  IntoOr P Q1 Q2 → IntoOr (▷ P) (▷ Q1) (▷ Q2).
+Proof. rewrite /IntoOr=>->. by rewrite later_or. Qed.
+
+(* FromExist *)
+Global Instance from_exist_exist {A} (Φ: A → uPred M): FromExist (∃ a, Φ a) Φ.
+Proof. done. Qed.
+Global Instance from_exist_rvs {A} P (Φ : A → uPred M) :
+  FromExist P Φ → FromExist (|=r=> P) (λ a, |=r=> Φ a)%I.
+Proof.
+  rewrite /FromExist=><-. apply exist_elim=> a. by rewrite -(exist_intro a).
+Qed.
+
+(* IntoExist *)
+Global Instance into_exist_exist {A} (Φ : A → uPred M) : IntoExist (∃ a, Φ a) Φ.
+Proof. done. Qed.
+Global Instance into_exist_later {A} P (Φ : A → uPred M) :
+  IntoExist P Φ → Inhabited A → IntoExist (▷ P) (λ a, ▷ (Φ a))%I.
+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.
+
+(* IntoExceptLast *)
+Global Instance into_except_last_except_last P : IntoExceptLast (â—‡ P) P.
+Proof. done. Qed.
+Global Instance into_except_last_timeless P : TimelessP P → IntoExceptLast (▷ P) P.
+Proof. done. Qed.
+
+(* IsExceptLast *)
+Global Instance is_except_last_except_last P : IsExceptLast (â—‡ P).
+Proof. by rewrite /IsExceptLast except_last_idemp. Qed.
+Global Instance is_except_last_later P : IsExceptLast (â–· P).
+Proof. by rewrite /IsExceptLast except_last_later. Qed.
+Global Instance is_except_last_rvs P : IsExceptLast P → IsExceptLast (|=r=> P).
+Proof.
+  rewrite /IsExceptLast=> HP.
+  by rewrite -{2}HP -(except_last_idemp P) -except_last_rvs -(except_last_intro P).
+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 50fdd54b659a0806122fbc5ef6221b167a0e6fb1..83c1f2e7c22378f98eca6118b94e7bb04c166042 100644
--- a/proofmode/classes.v
+++ b/proofmode/classes.v
@@ -1,5 +1,4 @@
 From iris.algebra Require Export upred.
-From iris.algebra Require Import upred_big_op gmap upred_tactics.
 Import uPred.
 
 Section classes.
@@ -8,311 +7,68 @@ Implicit Types P Q : uPred M.
 
 Class FromAssumption (p : bool) (P Q : uPred M) := from_assumption : □?p P ⊢ Q.
 Global Arguments from_assumption _ _ _ {_}.
-Global Instance from_assumption_exact p P : FromAssumption p P P.
-Proof. destruct p; by rewrite /FromAssumption /= ?always_elim. Qed.
-Global Instance from_assumption_always_l p P Q :
-  FromAssumption p P Q → FromAssumption p (□ P) Q.
-Proof. rewrite /FromAssumption=><-. by rewrite always_elim. Qed.
-Global Instance from_assumption_always_r P Q :
-  FromAssumption true P Q → FromAssumption true P (□ Q).
-Proof. rewrite /FromAssumption=><-. by rewrite always_always. Qed.
 
 Class IntoPure (P : uPred M) (φ : Prop) := into_pure : P ⊢ ■ φ.
 Global Arguments into_pure : clear implicits.
-Global Instance into_pure_pure φ : IntoPure (■ φ) φ.
-Proof. done. Qed.
-Global Instance into_pure_eq {A : cofeT} (a b : A) :
-  Timeless a → IntoPure (a ≡ b) (a ≡ b).
-Proof. intros. by rewrite /IntoPure timeless_eq. Qed.
-Global Instance into_pure_valid `{CMRADiscrete A} (a : A) : IntoPure (✓ a) (✓ a).
-Proof. by rewrite /IntoPure discrete_valid. Qed.
 
 Class FromPure (P : uPred M) (φ : Prop) := from_pure : φ → True ⊢ P.
 Global Arguments from_pure : clear implicits.
-Global Instance from_pure_pure φ : FromPure (■ φ) φ.
-Proof. intros ?. by apply pure_intro. Qed.
-Global Instance from_pure_eq {A : cofeT} (a b : A) : FromPure (a ≡ b) (a ≡ b).
-Proof. intros ->. apply eq_refl. Qed.
-Global Instance from_pure_valid {A : cmraT} (a : A) : FromPure (✓ a) (✓ a).
-Proof. intros ?. by apply valid_intro. Qed.
 
 Class IntoPersistentP (P Q : uPred M) := into_persistentP : P ⊢ □ Q.
 Global Arguments into_persistentP : clear implicits.
-Global Instance into_persistentP_always_trans P Q :
-  IntoPersistentP P Q → IntoPersistentP (□ P) Q | 0.
-Proof. rewrite /IntoPersistentP=> ->. by rewrite always_always. Qed.
-Global Instance into_persistentP_always P : IntoPersistentP (â–¡ P) P | 1.
-Proof. done. Qed.
-Global Instance into_persistentP_persistent P :
-  PersistentP P → IntoPersistentP P P | 100.
-Proof. done. Qed.
 
 Class IntoLater (P Q : uPred M) := into_later : P ⊢ ▷ Q.
 Global Arguments into_later _ _ {_}.
+
 Class FromLater (P Q : uPred M) := from_later : ▷ Q ⊢ P.
 Global Arguments from_later _ _ {_}.
 
-Global Instance into_later_default P : IntoLater P P | 1000.
-Proof. apply later_intro. Qed.
-Global Instance into_later_later P : IntoLater (â–· P) P.
-Proof. done. Qed.
-Global Instance into_later_and P1 P2 Q1 Q2 :
-  IntoLater P1 Q1 → IntoLater P2 Q2 → IntoLater (P1 ∧ P2) (Q1 ∧ Q2).
-Proof. intros ??; red. by rewrite later_and; apply and_mono. Qed.
-Global Instance into_later_or P1 P2 Q1 Q2 :
-  IntoLater P1 Q1 → IntoLater P2 Q2 → IntoLater (P1 ∨ P2) (Q1 ∨ Q2).
-Proof. intros ??; red. by rewrite later_or; apply or_mono. Qed.
-Global Instance into_later_sep P1 P2 Q1 Q2 :
-  IntoLater P1 Q1 → IntoLater P2 Q2 → IntoLater (P1 ★ P2) (Q1 ★ Q2).
-Proof. intros ??; red. by rewrite later_sep; apply sep_mono. Qed.
-
-Global Instance into_later_big_sepM `{Countable K} {A}
-    (Φ Ψ : K → A → uPred M) (m : gmap K A) :
-  (∀ x k, IntoLater (Φ k x) (Ψ k x)) →
-  IntoLater ([★ map] k ↦ x ∈ m, Φ k x) ([★ map] k ↦ x ∈ m, Ψ k x).
-Proof.
-  rewrite /IntoLater=> ?. rewrite big_sepM_later; by apply big_sepM_mono.
-Qed.
-Global Instance into_later_big_sepS `{Countable A}
-    (Φ Ψ : A → uPred M) (X : gset A) :
-  (∀ x, IntoLater (Φ x) (Ψ x)) →
-  IntoLater ([★ set] x ∈ X, Φ x) ([★ set] x ∈ X, Ψ x).
-Proof.
-  rewrite /IntoLater=> ?. rewrite big_sepS_later; by apply big_sepS_mono.
-Qed.
-
-Global Instance from_later_later P : FromLater (â–· P) P.
-Proof. done. Qed.
-Global Instance from_later_and P1 P2 Q1 Q2 :
-  FromLater P1 Q1 → FromLater P2 Q2 → FromLater (P1 ∧ P2) (Q1 ∧ Q2).
-Proof. intros ??; red. by rewrite later_and; apply and_mono. Qed.
-Global Instance from_later_or P1 P2 Q1 Q2 :
-  FromLater P1 Q1 → FromLater P2 Q2 → FromLater (P1 ∨ P2) (Q1 ∨ Q2).
-Proof. intros ??; red. by rewrite later_or; apply or_mono. Qed.
-Global Instance from_later_sep P1 P2 Q1 Q2 :
-  FromLater P1 Q1 → FromLater P2 Q2 → FromLater (P1 ★ P2) (Q1 ★ Q2).
-Proof. intros ??; red. by rewrite later_sep; apply sep_mono. Qed.
-
 Class IntoWand (R P Q : uPred M) := into_wand : R ⊢ P -★ Q.
 Global Arguments into_wand : clear implicits.
-Global Instance into_wand_wand P Q : IntoWand (P -★ Q) P Q.
-Proof. done. Qed.
-Global Instance into_wand_impl P Q : IntoWand (P → Q) P Q.
-Proof. apply impl_wand. Qed.
-Global Instance into_wand_iff_l P Q : IntoWand (P ↔ Q) P Q.
-Proof. by apply and_elim_l', impl_wand. Qed.
-Global Instance into_wand_iff_r P Q : IntoWand (P ↔ Q) Q P.
-Proof. apply and_elim_r', impl_wand. Qed.
-Global Instance into_wand_always R P Q : IntoWand R P Q → IntoWand (□ R) P Q.
-Proof. rewrite /IntoWand=> ->. apply always_elim. Qed.
 
 Class FromAnd (P Q1 Q2 : uPred M) := from_and : Q1 ∧ Q2 ⊢ P.
 Global Arguments from_and : clear implicits.
 
-Global Instance from_and_and P1 P2 : FromAnd (P1 ∧ P2) P1 P2.
-Proof. done. Qed.
-Global Instance from_and_sep_persistent_l P1 P2 :
-  PersistentP P1 → FromAnd (P1 ★ P2) P1 P2 | 9.
-Proof. intros. by rewrite /FromAnd always_and_sep_l. Qed.
-Global Instance from_and_sep_persistent_r P1 P2 :
-  PersistentP P2 → FromAnd (P1 ★ P2) P1 P2 | 10.
-Proof. intros. by rewrite /FromAnd always_and_sep_r. Qed.
-Global Instance from_and_always P Q1 Q2 :
-  FromAnd P Q1 Q2 → FromAnd (□ P) (□ Q1) (□ Q2).
-Proof. rewrite /FromAnd=> <-. by rewrite always_and. Qed.
-Global Instance from_and_later P Q1 Q2 :
-  FromAnd P Q1 Q2 → FromAnd (▷ P) (▷ Q1) (▷ Q2).
-Proof. rewrite /FromAnd=> <-. by rewrite later_and. Qed.
-
 Class FromSep (P Q1 Q2 : uPred M) := from_sep : Q1 ★ Q2 ⊢ P.
 Global Arguments from_sep : clear implicits.
 
-Global Instance from_sep_sep P1 P2 : FromSep (P1 ★ P2) P1 P2 | 100.
-Proof. done. Qed.
-Global Instance from_sep_always P Q1 Q2 :
-  FromSep P Q1 Q2 → FromSep (□ P) (□ Q1) (□ Q2).
-Proof. rewrite /FromSep=> <-. by rewrite always_sep. Qed.
-Global Instance from_sep_later P Q1 Q2 :
-  FromSep P Q1 Q2 → FromSep (▷ P) (▷ Q1) (▷ Q2).
-Proof. rewrite /FromSep=> <-. by rewrite later_sep. Qed.
+Class IntoAnd (p : bool) (P Q1 Q2 : uPred M) :=
+  into_and : P ⊢ if p then Q1 ∧ Q2 else Q1 ★ Q2.
+Global Arguments into_and : clear implicits.
 
-Global Instance from_sep_ownM (a b : M) :
-  FromSep (uPred_ownM (a â‹… b)) (uPred_ownM a) (uPred_ownM b) | 99.
-Proof. by rewrite /FromSep ownM_op. Qed.
-Global Instance from_sep_big_sepM
-    `{Countable K} {A} (Φ Ψ1 Ψ2 : K → A → uPred M) m :
-  (∀ k x, FromSep (Φ k x) (Ψ1 k x) (Ψ2 k x)) →
-  FromSep ([★ map] k ↦ x ∈ m, Φ k x)
-    ([★ map] k ↦ x ∈ m, Ψ1 k x) ([★ map] k ↦ x ∈ m, Ψ2 k x).
-Proof.
-  rewrite /FromSep=> ?. rewrite -big_sepM_sepM. by apply big_sepM_mono.
-Qed.
-Global Instance from_sep_big_sepS `{Countable A} (Φ Ψ1 Ψ2 : A → uPred M) X :
-  (∀ x, FromSep (Φ x) (Ψ1 x) (Ψ2 x)) →
-  FromSep ([★ set] x ∈ X, Φ x) ([★ set] x ∈ X, Ψ1 x) ([★ set] x ∈ X, Ψ2 x).
-Proof.
-  rewrite /FromSep=> ?. rewrite -big_sepS_sepS. by apply big_sepS_mono.
-Qed.
+Lemma mk_into_and_sep p P Q1 Q2 : (P ⊢ Q1 ★ Q2) → IntoAnd p P Q1 Q2.
+Proof. rewrite /IntoAnd=>->. destruct p; auto using sep_and. Qed.
 
-Class IntoSep (p: bool) (P Q1 Q2 : uPred M) := into_sep : □?p P ⊢ □?p (Q1 ★ Q2).
-Global Arguments into_sep : clear implicits.
 Class IntoOp {A : cmraT} (a b1 b2 : A) := into_op : a ≡ b1 ⋅ b2.
 Global Arguments into_op {_} _ _ _ {_}.
 
-Global Instance into_op_op {A : cmraT} (a b : A) : IntoOp (a â‹… b) a b.
-Proof. by rewrite /IntoOp. Qed.
-Global Instance into_op_persistent {A : cmraT} (a : A) :
-  Persistent a → IntoOp a a a.
-Proof. intros; apply (persistent_dup a). Qed.
-Global Instance into_op_pair {A B : cmraT} (a b1 b2 : A) (a' b1' b2' : B) :
-  IntoOp a b1 b2 → IntoOp a' b1' b2' →
-  IntoOp (a,a') (b1,b1') (b2,b2').
-Proof. by constructor. Qed.
-Global Instance into_op_Some {A : cmraT} (a : A) b1 b2 :
-  IntoOp a b1 b2 → IntoOp (Some a) (Some b1) (Some b2).
-Proof. by constructor. Qed.
-
-Global Instance into_sep_sep p P Q : IntoSep p (P ★ Q) P Q.
-Proof. rewrite /IntoSep. by rewrite always_if_sep. Qed.
-Global Instance into_sep_ownM p (a b1 b2 : M) :
-  IntoOp a b1 b2 →
-  IntoSep p (uPred_ownM a) (uPred_ownM b1) (uPred_ownM b2).
-Proof.
-  rewrite /IntoOp /IntoSep=> ->. by rewrite ownM_op always_if_sep.
-Qed.
-
-Global Instance into_sep_and P Q : IntoSep true (P ∧ Q) P Q.
-Proof. by rewrite /IntoSep /= always_and_sep. Qed.
-Global Instance into_sep_and_persistent_l P Q :
-  PersistentP P → IntoSep false (P ∧ Q) P Q.
-Proof. intros; by rewrite /IntoSep /= always_and_sep_l. Qed.
-Global Instance into_sep_and_persistent_r P Q :
-  PersistentP Q → IntoSep false (P ∧ Q) P Q.
-Proof. intros; by rewrite /IntoSep /= always_and_sep_r. Qed.
-
-Global Instance into_sep_later p P Q1 Q2 :
-  IntoSep p P Q1 Q2 → IntoSep p (▷ P) (▷ Q1) (▷ Q2).
-Proof. by rewrite /IntoSep -later_sep !always_if_later=> ->. Qed.
-
-Global Instance into_sep_big_sepM
-    `{Countable K} {A} (Φ Ψ1 Ψ2 : K → A → uPred M) p m :
-  (∀ k x, IntoSep p (Φ k x) (Ψ1 k x) (Ψ2 k x)) →
-  IntoSep p ([★ map] k ↦ x ∈ m, Φ k x)
-    ([★ map] k ↦ x ∈ m, Ψ1 k x) ([★ map] k ↦ x ∈ m, Ψ2 k x).
-Proof.
-  rewrite /IntoSep=> ?. rewrite -big_sepM_sepM !big_sepM_always_if.
-  by apply big_sepM_mono.
-Qed.
-Global Instance into_sep_big_sepS `{Countable A} (Φ Ψ1 Ψ2 : A → uPred M) p X :
-  (∀ x, IntoSep p (Φ x) (Ψ1 x) (Ψ2 x)) →
-  IntoSep p ([★ set] x ∈ X, Φ x) ([★ set] x ∈ X, Ψ1 x) ([★ set] x ∈ X, Ψ2 x).
-Proof.
-  rewrite /IntoSep=> ?. rewrite -big_sepS_sepS !big_sepS_always_if.
-  by apply big_sepS_mono.
-Qed.
-
 Class Frame (R P Q : uPred M) := frame : R ★ Q ⊢ P.
 Global Arguments frame : clear implicits.
 
-Global Instance frame_here R : Frame R R True.
-Proof. by rewrite /Frame right_id. Qed.
-
-Class MakeSep (P Q PQ : uPred M) := make_sep : P ★ Q ⊣⊢ PQ.
-Global Instance make_sep_true_l P : MakeSep True P P.
-Proof. by rewrite /MakeSep left_id. Qed.
-Global Instance make_sep_true_r P : MakeSep P True P.
-Proof. by rewrite /MakeSep right_id. Qed.
-Global Instance make_sep_default P Q : MakeSep P Q (P ★ Q) | 100.
-Proof. done. Qed.
-Global Instance frame_sep_l R P1 P2 Q Q' :
-  Frame R P1 Q → MakeSep Q P2 Q' → Frame R (P1 ★ P2) Q' | 9.
-Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc. Qed.
-Global Instance frame_sep_r R P1 P2 Q Q' :
-  Frame R P2 Q → MakeSep P1 Q Q' → Frame R (P1 ★ P2) Q' | 10.
-Proof. rewrite /Frame /MakeSep => <- <-. solve_sep_entails. Qed.
-
-Class MakeAnd (P Q PQ : uPred M) := make_and : P ∧ Q ⊣⊢ PQ.
-Global Instance make_and_true_l P : MakeAnd True P P.
-Proof. by rewrite /MakeAnd left_id. Qed.
-Global Instance make_and_true_r P : MakeAnd P True P.
-Proof. by rewrite /MakeAnd right_id. Qed.
-Global Instance make_and_default P Q : MakeSep P Q (P ★ Q) | 100.
-Proof. done. Qed.
-Global Instance frame_and_l R P1 P2 Q Q' :
-  Frame R P1 Q → MakeAnd Q P2 Q' → Frame R (P1 ∧ P2) Q' | 9.
-Proof. rewrite /Frame /MakeAnd => <- <-; eauto 10 with I. Qed.
-Global Instance frame_and_r R P1 P2 Q Q' :
-  Frame R P2 Q → MakeAnd P1 Q Q' → Frame R (P1 ∧ P2) Q' | 10.
-Proof. rewrite /Frame /MakeAnd => <- <-; eauto 10 with I. Qed.
-
-Class MakeOr (P Q PQ : uPred M) := make_or : P ∨ Q ⊣⊢ PQ.
-Global Instance make_or_true_l P : MakeOr True P True.
-Proof. by rewrite /MakeOr left_absorb. Qed.
-Global Instance make_or_true_r P : MakeOr P True True.
-Proof. by rewrite /MakeOr right_absorb. Qed.
-Global Instance make_or_default P Q : MakeOr P Q (P ∨ Q) | 100.
-Proof. done. Qed.
-Global Instance frame_or R P1 P2 Q1 Q2 Q :
-  Frame R P1 Q1 → Frame R P2 Q2 → MakeOr Q1 Q2 Q → Frame R (P1 ∨ P2) Q.
-Proof. rewrite /Frame /MakeOr => <- <- <-. by rewrite -sep_or_l. Qed.
-
-Global Instance frame_wand R P1 P2 Q2 :
-  Frame R P2 Q2 → Frame R (P1 -★ P2) (P1 -★ Q2).
-Proof.
-  rewrite /Frame=> ?. apply wand_intro_l.
-  by rewrite assoc (comm _ P1) -assoc wand_elim_r.
-Qed.
-
-Class MakeLater (P lP : uPred M) := make_later : ▷ P ⊣⊢ lP.
-Global Instance make_later_true : MakeLater True True.
-Proof. by rewrite /MakeLater later_True. Qed.
-Global Instance make_later_default P : MakeLater P (â–· P) | 100.
-Proof. done. Qed.
-
-Global Instance frame_later R P Q Q' :
-  Frame R P Q → MakeLater Q Q' → Frame R (▷ P) Q'.
-Proof.
-  rewrite /Frame /MakeLater=><- <-. by rewrite later_sep -(later_intro R).
-Qed.
-
-Global Instance frame_exist {A} R (Φ Ψ : A → uPred M) :
-  (∀ a, Frame R (Φ a) (Ψ a)) → Frame R (∃ x, Φ x) (∃ x, Ψ x).
-Proof. rewrite /Frame=> ?. by rewrite sep_exist_l; apply exist_mono. Qed.
-Global Instance frame_forall {A} R (Φ Ψ : A → uPred M) :
-  (∀ a, Frame R (Φ a) (Ψ a)) → Frame R (∀ x, Φ x) (∀ x, Ψ x).
-Proof. rewrite /Frame=> ?. by rewrite sep_forall_l; apply forall_mono. Qed.
-
 Class FromOr (P Q1 Q2 : uPred M) := from_or : Q1 ∨ Q2 ⊢ P.
 Global Arguments from_or : clear implicits.
-Global Instance from_or_or P1 P2 : FromOr (P1 ∨ P2) P1 P2.
-Proof. done. Qed.
 
 Class IntoOr P Q1 Q2 := into_or : P ⊢ Q1 ∨ Q2.
 Global Arguments into_or : clear implicits.
-Global Instance into_or_or P Q : IntoOr (P ∨ Q) P Q.
-Proof. done. Qed.
-Global Instance into_or_later P Q1 Q2 :
-  IntoOr P Q1 Q2 → IntoOr (▷ P) (▷ Q1) (▷ Q2).
-Proof. rewrite /IntoOr=>->. by rewrite later_or. Qed.
 
 Class FromExist {A} (P : uPred M) (Φ : A → uPred M) :=
   from_exist : (∃ x, Φ x) ⊢ P.
 Global Arguments from_exist {_} _ _ {_}.
-Global Instance from_exist_exist {A} (Φ: A → uPred M): FromExist (∃ a, Φ a) Φ.
-Proof. done. Qed.
 
 Class IntoExist {A} (P : uPred M) (Φ : A → uPred M) :=
   into_exist : P ⊢ ∃ x, Φ x.
 Global Arguments into_exist {_} _ _ {_}.
-Global Instance into_exist_exist {A} (Φ : A → uPred M) : IntoExist (∃ a, Φ a) Φ.
-Proof. done. Qed.
-Global Instance into_exist_later {A} P (Φ : A → uPred M) :
-  IntoExist P Φ → Inhabited A → IntoExist (▷ P) (λ a, ▷ (Φ a))%I.
-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.
 
-Class TimelessElim (Q : uPred M) :=
-  timeless_elim `{!TimelessP P} : ▷ P ★ (P -★ Q) ⊢ Q.
+Class IntoExceptLast (P Q : uPred M) := into_except_last : P ⊢ ◇ Q.
+Global Arguments into_except_last : clear implicits.
+
+Class IsExceptLast (Q : uPred M) := is_except_last : ◇ Q ⊢ Q.
+Global Arguments is_except_last : 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 36f52d72a7ec20f72712fb6ad7cd31d7a0b472fe..56699b25bd7b880ac5acea685103f4a0535bb9b0 100644
--- a/proofmode/coq_tactics.v
+++ b/proofmode/coq_tactics.v
@@ -1,5 +1,5 @@
 From iris.algebra Require Export upred.
-From iris.algebra Require Import upred_big_op upred_tactics gmap.
+From iris.algebra Require Import upred_big_op upred_tactics.
 From iris.proofmode Require Export environments classes.
 From iris.prelude Require Import stringmap hlist.
 Import uPred.
@@ -135,7 +135,7 @@ Proof. intros. rewrite envs_lookup_sound //. by rewrite always_if_elim. Qed.
 Lemma envs_lookup_persistent_sound Δ i P :
   envs_lookup i Δ = Some (true,P) → Δ ⊢ □ P ★ Δ.
 Proof.
-  intros. apply: always_entails_l. by rewrite envs_lookup_sound // sep_elim_l.
+  intros. apply (always_entails_l _ _). by rewrite envs_lookup_sound // sep_elim_l.
 Qed.
 
 Lemma envs_lookup_split Δ i p P :
@@ -277,8 +277,8 @@ Proof.
 Qed.
 Global Instance of_envs_proper : Proper (envs_Forall2 (⊣⊢) ==> (⊣⊢)) (@of_envs M).
 Proof.
-  intros Δ1 Δ2 ?; apply (anti_symm (⊢)); apply of_envs_mono;
-    eapply envs_Forall2_impl; [| |symmetry|]; eauto using equiv_entails.
+  intros Δ1 Δ2 HΔ; apply (anti_symm (⊢)); apply of_envs_mono;
+    eapply (envs_Forall2_impl (⊣⊢)); [| |symmetry|]; eauto using equiv_entails.
 Qed.
 Global Instance Envs_mono (R : relation (uPred M)) :
   Proper (env_Forall2 R ==> env_Forall2 R ==> envs_Forall2 R) (@Envs M).
@@ -375,19 +375,20 @@ Proof.
   by rewrite right_id always_and_sep_l' wand_elim_r HQ.
 Qed.
 
-Lemma tac_timeless Δ Δ' i p P Q :
-  TimelessElim Q →
-  envs_lookup i Δ = Some (p, ▷ P)%I → TimelessP P →
-  envs_simple_replace i p (Esnoc Enil i P) Δ = Some Δ' →
+Lemma tac_timeless Δ Δ' i p P P' Q :
+  IsExceptLast Q →
+  envs_lookup i Δ = Some (p, P) → IntoExceptLast P 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 right_id HQ -{2}(is_except_last Q).
+  by rewrite (into_except_last P) -except_last_always_if except_last_frame_r wand_elim_r.
 Qed.
 
 (** * Always *)
 Lemma tac_always_intro Δ Q : envs_persistent Δ = true → (Δ ⊢ Q) → Δ ⊢ □ Q.
-Proof. intros. by apply: always_intro. Qed.
+Proof. intros. by apply (always_intro _ _). Qed.
 
 Lemma tac_persistent Δ Δ' i p P P' Q :
   envs_lookup i Δ = Some (p, P) → IntoPersistentP P P' →
@@ -465,6 +466,8 @@ Class IntoAssert (P : uPred M) (Q : uPred M) (R : uPred M) :=
 Global Arguments into_assert _ _ _ {_}.
 Lemma into_assert_default P Q : IntoAssert P Q P.
 Proof. by rewrite /IntoAssert wand_elim_r. Qed.
+Global Instance to_assert_rvs P Q : IntoAssert P (|=r=> Q) (|=r=> P).
+Proof. by rewrite /IntoAssert rvs_frame_r wand_elim_r rvs_trans. Qed.
 
 Lemma tac_specialize_assert Δ Δ' Δ1 Δ2' j q lr js R P1 P2 P1' Q :
   envs_lookup_delete j Δ = Some (q, R, Δ') →
@@ -549,25 +552,13 @@ Proof.
   by rewrite right_id {1}(persistentP P) always_and_sep_l wand_elim_r.
 Qed.
 
-(** Whenever posing [lem : True ⊢ Q] as [H] we want it to appear as [H : Q] and
-not as [H : True -★ Q]. The class [IntoPosedProof] is used to strip off the
-[True]. Note that [to_posed_proof_True] is declared using a [Hint Extern] to
-make sure it is not used while posing [lem : ?P ⊢ Q] with [?P] an evar. *)
-Class IntoPosedProof (P1 P2 R : uPred M) :=
-  into_pose_proof : (P1 ⊢ P2) → True ⊢ R.
-Arguments into_pose_proof : clear implicits.
-Instance to_posed_proof_True P : IntoPosedProof True P P.
-Proof. by rewrite /IntoPosedProof. Qed.
-Global Instance to_posed_proof_wand P Q : IntoPosedProof P Q (P -★ Q).
-Proof. rewrite /IntoPosedProof. apply entails_wand. Qed.
-
-Lemma tac_pose_proof Δ Δ' j P1 P2 R Q :
-  (P1 ⊢ P2) → IntoPosedProof P1 P2 R →
-  envs_app true (Esnoc Enil j R) Δ = Some Δ' →
+Lemma tac_pose_proof Δ Δ' j P Q :
+  (True ⊢ P) →
+  envs_app true (Esnoc Enil j P) Δ = Some Δ' →
   (Δ' ⊢ Q) → Δ ⊢ Q.
 Proof.
-  intros HP ?? <-. rewrite envs_app_sound //; simpl.
-  by rewrite right_id -(into_pose_proof P1 P2 R) // always_pure wand_True.
+  intros HP ? <-. rewrite envs_app_sound //; simpl.
+  by rewrite right_id -HP always_pure wand_True.
 Qed.
 
 Lemma tac_pose_proof_hyp Δ Δ' Δ'' i p j P Q :
@@ -661,13 +652,27 @@ Proof.
 Qed.
 
 (** * Conjunction/separating conjunction elimination *)
-Lemma tac_sep_destruct Δ Δ' i p j1 j2 P P1 P2 Q :
-  envs_lookup i Δ = Some (p, P) → IntoSep p P P1 P2 →
+Lemma tac_and_destruct Δ Δ' i p j1 j2 P P1 P2 Q :
+  envs_lookup i Δ = Some (p, P) → IntoAnd p P P1 P2 →
   envs_simple_replace i p (Esnoc (Esnoc Enil j1 P1) j2 P2) Δ = Some Δ' →
   (Δ' ⊢ Q) → Δ ⊢ Q.
+Proof.
+  intros. rewrite envs_simple_replace_sound //; simpl. rewrite (into_and p P).
+  by destruct p; rewrite /= ?right_id (comm _ P1) ?always_and_sep wand_elim_r.
+Qed.
+
+(* Using this tactic, one can destruct a (non-separating) conjunction in the
+spatial context as long as one of the conjuncts is thrown away. It corresponds
+to the principle of "external choice" in linear logic. *)
+Lemma tac_and_destruct_choice Δ Δ' i p (lr : bool) j P P1 P2 Q :
+  envs_lookup i Δ = Some (p, P) → IntoAnd true P P1 P2 →
+  envs_simple_replace i p (Esnoc Enil j (if lr then P1 else P2)) Δ = Some Δ' →
+  (Δ' ⊢ Q) → Δ ⊢ Q.
 Proof.
   intros. rewrite envs_simple_replace_sound //; simpl.
-  by rewrite (into_sep p P) right_id (comm uPred_sep P1) wand_elim_r.
+  rewrite right_id (into_and true P). destruct lr.
+  - by rewrite and_elim_l wand_elim_r.
+  - by rewrite and_elim_r wand_elim_r.
 Qed.
 
 (** * Framing *)
@@ -744,7 +749,18 @@ Proof.
   apply exist_elim=> a; destruct (HΦ a) as (Δ'&?&?).
   rewrite envs_simple_replace_sound' //; simpl. by rewrite right_id wand_elim_r.
 Qed.
-End tactics.
 
-Hint Extern 0 (IntoPosedProof True _ _) =>
-  class_apply @to_posed_proof_True : typeclass_instances.
+(** * 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..a1a2804ed8db7286ab15f047b6bebb2fcfc5c206 100644
--- a/proofmode/ghost_ownership.v
+++ b/proofmode/ghost_ownership.v
@@ -3,12 +3,12 @@ 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 :
-  IntoOp a b1 b2 → IntoSep p (own γ a) (own γ b1) (own γ b2).
-Proof. rewrite /IntoOp /IntoSep => ->. by rewrite own_op. Qed.
+Global Instance into_and_own p γ a b1 b2 :
+  IntoOp a b1 b2 → IntoAnd p (own γ a) (own γ b1) (own γ b2).
+Proof. intros. apply mk_into_and_sep. by rewrite (into_op a) own_op. Qed.
 Global Instance from_sep_own γ a b :
   FromSep (own γ (a ⋅ b)) (own γ a) (own γ b) | 90.
 Proof. by rewrite /FromSep own_op. Qed.
diff --git a/proofmode/intro_patterns.v b/proofmode/intro_patterns.v
index 0f55cbf3b034bf6bb7f220ec16f66318c4c2db33..ac49db83957e8c257589e6b1940a64d5773e6f60 100644
--- a/proofmode/intro_patterns.v
+++ b/proofmode/intro_patterns.v
@@ -3,14 +3,18 @@ From iris.prelude Require Export strings.
 Inductive intro_pat :=
   | IName : string → intro_pat
   | IAnom : intro_pat
-  | IAnomPure : intro_pat
   | IDrop : intro_pat
   | IFrame : intro_pat
-  | IPersistent : intro_pat → intro_pat
   | IList : list (list intro_pat) → intro_pat
+  | IPureElim : intro_pat
+  | IAlwaysElim : intro_pat → intro_pat
+  | ILaterElim : intro_pat → intro_pat
+  | IVsElim : intro_pat → intro_pat
+  | IPureIntro : intro_pat
+  | IAlwaysIntro : intro_pat
+  | ILaterIntro : intro_pat
+  | IVsIntro : intro_pat
   | ISimpl : intro_pat
-  | IAlways : intro_pat
-  | INext : intro_pat
   | IForall : intro_pat
   | IAll : intro_pat
   | IClear : list (bool * string) → intro_pat. (* true = frame, false = clear *)
@@ -19,23 +23,27 @@ Module intro_pat.
 Inductive token :=
   | TName : string → token
   | TAnom : token
-  | TAnomPure : token
   | TDrop : token
   | TFrame : token
-  | TPersistent : token
   | TBar : token
   | TBracketL : token
   | TBracketR : token
   | TAmp : token
   | TParenL : token
   | TParenR : token
+  | TPureElim : token
+  | TAlwaysElim : token
+  | TLaterElim : token
+  | TVsElim : token
+  | TPureIntro : token
+  | TAlwaysIntro : token
+  | TLaterIntro : token
+  | TVsIntro : token
   | TSimpl : token
-  | TAlways : token
-  | TNext : token
-  | TClearL : token
-  | TClearR : token
   | TForall : token
-  | TAll : token.
+  | TAll : token
+  | TClearL : token
+  | TClearR : token.
 
 Fixpoint cons_name (kn : string) (k : list token) : list token :=
   match kn with "" => k | _ => TName (string_rev kn) :: k end.
@@ -44,18 +52,24 @@ Fixpoint tokenize_go (s : string) (k : list token) (kn : string) : list token :=
   | "" => rev (cons_name kn k)
   | String " " s => tokenize_go s (cons_name kn k) ""
   | String "?" s => tokenize_go s (TAnom :: cons_name kn k) ""
-  | String "%" s => tokenize_go s (TAnomPure :: cons_name kn k) ""
   | String "_" s => tokenize_go s (TDrop :: cons_name kn k) ""
   | String "$" s => tokenize_go s (TFrame :: cons_name kn k) ""
-  | String "#" s => tokenize_go s (TPersistent :: cons_name kn k) ""
   | String "[" s => tokenize_go s (TBracketL :: cons_name kn k) ""
   | String "]" s => tokenize_go s (TBracketR :: cons_name kn k) ""
   | String "|" s => tokenize_go s (TBar :: cons_name kn k) ""
   | String "(" s => tokenize_go s (TParenL :: cons_name kn k) ""
   | String ")" s => tokenize_go s (TParenR :: cons_name kn k) ""
   | String "&" s => tokenize_go s (TAmp :: cons_name kn k) ""
-  | String "!" s => tokenize_go s (TAlways :: cons_name kn k) ""
-  | String ">" s => tokenize_go s (TNext :: cons_name kn k) ""
+  | String "%" s => tokenize_go s (TPureElim :: cons_name kn k) ""
+  | String "#" s => tokenize_go s (TAlwaysElim :: cons_name kn k) ""
+  | String ">" s => tokenize_go s (TLaterElim :: cons_name kn k) ""
+  | String "=" (String "=" (String ">" s)) =>
+     tokenize_go s (TVsElim :: cons_name kn k) ""
+  | String "!" (String "%" s) => tokenize_go s (TPureIntro :: cons_name kn k) ""
+  | String "!" (String "#" s) => tokenize_go s (TAlwaysIntro :: cons_name kn k) ""
+  | String "!" (String ">" s) => tokenize_go s (TLaterIntro :: cons_name kn k) ""
+  | String "!" (String "=" (String "=" (String ">" s))) =>
+     tokenize_go s (TVsIntro :: cons_name kn k) ""
   | String "{" s => tokenize_go s (TClearL :: cons_name kn k) ""
   | String "}" s => tokenize_go s (TClearR :: cons_name kn k) ""
   | String "/" (String "=" s) => tokenize_go s (TSimpl :: cons_name kn k) ""
@@ -67,11 +81,13 @@ Definition tokenize (s : string) : list token := tokenize_go s [] "".
 
 Inductive stack_item :=
   | SPat : intro_pat → stack_item
-  | SPersistent : stack_item
   | SList : stack_item
   | SConjList : stack_item
   | SBar : stack_item
-  | SAmp : stack_item.
+  | SAmp : stack_item
+  | SAlwaysElim : stack_item
+  | SLaterElim : stack_item
+  | SVsElim : stack_item.
 Notation stack := (list stack_item).
 
 Fixpoint close_list (k : stack)
@@ -79,8 +95,11 @@ Fixpoint close_list (k : stack)
   match k with
   | SList :: k => Some (SPat (IList (ps :: pss)) :: k)
   | SPat pat :: k => close_list k (pat :: ps) pss
-  | SPersistent :: k =>
-     '(p,ps) ← maybe2 (::) ps; close_list k (IPersistent p :: ps) pss
+  | SAlwaysElim :: k =>
+     '(p,ps) ← maybe2 (::) ps; close_list k (IAlwaysElim p :: ps) pss
+  | SLaterElim :: k =>
+     '(p,ps) ← maybe2 (::) ps; close_list k (ILaterElim p :: ps) pss
+  | SVsElim :: k => '(p,ps) ← maybe2 (::) ps; close_list k (IVsElim p :: ps) pss
   | SBar :: k => close_list k [] (ps :: pss)
   | _ => None
   end.
@@ -102,7 +121,9 @@ Fixpoint close_conj_list (k : stack)
           end;
      Some (SPat (big_conj ps) :: k)
   | SPat pat :: k => guard (cur = None); close_conj_list k (Some pat) ps
-  | SPersistent :: k => p ← cur; close_conj_list k (Some (IPersistent p)) ps
+  | SAlwaysElim :: k => p ← cur; close_conj_list k (Some (IAlwaysElim p)) ps
+  | SLaterElim :: k => p ← cur; close_conj_list k (Some (ILaterElim p)) ps
+  | SVsElim :: k => p ← cur; close_conj_list k (Some (IVsElim p)) ps
   | SAmp :: k => p ← cur; close_conj_list k None (p :: ps)
   | _ => None
   end.
@@ -112,19 +133,23 @@ Fixpoint parse_go (ts : list token) (k : stack) : option stack :=
   | [] => Some k
   | TName s :: ts => parse_go ts (SPat (IName s) :: k)
   | TAnom :: ts => parse_go ts (SPat IAnom :: k)
-  | TAnomPure :: ts => parse_go ts (SPat IAnomPure :: k)
   | TDrop :: ts => parse_go ts (SPat IDrop :: k)
   | TFrame :: ts => parse_go ts (SPat IFrame :: k)
-  | TPersistent :: ts => parse_go ts (SPersistent :: k)
   | TBracketL :: ts => parse_go ts (SList :: k)
   | TBar :: ts => parse_go ts (SBar :: k)
   | TBracketR :: ts => close_list k [] [] ≫= parse_go ts
   | TParenL :: ts => parse_go ts (SConjList :: k)
   | TAmp :: ts => parse_go ts (SAmp :: k)
   | TParenR :: ts => close_conj_list k None [] ≫= parse_go ts
+  | TPureElim :: ts => parse_go ts (SPat IPureElim :: k)
+  | TAlwaysElim :: ts => parse_go ts (SAlwaysElim :: k)
+  | TLaterElim :: ts => parse_go ts (SLaterElim :: k)
+  | TVsElim :: ts => parse_go ts (SVsElim :: k)
+  | TPureIntro :: ts => parse_go ts (SPat IPureIntro :: k)
+  | TAlwaysIntro :: ts => parse_go ts (SPat IAlwaysIntro :: k)
+  | TLaterIntro :: ts => parse_go ts (SPat ILaterIntro :: k)
+  | TVsIntro :: ts => parse_go ts (SPat IVsIntro :: k)
   | TSimpl :: ts => parse_go ts (SPat ISimpl :: k)
-  | TAlways :: ts => parse_go ts (SPat IAlways :: k)
-  | TNext :: ts => parse_go ts (SPat INext :: k)
   | TAll :: ts => parse_go ts (SPat IAll :: k)
   | TForall :: ts => parse_go ts (SPat IForall :: k)
   | TClearL :: ts => parse_clear ts [] k
diff --git a/proofmode/invariants.v b/proofmode/invariants.v
index 46fb3cf1e01432776d1a73f861df26783c7da0b2..d5813093ec389667f2a166db0f299c1b7b0a712e 100644
--- a/proofmode/invariants.v
+++ b/proofmode/invariants.v
@@ -1,89 +1,29 @@
-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 || match goal with |- ?P => fail "iInv: cannot solve" P end
+    |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 df3203323e96350820f988dab219d639f553fec6..e3641da1c84232d2986ba90c6d1dd88db33e459c 100644
--- a/proofmode/pviewshifts.v
+++ b/proofmode/pviewshifts.v
@@ -1,157 +1,58 @@
-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.
+
 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 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.
 
-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.
+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.
 
-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 is_except_last_pvs E1 E2 P : IsExceptLast (|={E1,E2}=> P).
+Proof. by rewrite /IsExceptLast except_last_pvs. 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 from_vs_pvs E P : FromVs (|={E}=> P) P.
+Proof. by rewrite /FromVs -rvs_pvs. Qed.
 
-Lemma tac_pvs_intro Δ E1 E2 Q : E1 = E2 → (Δ ⊢ Q) → Δ ⊢ |={E1,E2}=> Q.
-Proof. intros -> ->. apply pvs_intro. 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.
-
-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(H) :=
-  iDestructHelp H as (fun H => iPvsCore H; last iDestruct H as "?").
-Tactic Notation "iPvs" open_constr(H) "as" constr(pat) :=
-  iDestructHelp H as (fun H => iPvsCore H; last iDestruct H as pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1) ")"
-    constr(pat) :=
-  iDestructHelp H as (fun H => iPvsCore H; last iDestruct H as ( x1 ) pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) ")" constr(pat) :=
-  iDestructHelp H as (fun H => iPvsCore H; last iDestruct H as ( x1 x2 ) pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
-  iDestructHelp H as (fun H => iPvsCore H; last iDestruct H as ( x1 x2 x3 ) pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
-    constr(pat) :=
-  iDestructHelp H as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 ) pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) ")" constr(pat) :=
-  iDestructHelp H as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 ) pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) :=
-  iDestructHelp H as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 ) pat).
-Tactic Notation "iPvs" open_constr(H) "as" "(" simple_intropattern(x1)
-    simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
-    simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")"
-    constr(pat) :=
-  iDestructHelp H as (fun H =>
-    iPvsCore H; last iDestruct H as ( x1 x2 x3 x4 x5 x6 x7 ) pat).
-Tactic Notation "iPvs" open_constr(H) "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) :=
-  iDestructHelp H 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/spec_patterns.v b/proofmode/spec_patterns.v
index 6489692eae596672c8a5e3ee1d2e0b99c171c8c6..03f6d40f97dbd8b10b6edb5d3c8f8009f4c67b25 100644
--- a/proofmode/spec_patterns.v
+++ b/proofmode/spec_patterns.v
@@ -1,6 +1,6 @@
 From iris.prelude Require Export strings.
 
-Inductive spec_goal_kind := GoalStd | GoalPvs.
+Inductive spec_goal_kind := GoalStd | GoalVs.
 
 Inductive spec_pat :=
   | SGoal : spec_goal_kind → bool → list string → spec_pat
@@ -18,7 +18,7 @@ Inductive token :=
   | TPersistent : token
   | TPure : token
   | TForall : token
-  | TPvs : token.
+  | TVs : token.
 
 Fixpoint cons_name (kn : string) (k : list token) : list token :=
   match kn with "" => k | _ => TName (string_rev kn) :: k end.
@@ -32,8 +32,7 @@ Fixpoint tokenize_go (s : string) (k : list token) (kn : string) : list token :=
   | String "#" s => tokenize_go s (TPersistent :: cons_name kn k) ""
   | String "%" s => tokenize_go s (TPure :: cons_name kn k) ""
   | String "*" s => tokenize_go s (TForall :: cons_name kn k) ""
-  | String "|" (String "=" (String "=" (String ">" s))) =>
-     tokenize_go s (TPvs :: cons_name kn k) ""
+  | String "=" (String "=" (String ">" s)) => tokenize_go s (TVs :: cons_name kn k) ""
   | String a s => tokenize_go s k (String a kn)
   end.
 Definition tokenize (s : string) : list token := tokenize_go s [] "".
@@ -49,8 +48,8 @@ Fixpoint parse_go (ts : list token) (k : list spec_pat) : option (list spec_pat)
   | TBracketL :: TPersistent :: TBracketR :: ts => parse_go ts (SGoalPersistent :: k)
   | TBracketL :: TPure :: TBracketR :: ts => parse_go ts (SGoalPure :: k)
   | TBracketL :: ts => parse_goal ts GoalStd false [] k
-  | TPvs :: TBracketL :: ts => parse_goal ts GoalPvs false [] k
-  | TPvs :: ts => parse_go ts (SGoal GoalPvs true [] :: k)
+  | TVs :: TBracketL :: ts => parse_goal ts GoalVs false [] k
+  | TVs :: ts => parse_go ts (SGoal GoalVs true [] :: k)
   | TPersistent :: TName s :: ts => parse_go ts (SName true s :: k)
   | TForall :: ts => parse_go ts (SForall :: k)
   | _ => None
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 3901386c43ad909a114baeac44a5608ad5a6b1fd..61646305a637a5b866715b161adcd58d7fe2dc8e 100644
--- a/proofmode/tactics.v
+++ b/proofmode/tactics.v
@@ -1,6 +1,7 @@
 From iris.proofmode Require Import coq_tactics intro_patterns spec_patterns.
 From iris.algebra Require Export upred.
 From iris.proofmode Require Export classes notation.
+From iris.proofmode Require Import class_instances.
 From iris.prelude Require Import stringmap hlist.
 
 Declare Reduction env_cbv := cbv [
@@ -25,7 +26,7 @@ Ltac iFresh :=
      first use [cbv] to compute the domain of [Δ] *)
      let Hs := eval cbv in (envs_dom Δ) in
      eval vm_compute in (fresh_string_of_set "~" (of_list Hs))
-  | _ => constr:"~"
+  | _ => constr:("~")
   end.
 
 Tactic Notation "iTypeOf" constr(H) tactic(tac):=
@@ -106,14 +107,14 @@ Local Tactic Notation "iPersistent" constr(H) :=
   eapply tac_persistent with _ H _ _ _; (* (i:=H) *)
     [env_cbv; reflexivity || fail "iPersistent:" H "not found"
     |let Q := match goal with |- IntoPersistentP ?Q _ => Q end in
-     apply _ || fail "iPersistent:" H ":" Q "not persistent"
+     apply _ || fail "iPersistent:" Q "not persistent"
     |env_cbv; reflexivity|].
 
 Local Tactic Notation "iPure" constr(H) "as" simple_intropattern(pat) :=
   eapply tac_pure with _ H _ _ _; (* (i:=H1) *)
     [env_cbv; reflexivity || fail "iPure:" H "not found"
     |let P := match goal with |- IntoPure ?P _ => P end in
-     apply _ || fail "iPure:" H ":" P "not pure"
+     apply _ || fail "iPure:" P "not pure"
     |intros pat].
 
 Tactic Notation "iPureIntro" :=
@@ -127,9 +128,9 @@ Record iTrm {X As} :=
 Arguments ITrm {_ _} _ _ _.
 
 Notation "( H $! x1 .. xn )" :=
-  (ITrm H (hcons x1 .. (hcons xn hnil) ..) "") (at level 0, x1, xn at level 0).
+  (ITrm H (hcons x1 .. (hcons xn hnil) ..) "") (at level 0, x1, xn at level 9).
 Notation "( H $! x1 .. xn 'with' pat )" :=
-  (ITrm H (hcons x1 .. (hcons xn hnil) ..) pat) (at level 0, x1, xn at level 0).
+  (ITrm H (hcons x1 .. (hcons xn hnil) ..) pat) (at level 0, x1, xn at level 9).
 Notation "( H 'with' pat )" := (ITrm H hnil pat) (at level 0).
 
 Local Tactic Notation "iSpecializeArgs" constr(H) open_constr(xs) :=
@@ -138,14 +139,15 @@ Local Tactic Notation "iSpecializeArgs" constr(H) open_constr(xs) :=
   | _ =>
     eapply tac_forall_specialize with _ H _ _ _ xs; (* (i:=H) (a:=x) *)
       [env_cbv; reflexivity || fail 1 "iSpecialize:" H "not found"
-      |apply _ || fail 1 "iSpecialize:" H "not a forall of the right arity or type"
+      |let P := match goal with |- ForallSpecialize _ ?P _ => P end in
+       apply _ || fail 1 "iSpecialize:" P "not a forall of the right arity or type"
       |cbn [himpl hcurry]; reflexivity|]
   end.
 
 Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
   let solve_to_wand H1 :=
     let P := match goal with |- IntoWand ?P _ _ => P end in
-    apply _ || fail "iSpecialize:" H1 ":" P "not an implication/wand" in
+    apply _ || fail "iSpecialize:" P "not an implication/wand" in
   let rec go H1 pats :=
     lazymatch pats with
     | [] => idtac
@@ -156,7 +158,7 @@ Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
          |env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
          |let P := match goal with |- IntoWand ?P ?Q _ => P end in
           let Q := match goal with |- IntoWand ?P ?Q _ => Q end in
-          apply _ || fail "iSpecialize: cannot instantiate" H1 ":" P "with" H2 ":" Q
+          apply _ || fail "iSpecialize: cannot instantiate" P "with" Q
          |env_cbv; reflexivity|go H1 pats]
     | SName true ?H2 :: ?pats =>
        eapply tac_specialize_persistent with _ _ H1 _ _ _ _;
@@ -195,7 +197,7 @@ Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
          |solve_to_wand H1
          |match k with
           | GoalStd => apply into_assert_default
-          | GoalPvs => apply _ || fail "iSpecialize: cannot generate pvs goal"
+          | GoalVs => apply _ || fail "iSpecialize: cannot generate view shifted goal"
           end
          |env_cbv; reflexivity || fail "iSpecialize:" Hs "not found"
          |(*goal*)
@@ -208,50 +210,82 @@ Tactic Notation "iSpecialize" open_constr(t) :=
   end.
 
 (** * Pose proof *)
-Local Tactic Notation "iPoseProofCore" open_constr(H1) "as" constr(H2) :=
-  lazymatch type of H1 with
-  | string =>
-     eapply tac_pose_proof_hyp with _ _ H1 _ H2 _;
-       [env_cbv; reflexivity || fail "iPoseProof:" H1 "not found"
-       |env_cbv; reflexivity || fail "iPoseProof:" H2 "not fresh"|]
-  | _ =>
-     eapply tac_pose_proof with _ H2 _ _ _; (* (j:=H) *)
-       [first [eapply H1|apply uPred.equiv_iff; eapply H1]
-       |apply _
-       |env_cbv; reflexivity || fail "iPoseProof:" H2 "not fresh"|]
+(* The tactic [iIntoEntails] tactic solves a goal [True ⊢ Q]. The arguments [t]
+is a Coq term whose type is of the following shape:
+
+- [∀ (x_1 : A_1) .. (x_n : A_n), True ⊢ Q]
+- [∀ (x_1 : A_1) .. (x_n : A_n), P1 ⊢ P2], in which case [Q] becomes [P1 -★ P2]
+- [∀ (x_1 : A_1) .. (x_n : A_n), P1 ⊣⊢ P2], in which case [Q] becomes [P1 ↔ P2]
+
+The tactic instantiates each dependent argument [x_i] with an evar and generates
+a goal [P] for non-dependent arguments [x_i : P]. *)
+Tactic Notation "iIntoEntails" open_constr(t) :=
+  let rec go t :=
+    let tT := type of t in
+    lazymatch eval hnf in tT with
+    | True ⊢ _ => apply t
+    | _ ⊢ _ => apply (uPred.entails_wand _ _ t)
+    (* need to use the unfolded version of [⊣⊢] due to the hnf *)
+    | uPred_equiv' _ _ => apply (uPred.equiv_iff _ _ t)
+    | ?P → ?Q => let H := fresh in assert P as H; [|go uconstr:(t H); clear H]
+    | ∀ _ : ?T, _ =>
+       (* Put [T] inside an [id] to avoid TC inference from being invoked. *)
+       (* This is a workarround for Coq bug #4969. *)
+       let e := fresh in evar (e:id T);
+       let e' := eval unfold e in e in clear e; go (t e')
+    end
+  in go t.
+
+Tactic Notation "iPoseProofCore" open_constr(lem) "as" tactic(tac) :=
+  let pose_trm t tac :=
+    let Htmp := iFresh in
+    lazymatch type of t with
+    | string =>
+       eapply tac_pose_proof_hyp with _ _ t _ Htmp _;
+         [env_cbv; reflexivity || fail "iPoseProof:" t "not found"
+         |env_cbv; reflexivity || fail "iPoseProof:" Htmp "not fresh"
+         |tac Htmp]
+    | _ =>
+       eapply tac_pose_proof with _ Htmp _; (* (j:=H) *)
+         [iIntoEntails t
+         |env_cbv; reflexivity || fail "iPoseProof:" Htmp "not fresh"
+         |tac Htmp]
+    end;
+    try (apply _) (* solve TC constraints. It is essential that this happens
+    after the continuation [tac] has been called. *)
+  in lazymatch lem with
+  | ITrm ?t ?xs ?pat =>
+     pose_trm t ltac:(fun Htmp =>
+       iSpecializeArgs Htmp xs; iSpecializePat Htmp pat; last (tac Htmp))
+  | _ => pose_trm lem tac
   end.
 
-Tactic Notation "iPoseProof" open_constr(t) "as" constr(H) :=
-  lazymatch t with
-  | ITrm ?H1 ?xs ?pat =>
-     iPoseProofCore H1 as H; last (iSpecializeArgs H xs; iSpecializePat H pat)
-  | _ => iPoseProofCore t as H
-  end.
+Tactic Notation "iPoseProof" open_constr(lem) "as" constr(H) :=
+  iPoseProofCore lem as (fun Htmp => iRename Htmp into H).
 
-Tactic Notation "iPoseProof" open_constr(t) :=
-  let H := iFresh in iPoseProof t as H.
+Tactic Notation "iPoseProof" open_constr(lem) :=
+  iPoseProofCore lem as (fun _ => idtac).
 
 (** * Apply *)
-Tactic Notation "iApply" open_constr(t) :=
+Tactic Notation "iApply" open_constr(lem) :=
   let finish H := first
     [iExact H
     |eapply tac_apply with _ H _ _ _;
        [env_cbv; reflexivity || fail 1 "iApply:" H "not found"
        |let P := match goal with |- IntoWand ?P _ _ => P end in
-        apply _ || fail 1 "iApply: cannot apply" H ":" P
+        apply _ || fail 1 "iApply: cannot apply" P
        |lazy beta (* reduce betas created by instantiation *)]] in
-  let Htmp := iFresh in
-  lazymatch t with
-  | ITrm ?H ?xs ?pat =>
-     iPoseProofCore H as Htmp; last (
+  lazymatch lem with
+  | ITrm ?t ?xs ?pat =>
+     iPoseProofCore t as (fun Htmp =>
        iSpecializeArgs Htmp xs;
        try (iSpecializeArgs Htmp (hcons _ _));
        iSpecializePat Htmp pat; last finish Htmp)
   | _ =>
-     iPoseProofCore t as Htmp; last (
+     iPoseProofCore lem as (fun Htmp =>
        try (iSpecializeArgs Htmp (hcons _ _));
        finish Htmp)
-  end; try apply _.
+  end.
 
 (** * Revert *)
 Local Tactic Notation "iForallRevert" ident(x) :=
@@ -331,7 +365,7 @@ Local Tactic Notation "iOrDestruct" constr(H) "as" constr(H1) constr(H2) :=
   eapply tac_or_destruct with _ _ H _ H1 H2 _ _ _; (* (i:=H) (j1:=H1) (j2:=H2) *)
     [env_cbv; reflexivity || fail "iOrDestruct:" H "not found"
     |let P := match goal with |- IntoOr ?P _ _ => P end in
-     apply _ || fail "iOrDestruct:" H ":" P "not a disjunction"
+     apply _ || fail "iOrDestruct: cannot destruct" P
     |env_cbv; reflexivity || fail "iOrDestruct:" H1 "not fresh"
     |env_cbv; reflexivity || fail "iOrDestruct:" H2 "not fresh"| |].
 
@@ -350,46 +384,32 @@ Tactic Notation "iSplitL" constr(Hs) :=
   eapply tac_sep_split with _ _ false Hs _ _; (* (js:=Hs) *)
     [let P := match goal with |- FromSep ?P _ _ => P end in
      apply _ || fail "iSplitL:" P "not a separating conjunction"
-    |env_cbv; reflexivity || fail "iSplitL: hypotheses" Hs "not found"| |].
+    |env_cbv; reflexivity || fail "iSplitL: hypotheses" Hs
+                                  "not found in the spatial context"| |].
 Tactic Notation "iSplitR" constr(Hs) :=
   let Hs := words Hs in
   eapply tac_sep_split with _ _ true Hs _ _; (* (js:=Hs) *)
     [let P := match goal with |- FromSep ?P _ _ => P end in
      apply _ || fail "iSplitR:" P "not a separating conjunction"
-    |env_cbv; reflexivity || fail "iSplitR: hypotheses" Hs "not found"| |].
+    |env_cbv; reflexivity || fail "iSplitR: hypotheses" Hs
+                                  "not found in the spatial context"| |].
 
 Tactic Notation "iSplitL" := iSplitR "".
 Tactic Notation "iSplitR" := iSplitL "".
 
-Local Tactic Notation "iSepDestruct" constr(H) "as" constr(H1) constr(H2) :=
-  eapply tac_sep_destruct with _ H _ H1 H2 _ _ _; (* (i:=H) (j1:=H1) (j2:=H2) *)
-    [env_cbv; reflexivity || fail "iSepDestruct:" H "not found"
-    |let P := match goal with |- IntoSep _ ?P _ _ => P end in
-     apply _ || fail "iSepDestruct:" H ":" P "not separating destructable"
-    |env_cbv; reflexivity || fail "iSepDestruct:" H1 "or" H2 " not fresh"|].
-
-Tactic Notation "iFrame" constr(Hs) :=
-  let rec go Hs :=
-    match Hs with
-    | [] => idtac
-    | ?H :: ?Hs =>
-       eapply tac_frame with _ H _ _ _;
-         [env_cbv; reflexivity || fail "iFrame:" H "not found"
-         |let R := match goal with |- Frame ?R _ _ => R end in
-          apply _ || fail "iFrame: cannot frame" R
-         |lazy iota beta; go Hs]
-    end
-  in let Hs := words Hs in go Hs.
+Local Tactic Notation "iAndDestruct" constr(H) "as" constr(H1) constr(H2) :=
+  eapply tac_and_destruct with _ H _ H1 H2 _ _ _; (* (i:=H) (j1:=H1) (j2:=H2) *)
+    [env_cbv; reflexivity || fail "iAndDestruct:" H "not found"
+    |let P := match goal with |- IntoAnd _ ?P _ _ => P end in
+     apply _ || fail "iAndDestruct: cannot destruct" P
+    |env_cbv; reflexivity || fail "iAndDestruct:" H1 "or" H2 " not fresh"|].
 
-Tactic Notation "iFrame" :=
-  let rec go Hs :=
-    match Hs with
-    | [] => idtac
-    | ?H :: ?Hs => try iFrame H; go Hs
-    end in
-  match goal with
-  | |- of_envs ?Δ ⊢ _ => let Hs := eval cbv in (env_dom (env_spatial Δ)) in go Hs
-  end.
+Local Tactic Notation "iAndDestructChoice" constr(H) "as" constr(lr) constr(H') :=
+  eapply tac_and_destruct_choice with _ H _ lr H' _ _ _;
+    [env_cbv; reflexivity || fail "iAndDestruct:" H "not found"
+    |let P := match goal with |- IntoAnd _ ?P _ _ => P end in
+     apply _ || fail "iAndDestruct: cannot destruct" P
+    |env_cbv; reflexivity || fail "iAndDestruct:" H' " not fresh"|].
 
 Tactic Notation "iCombine" constr(H1) constr(H2) "as" constr(H) :=
   eapply tac_combine with _ _ _ H1 _ _ H2 _ _ H _;
@@ -397,9 +417,45 @@ Tactic Notation "iCombine" constr(H1) constr(H2) "as" constr(H) :=
     |env_cbv; reflexivity || fail "iCombine:" H2 "not found"
     |let P1 := match goal with |- FromSep _ ?P1 _ => P1 end in
      let P2 := match goal with |- FromSep _ _ ?P2 => P2 end in
-     apply _ || fail "iCombine: cannot combine" H1 ":" P1 "and" H2 ":" P2
+     apply _ || fail "iCombine: cannot combine" P1 "and" P2
     |env_cbv; reflexivity || fail "iCombine:" H "not fresh"|].
 
+(** Framing *)
+Local Ltac iFrameHyp H :=
+  eapply tac_frame with _ H _ _ _;
+    [env_cbv; reflexivity || fail "iFrame:" H "not found"
+    |let R := match goal with |- Frame ?R _ _ => R end in
+     apply _ || fail "iFrame: cannot frame" R
+    |lazy iota beta].
+
+Local Ltac iFramePersistent :=
+  let rec go Hs :=
+    match Hs with [] => idtac | ?H :: ?Hs => repeat iFrameHyp H; go Hs end in
+  match goal with
+  | |- of_envs ?Δ ⊢ _ =>
+     let Hs := eval cbv in (env_dom (env_persistent Δ)) in go Hs
+  end.
+
+Local Ltac iFrameSpatial :=
+  let rec go Hs :=
+    match Hs with [] => idtac | ?H :: ?Hs => try iFrameHyp H; go Hs end in
+  match goal with
+  | |- of_envs ?Δ ⊢ _ =>
+     let Hs := eval cbv in (env_dom (env_spatial Δ)) in go Hs
+  end.
+
+Tactic Notation "iFrame" constr(Hs) :=
+  let rec go Hs :=
+    match Hs with
+    | [] => idtac
+    | "#" :: ?Hs => iFramePersistent; go Hs
+    | "★" :: ?Hs => iFrameSpatial; go Hs
+    | ?H :: ?Hs => iFrameHyp H; go Hs
+    end
+  in let Hs := words Hs in go Hs.
+
+Tactic Notation "iFrame" := iFrameSpatial.
+
 (** * Existential *)
 Tactic Notation "iExists" uconstr(x1) :=
   eapply tac_exist;
@@ -433,26 +489,66 @@ Local Tactic Notation "iExistDestruct" constr(H)
   eapply tac_exist_destruct with H _ Hx _ _; (* (i:=H) (j:=Hx) *)
     [env_cbv; reflexivity || fail "iExistDestruct:" H "not found"
     |let P := match goal with |- IntoExist ?P _ => P end in
-     apply _ || fail "iExistDestruct:" H ":" P "not an existential"|];
+     apply _ || fail "iExistDestruct: cannot destruct" P|];
   let y := fresh in
   intros y; eexists; split;
     [env_cbv; reflexivity || fail "iExistDestruct:" Hx "not fresh"
     |revert y; intros x].
 
+(** * Always *)
+Tactic Notation "iAlways":=
+  apply tac_always_intro;
+    [reflexivity || fail "iAlways: spatial context non-empty"|].
+
+(** * Later *)
+Tactic Notation "iNext":=
+  eapply tac_next;
+    [apply _
+    |let P := match goal with |- FromLater ?P _ => P end in
+     apply _ || fail "iNext:" P "does not contain laters"|].
+
+Tactic Notation "iTimeless" constr(H) :=
+  eapply tac_timeless with _ H _ _ _;
+    [let Q := match goal with |- IsExceptLast ?Q => Q end in
+     apply _ || fail "iTimeless: cannot remove later when goal is" Q
+    |env_cbv; reflexivity || fail "iTimeless:" H "not found"
+    |let P := match goal with |- IntoExceptLast ?P _ => P end in
+     apply _ || fail "iTimeless:" P "not timeless"
+    |env_cbv; reflexivity|].
+
+(** * 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 view shift"|].
+
+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 run" P "in" Q
+                     "because the goal or hypothesis is not a view shift"
+    |env_cbv; reflexivity|].
+
 (** * Basic destruct tactic *)
 Local Tactic Notation "iDestructHyp" constr(H) "as" constr(pat) :=
   let rec go Hz pat :=
     lazymatch pat with
     | IAnom => idtac
-    | IAnomPure => iPure Hz as ?
     | IDrop => iClear Hz
     | IFrame => iFrame Hz
     | IName ?y => iRename Hz into y
-    | IPersistent ?pat => iPersistent Hz; go Hz pat
     | IList [[]] => iExFalso; iExact Hz
+    | IList [[?pat1; IDrop]] => iAndDestructChoice Hz as true Hz; go Hz pat1
+    | IList [[IDrop; ?pat2]] => iAndDestructChoice Hz as false Hz; go Hz pat2
     | IList [[?pat1; ?pat2]] =>
-       let Hy := iFresh in iSepDestruct Hz as Hz Hy; go Hz pat1; go Hy pat2
+       let Hy := iFresh in iAndDestruct Hz as Hz Hy; go Hz pat1; go Hy pat2
     | IList [[?pat1];[?pat2]] => iOrDestruct Hz as Hz Hz; [go Hz pat1|go Hz pat2]
+    | IPureElim => iPure Hz as ?
+    | IAlwaysElim ?pat => iPersistent Hz; go Hz pat
+    | ILaterElim ?pat => iTimeless Hz; go Hz pat
+    | IVsElim ?pat => iVsCore Hz; go Hz pat
     | _ => fail "iDestruct:" pat "invalid"
     end
   in let pat := intro_pat.parse_one pat in go H pat.
@@ -489,27 +585,6 @@ Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
     simple_intropattern(x8) ")" constr(pat) :=
   iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 x4 x5 x6 x7 x8 ) pat.
 
-(** * Always *)
-Tactic Notation "iAlways":=
-  apply tac_always_intro;
-    [reflexivity || fail "iAlways: spatial context non-empty"|].
-
-(** * Later *)
-Tactic Notation "iNext":=
-  eapply tac_next;
-    [apply _
-    |let P := match goal with |- FromLater ?P _ => P end in
-     apply _ || fail "iNext:" P "does not contain laters"|].
-
-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
-    |env_cbv; reflexivity || fail "iTimeless:" H "not found"
-    |let P := match goal with |- TimelessP ?P => P end in
-     apply _ || fail "iTimeless: " P "not timeless"
-    |env_cbv; reflexivity|].
-
 (** * Introduction tactic *)
 Local Tactic Notation "iIntro" "(" simple_intropattern(x) ")" := first
   [ (* (∀ _, _) *) apply tac_forall_intro; intros x
@@ -564,11 +639,18 @@ Tactic Notation "iIntros" constr(pat) :=
   let rec go pats :=
     lazymatch pats with
     | [] => idtac
+    | IPureElim :: ?pats => iIntro (?); go pats
+    | IAlwaysElim IAnom :: ?pats => let H := iFresh in iIntro #H; go pats
+    | IAnom :: ?pats => let H := iFresh in iIntro H; go pats
+    | IAlwaysElim (IName ?H) :: ?pats => iIntro #H; go pats
+    | IName ?H :: ?pats => iIntro H; go pats
+    | IPureIntro :: ?pats => iPureIntro; go pats
+    | IAlwaysIntro :: ?pats => iAlways; go pats
+    | ILaterIntro :: ?pats => iNext; go pats
+    | IVsIntro :: ?pats => iVsIntro; go pats
+    | ISimpl :: ?pats => simpl; go pats
     | IForall :: ?pats => repeat iIntroForall; go pats
     | IAll :: ?pats => repeat (iIntroForall || iIntro); go pats
-    | ISimpl :: ?pats => simpl; go pats
-    | IAlways :: ?pats => iAlways; go pats
-    | INext :: ?pats => iNext; go pats
     | IClear ?cpats :: ?pats =>
        let rec clr cpats :=
          match cpats with
@@ -576,19 +658,13 @@ Tactic Notation "iIntros" constr(pat) :=
          | (false,?H) :: ?cpats => iClear H; clr cpats
          | (true,?H) :: ?cpats => iFrame H; clr cpats
          end in clr cpats
-    | IPersistent (IName ?H) :: ?pats => iIntro #H; go pats
-    | IName ?H :: ?pats => iIntro H; go pats
-    | IPersistent IAnom :: ?pats => let H := iFresh in iIntro #H; go pats
-    | IAnom :: ?pats => let H := iFresh in iIntro H; go pats
-    | IAnomPure :: ?pats => iIntro (?); go pats
-    | IPersistent ?pat :: ?pats =>
+    | IAlwaysElim ?pat :: ?pats =>
        let H := iFresh in iIntro #H; iDestructHyp H as pat; go pats
     | ?pat :: ?pats =>
        let H := iFresh in iIntro H; iDestructHyp H as pat; go pats
-    | _ => fail "iIntro: failed with" pats
     end
   in let pats := intro_pat.parse pat in try iProof; go pats.
-Tactic Notation "iIntros" := iIntros "**".
+Tactic Notation "iIntros" := iIntros [IAll].
 
 Tactic Notation "iIntros" "(" simple_intropattern(x1) ")" :=
   try iProof; iIntro ( x1 ).
@@ -647,7 +723,7 @@ Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
   iIntros ( x1 x2 x3 x4 x5 x6 x7 x8 ); iIntros p.
 
 (** * Destruct tactic *)
-Tactic Notation "iDestructHelp" open_constr(lem) "as" tactic(tac) :=
+Tactic Notation "iDestructCore" open_constr(lem) "as" tactic(tac) :=
   let intro_destruct n :=
     let rec go n' :=
       lazymatch n' with
@@ -663,49 +739,48 @@ Tactic Notation "iDestructHelp" open_constr(lem) "as" tactic(tac) :=
   | string => tac lem
   | iTrm =>
      lazymatch lem with
-     | @iTrm string ?H _ hnil ?pat =>
-        iSpecializePat H pat; last tac H
-     | _ => let H := iFresh in iPoseProof lem as H; last tac H; try apply _
+     | @iTrm string ?H _ hnil ?pat => iSpecializePat H pat; last (tac H)
+     | _ => iPoseProofCore lem as tac
      end
-  | _ => let H := iFresh in iPoseProof lem as H; last tac H; try apply _
+  | _ => iPoseProofCore lem as tac
   end.
 
-Tactic Notation "iDestruct" open_constr(H) "as" constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1) ")"
+Tactic Notation "iDestruct" open_constr(lem) "as" constr(pat) :=
+  iDestructCore lem as (fun H => iDestructHyp H as pat).
+Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1) ")"
     constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 ) pat).
+Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
     simple_intropattern(x2) ")" constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 ) pat).
+Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
     simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 x3 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 x3 ) pat).
+Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
     simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
     constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 x3 x4 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 x3 x4 ) pat).
+Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
     simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
     simple_intropattern(x5) ")" constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 ) pat).
+Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
     simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
     simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 ) pat).
+Tactic Notation "iDestruct" 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) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 ) pat).
-Tactic Notation "iDestruct" open_constr(H) "as" "(" simple_intropattern(x1)
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 ) pat).
+Tactic Notation "iDestruct" 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) :=
-  iDestructHelp H as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat).
+  iDestructCore lem as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat).
 
-Tactic Notation "iDestruct" open_constr(H) "as" "%" simple_intropattern(pat) :=
-  let Htmp := iFresh in iDestruct H as Htmp; last iPure Htmp as pat.
+Tactic Notation "iDestruct" open_constr(lem) "as" "%" simple_intropattern(pat) :=
+  iDestructCore lem as (fun H => iPure H as pat).
 
 (* This is pretty ugly, but without Ltac support for manipulating lists of
 idents I do not know how to do this better. *)
@@ -762,7 +837,7 @@ Tactic Notation "iAssert" open_constr(Q) "with" constr(Hs) "as" constr(pat) :=
      eapply tac_assert with _ _ _ lr Hs H Q _; (* (js:=Hs) (j:=H) (P:=Q) *)
        [match k with
         | GoalStd => apply into_assert_default
-        | GoalPvs => apply _ || fail "iAssert: cannot generate pvs goal"
+        | GoalVs => apply _ || fail "iAssert: cannot generate view shifted goal"
         end
        |env_cbv; reflexivity || fail "iAssert:" Hs "not found"
        |env_cbv; reflexivity|
@@ -781,38 +856,80 @@ Local Ltac iRewriteFindPred :=
      match goal with |- (∀ y, @?Ψ y ⊣⊢ _) => unify Φ Ψ; reflexivity end
   end.
 
-Local Tactic Notation "iRewriteCore" constr(lr) open_constr(t) :=
-  let Heq := iFresh in iPoseProof t as Heq; last (
-  eapply (tac_rewrite _ Heq _ _ lr);
-    [env_cbv; reflexivity || fail "iRewrite:" Heq "not found"
-    |let P := match goal with |- ?P ⊢ _ => P end in
-     reflexivity || fail "iRewrite:" Heq ":" P "not an equality"
-    |iRewriteFindPred
-    |intros ??? ->; reflexivity|lazy beta; iClear Heq]).
-
-Tactic Notation "iRewrite" open_constr(t) := iRewriteCore false t.
-Tactic Notation "iRewrite" "-" open_constr(t) := iRewriteCore true t.
-
-Local Tactic Notation "iRewriteCore" constr(lr) open_constr(t) "in" constr(H) :=
-  let Heq := iFresh in iPoseProof t as Heq; last (
-  eapply (tac_rewrite_in _ Heq _ _ H _ _ lr);
-    [env_cbv; reflexivity || fail "iRewrite:" Heq "not found"
-    |env_cbv; reflexivity || fail "iRewrite:" H "not found"
-    |let P := match goal with |- ?P ⊢ _ => P end in
-     reflexivity || fail "iRewrite:" Heq ":" P "not an equality"
-    |iRewriteFindPred
-    |intros ??? ->; reflexivity
-    |env_cbv; reflexivity|lazy beta; iClear Heq]).
-
-Tactic Notation "iRewrite" open_constr(t) "in" constr(H) :=
-  iRewriteCore false t in H.
-Tactic Notation "iRewrite" "-" open_constr(t) "in" constr(H) :=
-  iRewriteCore true t in H.
+Local Tactic Notation "iRewriteCore" constr(lr) open_constr(lem) :=
+  iPoseProofCore lem as (fun Heq =>
+    eapply (tac_rewrite _ Heq _ _ lr);
+      [env_cbv; reflexivity || fail "iRewrite:" Heq "not found"
+      |let P := match goal with |- ?P ⊢ _ => P end in
+       reflexivity || fail "iRewrite:" P "not an equality"
+      |iRewriteFindPred
+      |intros ??? ->; reflexivity|lazy beta; iClear Heq]).
+
+Tactic Notation "iRewrite" open_constr(lem) := iRewriteCore false lem.
+Tactic Notation "iRewrite" "-" open_constr(lem) := iRewriteCore true lem.
+
+Local Tactic Notation "iRewriteCore" constr(lr) open_constr(lem) "in" constr(H) :=
+  iPoseProofCore lem as (fun Heq =>
+    eapply (tac_rewrite_in _ Heq _ _ H _ _ lr);
+      [env_cbv; reflexivity || fail "iRewrite:" Heq "not found"
+      |env_cbv; reflexivity || fail "iRewrite:" H "not found"
+      |let P := match goal with |- ?P ⊢ _ => P end in
+       reflexivity || fail "iRewrite:" P "not an equality"
+      |iRewriteFindPred
+      |intros ??? ->; reflexivity
+      |env_cbv; reflexivity|lazy beta; iClear Heq]).
+
+Tactic Notation "iRewrite" open_constr(lem) "in" constr(H) :=
+  iRewriteCore false lem in H.
+Tactic Notation "iRewrite" "-" open_constr(lem) "in" constr(H) :=
+  iRewriteCore true lem in H.
 
 Ltac iSimplifyEq := repeat (
   iMatchGoal ltac:(fun H P => match P with (_ = _)%I => iDestruct H as %? end)
   || simplify_eq/=).
 
+(** * View shifts *)
+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.
@@ -823,11 +940,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..90b01fb5c434365d6abef7081b09ff403f679f74 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_except_last_wp E e Φ : IsExceptLast (WP e @ E {{ Φ }}).
+Proof. by rewrite /IsExceptLast -{2}pvs_wp -except_last_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/atomic.v b/tests/atomic.v
new file mode 100644
index 0000000000000000000000000000000000000000..a26f9ee60e422adbe23299f222e75ee95aa913a5
--- /dev/null
+++ b/tests/atomic.v
@@ -0,0 +1,148 @@
+From iris.program_logic Require Export hoare weakestpre pviewshifts ownership.
+From iris.algebra Require Import upred_big_op.
+From iris.prelude Require Export coPset.
+From iris.proofmode Require Import tactics pviewshifts weakestpre.
+Import uPred.
+
+Section atomic.
+  Context `{irisG Λ Σ} {A: Type}.
+
+  (* logically atomic triple: <x, α> e @ E_i, E_o <v, β x v> *)
+  Definition atomic_triple
+             (α: A → iProp Σ)
+             (β: A → val _ → iProp Σ)
+             (Ei Eo: coPset)
+             (e: expr _) : iProp Σ :=
+    (∀ P Q, (P ={Eo, Ei}=> ∃ x:A,
+                                α x ★
+                                ((α x ={Ei, Eo}=★ P) ∧
+                                 (∀ v, β x v ={Ei, Eo}=★ Q x v))
+            ) -★ {{ P }} e @ ⊤ {{ v, (∃ x: A, Q x v) }})%I.
+
+  (* Weakest-pre version of the above one. Also weaker in some sense *)
+  Definition atomic_triple_wp
+             (α: A → iProp Σ)
+             (β: A → val _ → iProp Σ)
+             (Ei Eo: coPset)
+             (e: expr _) : iProp Σ :=
+    (∀ P Q, (P ={Eo, Ei}=> ∃ x,
+                              α x ★
+                              ((α x ={Ei, Eo}=★ P) ∧
+                               (∀ v, β x v ={Ei, Eo}=★ Q x v))
+            ) -★ P -★ WP e @ ⊤ {{ v, (∃ x, Q x v) }})%I.
+
+  Lemma atomic_triple_weaken α β Ei Eo e:
+    atomic_triple α β Ei Eo e ⊢ atomic_triple_wp α β Ei Eo e.
+  Proof.
+    iIntros "H". iIntros (P Q) "Hvs Hp".
+    by iApply ("H" $! P Q with "Hvs").
+  Qed.
+
+  Arguments atomic_triple {_} _ _ _ _.
+End atomic.
+
+From iris.heap_lang Require Export lang proofmode notation.
+From iris.proofmode Require Import invariants.
+
+Section incr.
+  Context `{!heapG Σ} (N : namespace).
+
+  Definition incr: val :=
+    rec: "incr" "l" :=
+       let: "oldv" := !"l" in
+       if: CAS "l" "oldv" ("oldv" + #1)
+         then "oldv" (* return old value if success *)
+         else "incr" "l".
+
+  Global Opaque incr.
+
+  (* TODO: Can we have a more WP-style definition and avoid the equality? *)
+  Definition incr_triple (l: loc) :=
+    atomic_triple (fun (v: Z) => l ↦ #v)%I
+                  (fun v ret => ret = #v ★ l ↦ #(v + 1))%I
+                  (nclose heapN)
+                  ⊤
+                  (incr #l).
+
+  Lemma incr_atomic_spec: ∀ (l: loc), heapN ⊥ N → heap_ctx ⊢ incr_triple l.
+  Proof.
+    iIntros (l HN) "#?".
+    rewrite /incr_triple.
+    rewrite /atomic_triple.
+    iIntros (P Q) "#Hvs".
+    iLöb as "IH".
+    iIntros "!# HP".
+    wp_rec.
+    wp_bind (! _)%E.
+    iVs ("Hvs" with "HP") as (x) "[Hl [Hvs' _]]".
+    wp_load.
+    iVs ("Hvs'" with "Hl") as "HP".
+    iVsIntro. wp_let. wp_bind (CAS _ _ _). wp_op.
+    iVs ("Hvs" with "HP") as (x') "[Hl Hvs']".
+    destruct (decide (x = x')).
+    - subst.
+      iDestruct "Hvs'" as "[_ Hvs']".
+      iSpecialize ("Hvs'" $! #x').
+      wp_cas_suc.
+      iVs ("Hvs'" with "[Hl]") as "HQ"; first by iFrame.
+      iVsIntro. wp_if. iVsIntro. by iExists x'.
+    - iDestruct "Hvs'" as "[Hvs' _]".
+      wp_cas_fail.
+      iVs ("Hvs'" with "[Hl]") as "HP"; first by iFrame.
+      iVsIntro. wp_if. by iApply "IH".
+  Qed.
+End incr.
+
+From iris.heap_lang.lib Require Import par.
+
+Section user.
+  Context `{!heapG Σ, !spawnG Σ} (N : namespace).
+
+  Definition incr_2 : val :=
+    λ: "x",
+       let: "l" := ref "x" in
+       incr "l" || incr "l";;
+       !"l".
+
+  (* prove that incr is safe w.r.t. data race. TODO: prove a stronger post-condition *)
+  Lemma incr_2_safe:
+    ∀ (x: Z), heapN ⊥ N -> heap_ctx ⊢ WP incr_2 #x {{ _, True }}.
+  Proof.
+    iIntros (x HN) "#Hh".
+    rewrite /incr_2.
+    wp_let.
+    wp_alloc l as "Hl".
+    iVs (inv_alloc N _ (∃x':Z, l ↦ #x')%I with "[Hl]") as "#?"; first eauto.
+    wp_let.
+    wp_bind (_ || _)%E.
+    iApply (wp_par (λ _, True%I) (λ _, True%I)).
+    iFrame "Hh".
+    (* prove worker triple *)
+    iDestruct (incr_atomic_spec N l with "Hh") as "Hincr"=>//.
+    rewrite /incr_triple /atomic_triple.
+    iSpecialize ("Hincr"  $! True%I (fun _ _ => True%I) with "[]").
+    - iIntros "!# _".
+      (* open the invariant *)
+      iInv N as (x') ">Hl'" "Hclose".
+      (* mask magic *)
+      iApply pvs_intro'.
+      { apply ndisj_subseteq_difference; auto. }
+      iIntros "Hvs".
+      iExists x'.
+      iFrame "Hl'".
+      iSplit.
+      + (* provide a way to rollback *)
+        iIntros "Hl'".
+        iVs "Hvs". iVs ("Hclose" with "[Hl']"); eauto.
+      + (* provide a way to commit *)
+        iIntros (v) "[Heq Hl']".
+        iVs "Hvs". iVs ("Hclose" with "[Hl']"); eauto.
+    - iDestruct "Hincr" as "#HIncr".
+      iSplitL; [|iSplitL]; try (iApply wp_wand_r;iSplitL; [by iApply "HIncr"|auto]).
+      iIntros (v1 v2) "_ !>".
+      wp_seq.
+      iInv N as (x') ">Hl" "Hclose".
+      wp_load.
+      iApply "Hclose". eauto.
+  Qed.
+End user.
diff --git a/tests/barrier_client.v b/tests/barrier_client.v
index ae0d5df08505a7e75841cb1f226ebb37d937ce72..dac3e09c35a59dce6faa957713e42acca4a7f0ff 100644
--- a/tests/barrier_client.v
+++ b/tests/barrier_client.v
@@ -1,8 +1,9 @@
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
 From iris.heap_lang.lib.barrier Require Import proof.
 From iris.heap_lang Require Import par.
 From iris.program_logic Require Import auth sts saved_prop hoare ownership.
-From iris.heap_lang Require Import proofmode.
-Import uPred.
+From iris.heap_lang Require Import adequacy proofmode.
 
 Definition worker (n : Z) : val :=
   λ: "b" "y", wait "b" ;; !"y" #n.
@@ -14,10 +15,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).
@@ -47,26 +47,26 @@ Section client.
       wp_store. iApply signal_spec; iFrame "Hs"; iSplit; [|done].
       iExists _; iSplitL; [done|]. iAlways; iIntros (n). wp_let. by wp_op.
     - (* The two spawned threads, the waiters. *)
-      iSplitL; [|by iIntros (_ _) "_ >"].
+      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 (_ _) "_ >"]];
+      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 Σ).
 
-  Lemma client_safe_closed σ : {{ ownP σ : iProp }} client {{ v, True }}.
-  Proof.
-    iIntros "! Hσ".
-    iPvs (heap_alloc with "Hσ") as (h) "[#Hh _]"; first done.
-    iApply (client_safe (nroot .@ "barrier")); auto with ndisj.
-  Qed.
+Let Σ : gFunctors := #[ heapΣ ; barrierΣ ; spawnΣ ].
+
+Lemma client_adequate σ : adequate client σ (λ _, True).
+Proof.
+  apply (heap_adequacy Σ)=> ?.
+  apply (client_safe (nroot .@ "barrier")); auto with ndisj.
+Qed.
 
-  Print Assumptions client_safe_closed.
 End ClosedProofs.
+
+Print Assumptions client_adequate.
diff --git a/tests/counter.v b/tests/counter.v
new file mode 100644
index 0000000000000000000000000000000000000000..c5e84d11d76f6f8bcbcd1098e63bc885a3fa5825
--- /dev/null
+++ b/tests/counter.v
@@ -0,0 +1,138 @@
+(* This file contains a formalization of the monotone counter, but with an
+explicit contruction of the monoid, as we have also done in the proof mode
+paper. A version that uses the authoritative monoid and natural number monoid
+under max can be found in `heap_lang/lib/counter.v`. *)
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
+From iris.program_logic Require Export hoare.
+From iris.proofmode Require Import invariants tactics.
+From iris.heap_lang Require Import proofmode notation.
+Import uPred.
+
+Definition newcounter : val := λ: <>, ref #0.
+Definition inc : val :=
+  rec: "inc" "l" :=
+    let: "n" := !"l" in
+    if: CAS "l" "n" (#1 + "n") then #() else "inc" "l".
+Definition read : val := λ: "l", !"l".
+Global Opaque newcounter inc read.
+
+(** The CMRA we need. *)
+Inductive M := Auth : nat → M | Frag : nat → M | Bot.
+
+Section M.
+  Arguments cmra_op _ !_ !_/.
+  Arguments op _ _ !_ !_/.
+  Arguments core _ _ !_/.
+
+  Canonical Structure M_C : cofeT := leibnizC M.
+
+  Instance M_valid : Valid M := λ x, x ≠ Bot.
+  Instance M_op : Op M := λ x y,
+    match x, y with
+    | Auth n, Frag j | Frag j, Auth n => if decide (j ≤ n)%nat then Auth n else Bot
+    | Frag i, Frag j => Frag (max i j)
+    | _, _ => Bot
+    end.
+  Instance M_pcore : PCore M := λ x,
+    Some match x with Auth j | Frag j => Frag j | _ => Bot end.
+  Instance M_empty : Empty M := Frag 0.
+
+  Definition M_ra_mixin : RAMixin M.
+  Proof.
+    apply ra_total_mixin; try solve_proper || eauto.
+    - intros [n1|i1|] [n2|i2|] [n3|i3|];
+        repeat (simpl; case_decide); f_equal/=; lia.
+    - intros [n1|i1|] [n2|i2|]; repeat (simpl; case_decide); f_equal/=; lia.
+    - intros [n|i|]; repeat (simpl; case_decide); f_equal/=; lia.
+    - by intros [n|i|].
+    - intros [n1|i1|] y [[n2|i2|] ?]; exists (core y); simplify_eq/=;
+        repeat (simpl; case_decide); f_equal/=; lia.
+    - intros [n1|i1|] [n2|i2|]; simpl; by try case_decide.
+  Qed.
+  Canonical Structure M_R : cmraT := discreteR M M_ra_mixin.
+  Definition M_ucmra_mixin : UCMRAMixin M.
+  Proof.
+    split; try (done || apply _).
+    intros [?|?|]; simpl; try case_decide; f_equal/=; lia.
+  Qed.
+  Canonical Structure M_UR : ucmraT := discreteUR M M_ra_mixin M_ucmra_mixin.
+
+  Global Instance frag_persistent n : Persistent (Frag n).
+  Proof. by constructor. Qed.
+  Lemma auth_frag_valid j n : ✓ (Auth n ⋅ Frag j) → (j ≤ n)%nat.
+  Proof. simpl. case_decide. done. by intros []. Qed.
+  Lemma auth_frag_op (j n : nat) : (j ≤ n)%nat → Auth n = Auth n ⋅ Frag j.
+  Proof. intros. by rewrite /= decide_True. Qed.
+
+  Lemma M_update n : Auth n ~~> Auth (S n).
+  Proof.
+    apply cmra_discrete_update=>-[m|j|] /= ?; repeat case_decide; done || lia.
+  Qed.
+End M.
+
+Class counterG Σ := CounterG { counter_tokG :> inG Σ M_UR }.
+Definition counterΣ : gFunctors := #[GFunctor (constRF M_UR)].
+Instance subG_counterΣ {Σ} : subG counterΣ Σ → counterG Σ.
+Proof. intros [?%subG_inG _]%subG_inv. split; apply _. Qed.
+
+Section proof.
+Context `{!heapG Σ, !counterG Σ}.
+Implicit Types l : loc.
+
+Definition I (γ : gname) (l : loc) : iProp Σ :=
+  (∃ c : nat, l ↦ #c ★ own γ (Auth c))%I.
+
+Definition C (l : loc) (n : nat) : iProp Σ :=
+  (∃ N γ, heapN ⊥ N ∧ heap_ctx ∧ inv N (I γ l) ∧ own γ (Frag n))%I.
+
+(** The main proofs. *)
+Global Instance C_persistent l n : PersistentP (C l n).
+Proof. apply _. Qed.
+
+Lemma newcounter_spec N :
+  heapN ⊥ N →
+  heap_ctx ⊢ {{ True }} newcounter #() {{ v, ∃ l, v = #l ∧ C l 0 }}.
+Proof.
+  iIntros (?) "#Hh !# _ /=". rewrite /newcounter /=. wp_seq. wp_alloc l as "Hl".
+  iVs (own_alloc (Auth 0)) as (γ) "Hγ"; first done.
+  rewrite (auth_frag_op 0 0) //; iDestruct "Hγ" as "[Hγ Hγf]".
+  iVs (inv_alloc N _ (I γ l) with "[Hl Hγ]") as "#?".
+  { iIntros "!>". iExists 0%nat. by iFrame. }
+  iVsIntro. rewrite /C; eauto 10.
+Qed.
+
+Lemma inc_spec l n :
+  {{ C l n }} inc #l {{ v, v = #() ∧ C l (S n) }}.
+Proof.
+  iIntros "!# Hl /=". iLöb as "IH". wp_rec.
+  iDestruct "Hl" as (N γ) "(% & #Hh & #Hinv & Hγf)".
+  wp_bind (! _)%E; iInv N as (c) "[Hl Hγ]" "Hclose".
+  wp_load. iVs ("Hclose" with "[Hl Hγ]"); [iNext; iExists c; by iFrame|].
+  iVsIntro. wp_let. wp_op.
+  wp_bind (CAS _ _ _). iInv N as (c') ">[Hl Hγ]" "Hclose".
+  destruct (decide (c' = c)) as [->|].
+  - iCombine "Hγ" "Hγf" as "Hγ".
+    iDestruct (own_valid with "#Hγ") as %?%auth_frag_valid; rewrite -auth_frag_op //.
+    iVs (own_update with "Hγ") as "Hγ"; first apply M_update.
+    rewrite (auth_frag_op (S n) (S c)); last lia; iDestruct "Hγ" as "[Hγ Hγf]".
+    wp_cas_suc. iVs ("Hclose" with "[Hl Hγ]").
+    { iNext. iExists (S c). rewrite Nat2Z.inj_succ Z.add_1_l. by iFrame. }
+    iVsIntro. wp_if. iVsIntro; rewrite {3}/C; eauto 10.
+  - wp_cas_fail; first (intros [=]; abstract omega).
+    iVs ("Hclose" with "[Hl Hγ]"); [iNext; iExists c'; by iFrame|].
+    iVsIntro. wp_if. iApply ("IH" with "[Hγf]"). rewrite {3}/C; eauto 10.
+Qed.
+
+Lemma read_spec l n :
+  {{ C l n }} read #l {{ v, ∃ m : nat, ■ (v = #m ∧ n ≤ m) ∧ C l m }}.
+Proof.
+  iIntros "!# Hl /=". iDestruct "Hl" as (N γ) "(% & #Hh & #Hinv & Hγf)".
+  rewrite /read /=. wp_let. iInv N as (c) "[Hl Hγ]" "Hclose". wp_load.
+  iDestruct (own_valid γ (Frag n ⋅ Auth c) with "[#]") as % ?%auth_frag_valid.
+  { iApply own_op. by iFrame. }
+  rewrite (auth_frag_op c c); last lia; iDestruct "Hγ" as "[Hγ Hγf']".
+  iVs ("Hclose" with "[Hl Hγ]"); [iNext; iExists c; by iFrame|].
+  iVsIntro; rewrite /C; eauto 10 with omega.
+Qed.
+End proof.
diff --git a/tests/heap_lang.v b/tests/heap_lang.v
index b259757ffa4bd204caa68a5f4ef2a24e58138166..fec4903222f412abc0f6832a478be7ac4ae56b8f 100644
--- a/tests/heap_lang.v
+++ b/tests/heap_lang.v
@@ -1,25 +1,14 @@
 (** This file is essentially a bunch of testcases. *)
-From iris.program_logic Require Import ownership hoare auth.
+From iris.program_logic Require Export weakestpre hoare.
+From iris.heap_lang Require Export lang.
+From iris.heap_lang Require Import adequacy.
+From iris.program_logic Require Import ownership.
 From iris.heap_lang Require Import proofmode notation.
-Import uPred.
-
-Section LangTests.
-  Definition add : expr := (#21 + #21)%E.
-  Goal ∀ σ, head_step add σ (#42) σ None.
-  Proof. intros; do_head_step done. Qed.
-  Definition rec_app : expr := ((rec: "f" "x" := "f" "x") #0)%E.
-  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.
-  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 +60,9 @@ 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 Σ).
-
-  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.
-    by iApply heap_e_spec; first solve_ndisj.
-  Qed.
-End ClosedProofs.
+Lemma heap_e_adequate σ : adequate heap_e σ (λ v, v = #2).
+Proof. eapply (heap_adequacy heapΣ)=> ?. by apply heap_e_spec. Qed.
diff --git a/tests/joining_existentials.v b/tests/joining_existentials.v
index 36077d52ee263ed97797971fc0777c21f8f99349..e06e47214277a05745f7981127303fd4ba54828a 100644
--- a/tests/joining_existentials.v
+++ b/tests/joining_existentials.v
@@ -1,22 +1,22 @@
-From iris.algebra Require Import csum.
-From iris.program_logic Require Import hoare.
+From iris.program_logic Require Export weakestpre hoare.
+From iris.heap_lang Require Export lang.
+From iris.algebra Require Import excl agree csum.
 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).
-Definition oneShotGF (F : cFunctor) : gFunctor :=
-  GFunctor (csumRF (exclRF unitC) (agreeRF (â–¶ F))).
-Instance inGF_oneShotG  `{inGF Λ Σ (oneShotGF F)} : oneShotG Λ Σ F.
-Proof. apply: inGF_inG. Qed.
+Class oneShotG (Σ : gFunctors) (F : cFunctor) :=
+  one_shot_inG :> inG Σ (one_shotR Σ F).
+Definition oneShotΣ (F : cFunctor) : gFunctors :=
+  #[ GFunctor (csumRF (exclRF unitC) (agreeRF (â–¶ F))) ].
+Instance subG_oneShotΣ {Σ F} : subG (oneShotΣ F) Σ → oneShotG Σ F.
+Proof. apply subG_inG. Qed.
 
 Definition client eM eW1 eW2 : expr :=
   let: "b" := newbarrier #() in
@@ -24,25 +24,24 @@ 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.
-  iIntros "[Hl #He]". wp_apply wait_spec; iFrame "Hl".
+  iIntros "[Hl #He]". wp_apply wait_spec; simpl; iFrame "Hl".
   iDestruct 1 as (x) "[#Hγ Hx]".
   wp_seq. iApply wp_wand_l. iSplitR; [|by iApply "He"].
   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}.
 
@@ -56,9 +55,9 @@ Lemma Q_res_join γ : barrier_res γ Ψ1 ★ barrier_res γ Ψ2 ⊢ ▷ barrier_
 Proof.
   iIntros "[Hγ Hγ']";
   iDestruct "Hγ" as (x) "[#Hγ Hx]"; iDestruct "Hγ'" as (x') "[#Hγ' Hx']".
-  iAssert (▷ (x ≡ x'))%I as "Hxx" .
+  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,20 +75,20 @@ 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).
   wp_let. wp_apply (wp_par  (λ _, True)%I workers_post); iFrame "Hh".
   iSplitL "HP Hs Hγ"; [|iSplitL "Hr"].
-  - wp_focus eM. iApply wp_wand_l; iSplitR "HP"; [|by iApply "He"].
+  - wp_bind 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..4ab90f4c54e09c4e985c21c8c9415ac464f98a25 100644
--- a/tests/list_reverse.v
+++ b/tests/list_reverse.v
@@ -1,14 +1,14 @@
 (** Correctness of in-place list reversal *)
+From iris.program_logic Require Export weakestpre hoare.
+From iris.heap_lang Require Export lang.
 From iris.proofmode Require Export tactics.
-From iris.program_logic Require Export hoare.
 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 +26,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 +42,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 faaac93142802f4d9b518ff50fa4440cf39bbc83..3e610b65d9ad5dd6d1549dc8575c201d0b16df37 100644
--- a/tests/one_shot.v
+++ b/tests/one_shot.v
@@ -1,8 +1,8 @@
-From iris.algebra Require Import dec_agree csum.
-From iris.program_logic Require Import hoare.
+From iris.program_logic Require Export weakestpre hoare.
+From iris.heap_lang Require Export lang.
+From iris.algebra Require Import excl dec_agree csum.
 From iris.heap_lang Require Import assert proofmode notation.
-From iris.proofmode Require Import invariants ghost_ownership.
-Import uPred.
+From iris.proofmode Require Import invariants.
 
 Definition one_shot_example : val := λ: <>,
   let: "x" := ref NONE in (
@@ -24,73 +24,76 @@ 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.
-Definition one_shotGF : gFunctorList := [GFunctor (constRF one_shotR)].
-Instance inGF_one_shotG Σ : inGFs heap_lang Σ one_shotGF → one_shotG Σ.
-Proof. intros [? _]; apply: inGF_inG. Qed.
+Class one_shotG Σ := { one_shot_inG :> inG Σ one_shotR }.
+Definition one_shotΣ : gFunctors := #[GFunctor (constRF one_shotR)].
+Instance subG_one_shotΣ {Σ} : subG one_shotΣ Σ → one_shotG Σ.
+Proof. intros [?%subG_inG _]%subG_inv. split; apply _. 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)
   ⊢ WP one_shot_example #() {{ Φ }}.
 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.
+  rewrite /one_shot_example /=. wp_seq. wp_alloc l as "Hl". wp_let.
+  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.
-  - 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γ".
+  iVsIntro. iApply "Hf"; iSplit.
+  - iIntros (n) "!#". wp_let.
+    iInv N as ">[[Hl Hγ]|H]" "Hclose"; 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_bind (! _)%E.
+    iInv N as ">Hγ" "Hclose".
     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γ]".
-    { iCombine "Hγ" "Hγ'" as "Hγ". by iDestruct (@own_valid with "Hγ") as %?. }
-    wp_load; iPvsIntro.
+    wp_match. wp_bind (! _)%E.
+    iInv N as ">[[Hl Hγ]|H]" "Hclose"; last iDestruct "H" as (m') "[Hl Hγ]".
+    { iCombine "Hγ" "Hγ'" as "Hγ". by iDestruct (own_valid with "Hγ") as %?. }
+    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.
+    iDestruct (own_valid with "#Hγ") as %[=->]%dec_agree_op_inv.
+    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 }}) ★
       {{ True }} Snd ff #() {{ g, {{ True }} g #() {{ _, True }} }}
     }}.
 Proof.
-  iIntros "#? ! _". iApply wp_one_shot. iSplit; first done.
+  iIntros "#? !# _". iApply wp_one_shot. iSplit; first done.
   iIntros (f1 f2) "[#Hf1 #Hf2]"; iSplit.
-  - iIntros (n) "! _". wp_proj. iApply "Hf1".
-  - iIntros "! _". wp_proj.
-    iApply wp_wand_l; iFrame "Hf2". by iIntros (v) "#? ! _".
+  - iIntros (n) "!# _". wp_proj. iApply "Hf1".
+  - iIntros "!# _". wp_proj.
+    iApply wp_wand_l; iFrame "Hf2". by iIntros (v) "#? !# _".
 Qed.
 End proof.
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..49e2160cc8642d273d41f5d4b24a6293ff2db358 100644
--- a/tests/proofmode.v
+++ b/tests/proofmode.v
@@ -20,7 +20,7 @@ Lemma demo_1 (M : ucmraT) (P1 P2 P3 : nat → uPred M) :
     ▷ (∀ n m : nat, P1 n → □ ((True ∧ P2 n) → □ (n = n ↔ P3 n))) -★
     ▷ (x = 0) ∨ ∃ x z, ▷ P3 (x + z) ★ uPred_ownM b ★ uPred_ownM (core b)).
 Proof.
-  iIntros (i [|j] a b ?) "! [Ha Hb] H1 #H2 H3"; setoid_subst.
+  iIntros (i [|j] a b ?) "!# [Ha Hb] H1 #H2 H3"; setoid_subst.
   { iLeft. by iNext. }
   iRight.
   iDestruct "H1" as (z1 z2 c) "(H1&_&#Hc)".
@@ -81,26 +81,20 @@ Proof.
   by iIntros "# _".
 Qed.
 
+Lemma demo_7 (M : ucmraT) (P Q1 Q2 : uPred M) : P ★ (Q1 ∧ Q2) ⊢ P ★ Q1.
+Proof. iIntros "[H1 [H2 _]]". by iFrame. Qed.
+
 Section iris.
-  Context {Λ : language} {Σ : iFunctor}.
+  Context `{irisG Λ Σ}.
   Implicit Types E : coPset.
-  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.
+  Implicit Types P Q : iProp Σ.
 
   Lemma demo_8 N E P Q R :
     nclose N ⊆ E →
     (True -★ P -★ inv N Q -★ True -★ R) ⊢ P -★ ▷ Q ={E}=★ R.
   Proof.
     iIntros (?) "H HP HQ".
-    iApply ("H" with "[#] HP |==>[HQ] |==>").
+    iApply ("H" with "[#] HP ==>[HQ] ==>").
     - done.
     - by iApply inv_alloc.
     - done.
diff --git a/tests/tree_sum.v b/tests/tree_sum.v
index 6d07b098cde1be59189a4e083999129f6db43a54..73a023c116a2b864379247a04e715298e9ce5ce6 100644
--- a/tests/tree_sum.v
+++ b/tests/tree_sum.v
@@ -1,3 +1,5 @@
+From iris.program_logic Require Export weakestpre.
+From iris.heap_lang Require Export lang.
 From iris.proofmode Require Export tactics.
 From iris.heap_lang Require Import proofmode notation.
 
@@ -5,7 +7,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 +35,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 {{ Φ }}.
@@ -58,7 +60,7 @@ Lemma sum_wp `{!heapG Σ} v t Φ :
   heap_ctx ★ is_tree v t ★ (is_tree v t -★ Φ #(sum t))
   ⊢ WP sum' v {{ Φ }}.
 Proof.
-  iIntros "(#Hh & Ht & HΦ)". rewrite /sum'.
+  iIntros "(#Hh & Ht & HΦ)". rewrite /sum' /=.
   wp_let. wp_alloc l as "Hl". wp_let.
   wp_apply sum_loop_wp; iFrame "Hh Ht Hl".
   rewrite Z.add_0_r.