From c79e7e716af130077fadc0535a85cc9169d47730 Mon Sep 17 00:00:00 2001
From: William Mansky <wmansky@seas.upenn.edu>
Date: Wed, 3 Jan 2018 16:17:16 -0500
Subject: [PATCH] Merge, started kvnode example.

---
 LICENSE                                  |   2 +-
 _CoqProject                              |   8 +
 opam                                     |   2 +-
 theories/adequacy.v                      |   4 +-
 theories/base/accessors.v                |  10 +-
 theories/base/alloc.v                    |   2 +-
 theories/base/at_cas.v                   |   5 +-
 theories/base/at_fai.v                   |  10 +-
 theories/base/at_read.v                  |   2 +-
 theories/base/at_write.v                 |  10 +-
 theories/base/atomic.v                   |   2 +-
 theories/base/base.v                     |   2 +-
 theories/base/dealloc.v                  |   6 +-
 theories/base/ghosts.v                   |  23 +-
 theories/base/helpers.v                  |   2 +-
 theories/base/na_read.v                  |   5 +-
 theories/base/na_write.v                 |   6 +-
 theories/bigop.v                         |   2 +-
 theories/blocks.v                        |   2 +-
 theories/blocks_generic.v                |  16 +-
 theories/derived.v                       |   2 +-
 theories/derived_vProp.v                 |   2 +-
 theories/escrows.v                       |  18 +-
 theories/examples/circ_buffer.v          |   6 +-
 theories/examples/for_loop.v             |   4 +-
 theories/examples/hashtable.v            | 161 +--------
 theories/examples/hist_protocol.v        | 131 ++++++++
 theories/examples/kvnode.v               | 332 +++++++++++++++++++
 theories/examples/message_passing.v      | 151 +++++----
 theories/examples/message_passing_base.v |   2 +-
 theories/examples/msqueue.v              |  16 +-
 theories/examples/protocols.v            |   2 +-
 theories/examples/rcu.v                  |  50 +--
 theories/examples/rcu_data.v             |  10 +-
 theories/examples/repeat.v               |   8 +-
 theories/examples/sc_stack.v             |   6 +-
 theories/examples/snapshot.v             |  29 ++
 theories/examples/spin_lock.v            |   6 +-
 theories/examples/ticket_lock.v          |  28 +-
 theories/examples/tstack.v               |  14 +-
 theories/examples/unit_token.v           |   2 +-
 theories/fork.v                          |  24 ++
 theories/fractor.v                       |   2 +-
 theories/gps/cas.v                       |  69 ++--
 theories/gps/fai.v                       |   5 +-
 theories/gps/fractional.v                |  59 ++--
 theories/gps/init.v                      |  16 +-
 theories/gps/plain.v                     |  87 ++---
 theories/gps/raw.v                       |   2 +-
 theories/gps/read.v                      |  28 +-
 theories/gps/recursive.v                 |  14 +-
 theories/gps/shared.v                    |   2 +-
 theories/gps/singlewriter.v              | 127 ++++----
 theories/gps/write.v                     |   4 +-
 theories/history.v                       |  12 +-
 theories/infrastructure.v                | 152 ++++-----
 theories/lang/base.v                     |   5 +-
 theories/lang/ra_lang.v                  |  13 +-
 theories/lang/surface.v                  |   3 +-
 theories/lifting.v                       |  30 +-
 theories/lifting_vProp.v                 |   4 +-
 theories/logically_atomic.v              | 129 +++++++-
 theories/machine.v                       |  53 ++-
 theories/malloc.v                        |  10 +-
 theories/na.v                            |  29 +-
 theories/notation.v                      |   6 +-
 theories/persistor.v                     |   2 +-
 theories/proofmode.v                     |  16 +-
 theories/rsl.v                           | 399 ++++++++++-------------
 theories/rsl_instances.v                 | 200 +++++-------
 theories/rsl_sts.v                       |  22 +-
 theories/tactics.v                       |  19 +-
 theories/tactics_vProp.v                 |   8 +-
 theories/tests/message_passing.v         |   2 +-
 theories/types.v                         |   8 +-
 theories/weakestpre.v                    | 125 ++++++-
 theories/wp_tactics_vProp.v              |   2 +-
 77 files changed, 1595 insertions(+), 1194 deletions(-)
 create mode 100644 theories/examples/hist_protocol.v
 create mode 100644 theories/examples/kvnode.v
 create mode 100644 theories/fork.v

diff --git a/LICENSE b/LICENSE
index c7ea1132..4aed1f55 100644
--- a/LICENSE
+++ b/LICENSE
@@ -1,4 +1,4 @@
-All files in this development are distributed under the terms of the BSD 
+All files in this development are distributed under the terms of the BSD
 license, included below.
 
 ------------------------------------------------------------------------------
diff --git a/_CoqProject b/_CoqProject
index 28386415..121e47bc 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -22,11 +22,13 @@ theories/blocks_generic.v
 theories/blocks.v
 theories/malloc.v
 theories/persistor.v
+theories/fork.v
 theories/fractor.v
 theories/funbi.v
 theories/viewpred.v
 theories/weakestpre.v
 theories/logically_atomic.v
+theories/invariants.v
 theories/tactics_vProp.v
 theories/wp_tactics_vProp.v
 theories/bigop.v
@@ -79,4 +81,10 @@ theories/examples/msqueue.v
 theories/examples/sc_stack.v
 theories/examples/rcu_data.v
 theories/examples/rcu.v
+theories/examples/snapshot.v
+theories/examples/hist_protocol.v
+theories/examples/abs_hashtable.v
+theories/examples/hashtable.v
+theories/examples/for_loop.v
+theories/examples/kvnode.v
 theories/tests/message_passing.v
diff --git a/opam b/opam
index cea850f2..cae6eb71 100644
--- a/opam
+++ b/opam
@@ -9,5 +9,5 @@ build: [make "-j%{jobs}%"]
 install: [make "install"]
 remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/igps" ]
 depends: [
-  "coq-iris" { (= "branch.gen_proofmode.2017-12-04.6") | (= "dev") }
+  "coq-iris" { (= "branch.gen_proofmode.2017-12-29.0") | (= "dev") }
 ]
diff --git a/theories/adequacy.v b/theories/adequacy.v
index 2aca6093..b924b47c 100644
--- a/theories/adequacy.v
+++ b/theories/adequacy.v
@@ -46,7 +46,7 @@ Proof.
   rewrite -Heqo. move/I => ?. subst. auto.
 Qed.
 
-Lemma Some_helper {X : Type} (x y : X) : Some x = Some y ↔ x = y. 
+Lemma Some_helper {X : Type} (x y : X) : Some x = Some y ↔ x = y.
 Proof. naive_solver. Qed.
 
 Lemma dom_map_imap `{Countable K} {B} (f : K -> () -> B) (g: gmap _ _) :
@@ -182,7 +182,7 @@ Definition base_adequacy Σ `{!basePreG Σ} e V σ φ :
   phys_inv σ ->
   view_ok (mem σ) V ->
   (∀ `{!foundationG Σ} π, PSCtx ∗ Seen π V ⊢ WP e {{ v, ⌜φ v⌝ }}) →
-  adequate e σ φ.
+  adequate NotStuck e σ φ.
 Proof.
   intros Inv VOk Hwp; eapply (ownP_adequacy _ _). iIntros (?).
   iMod (own_alloc (● {[1%positive := Excl V]} ⋅ ◯ {[1%positive:=Excl V]})) as (γV) "HV".
diff --git a/theories/base/accessors.v b/theories/base/accessors.v
index 29f5fdf6..34329a8e 100644
--- a/theories/base/accessors.v
+++ b/theories/base/accessors.v
@@ -69,7 +69,7 @@ Section Accessors.
           ={E∖↑physN,E}=∗ Seen π V' ∗ Hist l {[(A,V')]} ∗ Info l 1 v.
   Proof.
     iIntros (?) "[#HInv >FV]".
-    iInv physN as (Ï‚ VIEW HIST INFO) 
+    iInv physN as (Ï‚ VIEW HIST INFO)
                 "(HoP & >AV & >AH & >AI & >% & DInv & >% & >% & >%)"
                 "HClose".
 
@@ -139,7 +139,7 @@ Section Accessors.
     ={E∖↑physN,E}=∗ Seen π V' ∗ Hist l h'.
   Proof.
     iIntros (?) "(#HInv & >FV & >FH)".
-    iInv physN as (Ï‚ VIEW HIST INFO) 
+    iInv physN as (Ï‚ VIEW HIST INFO)
                 "(HoP & >AV & >AH & >AI & >% & DInv & >% & >% & >%)" "HClose".
     iCombine "AV" "FV" as "HV".
     iCombine "AH" "FH" as "HH".
@@ -156,7 +156,7 @@ Section Accessors.
     iFrame "HH2".
     iApply "HClose".
     iNext. iExists _, _, _ ,_.
-    iFrame (H4 H9) "∗". iPureIntro. 
+    iFrame (H4 H9) "∗". iPureIntro.
     split; last by apply (IInv_update_addins Ï‚ _ HIST _ l h h').
     move => ? ?. rewrite lookup_insert_Some => [] [[_ [<- //]]|[_ /H1]].
     apply: view_ok_mono => //. by apply H7.
@@ -182,7 +182,7 @@ Section Accessors.
     ={E∖↑physN,E}=∗ Seen π V'.
   Proof.
     iIntros (?) "(#HInv & >FV & >FH)".
-    iInv physN as (Ï‚ VIEW HIST INFO) 
+    iInv physN as (Ï‚ VIEW HIST INFO)
                 "(HoP & >AV & >AH & >AI & >% & DInv & >% & >% & >%)" "HClose".
     iCombine "AV" "FV" as "HV".
     iCombine "AH" "FH" as "HH".
@@ -207,4 +207,4 @@ Section Accessors.
     apply: view_ok_mono => //. by apply H7.
   Qed.
 
-End Accessors.
\ No newline at end of file
+End Accessors.
diff --git a/theories/base/alloc.v b/theories/base/alloc.v
index fd45a569..17cd6d27 100644
--- a/theories/base/alloc.v
+++ b/theories/base/alloc.v
@@ -146,4 +146,4 @@ Section Alloc.
     - repeat eexists; first exact/elem_of_singleton; eauto. solve_jsl.
     - by eexists.
   Qed.
-End Alloc.
\ No newline at end of file
+End Alloc.
diff --git a/theories/base/at_cas.v b/theories/base/at_cas.v
index bd568317..eb1a9dc0 100644
--- a/theories/base/at_cas.v
+++ b/theories/base/at_cas.v
@@ -25,7 +25,6 @@ Proof.
   iMod (PSInv_open_VH with "[$I $Seen $Hist]")
     as (Ï‚ HIST) "(oPS & % & % & % & DInv & % & HClose)"; first auto.
 
-
   destruct H3 as [HInv1 HInv2]. symmetry in H2.
   move: (HInv2 _ _ H2) => [[v0 V0] [In0 [Min0 [/= NA0 [Eqh Val]]]]].
 
@@ -90,7 +89,7 @@ Proof.
 
       iMod ("HClose" $! Ï‚' _ h' with "[DInv $HP]") as "[Seen' Hist']".
         { (* re-establish PSInv *)
-          iFrame (H H3 H4 HInv'). 
+          iFrame (H H3 H4 HInv').
           iNext. iApply (DInv_update_value with "DInv"); last eauto.
             move => ? Eq2. move: Inh'.
             rewrite Eq2 elem_of_singleton. by inversion 1. }
@@ -163,4 +162,4 @@ Proof.
         + iRight. repeat (iSplit; auto). iPureIntro.
           rewrite VUpd lookup_join_max. apply: (anti_symm (⊑)); [solve_jsl|].
           rewrite Local Eqtr. solve_jsl. }
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/base/at_fai.v b/theories/base/at_fai.v
index f8624740..e56db2ec 100644
--- a/theories/base/at_fai.v
+++ b/theories/base/at_fai.v
@@ -87,7 +87,7 @@ Proof.
 
       iMod ("HClose" $! Ï‚' _ h' with "[DInv $HP]") as "[Seen' Hist']".
         { (* re-establish PSInv *)
-          iFrame (H H3 H4 HInv'). 
+          iFrame (H H3 H4 HInv').
           iNext. iApply (DInv_update_value with "DInv"); last eauto.
           move => ? Eq2. move: Inh'.
           rewrite Eq2 elem_of_singleton. by inversion 1. }
@@ -101,7 +101,7 @@ Proof.
       + iExists (mview m), v'.
         rewrite -EqView'. iFrame ((Logic.eq_refl v') Eqh'2 Disjh' Inh').
         repeat (iSplit); iPureIntro.
-        * rewrite EqView' VUpd. etrans; [done|apply join_ext_r].
+        * rewrite EqView' VUpd. solve_jsl.
         * rewrite -EqVal /value_at Eqh2 elem_of_map_gset.
           exists m. by rewrite elem_of_hist_from.
         * by rewrite Eqtr Local.
@@ -110,12 +110,12 @@ Proof.
             { by rewrite lookup_insert. }
           move: (proj2 HInv' l h' EqL) => [vV [? [? [_ [_ IS]]]]].
           exists vV. repeat (split; auto).
-          move => [? ?]. by apply IS. 
+          move => [? ?]. by apply IS.
         * exists tr. by rewrite EqView' -EqVt EqTime'.
         * move => [[v2 V2] [In2 [t2 [Eq2 /= Eq2']]]].
           rewrite Eqtr in Eq2.
           inversion Eq2. subst t2. clear Eq2.
-          move : In2. 
+          move : In2.
           rewrite Eqh2 elem_of_map_gset.
           move  => [m2 [[Eqv2 EqV2] /elem_of_hist_from [Eql2 [Inm2 _]]]].
           destruct (IMOk _ H0 _ Inm2) as [VT2 _].
@@ -123,4 +123,4 @@ Proof.
           { by rewrite EqLoc' Eql2. }
           { rewrite EqTime'. rewrite EqV2 -Eql2 VT2 in Eq2'.
             by inversion Eq2'. }
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/base/at_read.v b/theories/base/at_read.v
index 5f8edf0b..1359ff7b 100644
--- a/theories/base/at_read.v
+++ b/theories/base/at_read.v
@@ -60,4 +60,4 @@ Proof.
           change (Some t ⊑ Some t').
           rewrite -Local -EqV0. etrans; eauto.
           etrans; first by apply (Min0 _ In'). by simpl.
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/base/at_write.v b/theories/base/at_write.v
index 3633fff7..649655fb 100644
--- a/theories/base/at_write.v
+++ b/theories/base/at_write.v
@@ -4,11 +4,11 @@ From igps Require Import proofmode.
 Import lang.base.
 
 Lemma f_write_at `{fG : !foundationG Σ} π V_l l h v:
-  {{{ PSCtx 
+  {{{ PSCtx
         ∗ ▷ Seen π V_l ∗ ▷ Hist l h ∗ ⌜alloc_local h V_l⌝ }}}
     ([ #l ]_at <- #v, V_l) @ ↑physN
     {{{ V' h', RET (LitV LitUnit, V');
-        ⌜V_l ⊑ V'⌝ ∗ Seen π V' 
+        ⌜V_l ⊑ V'⌝ ∗ Seen π V'
          ∗ Hist l h' ∗ ⌜init_local h' V'⌝ ∗ ⌜init l h'⌝
          ∗ ⌜h' ≡ {[VInj v, V']} ∪ h⌝
          ∗ ⌜h ## {[VInj v, V']}⌝
@@ -35,7 +35,7 @@ Proof.
   iApply (wp_store_at_pst with "oPS"); [by auto|by auto|..].
   - (* drf_pre ς π l *)
     econstructor; eauto. rewrite NA0. by move: (Min0 _ In') => /= ->.
-  - (* alloc_pre ς π l *) 
+  - (* alloc_pre ς π l *)
     eapply alloc_pre_helper; eauto.
     + eapply alloc_local_proper; first by auto. exists v', V'. by erewrite <-Eqh2.
     + exists (v0, V0). repeat setoid_rewrite <-Eqh2.
@@ -47,7 +47,7 @@ Proof.
 
     inversion H5.
 
-    assert (Le: (V0 !! l : option positive) ⊑ V_l !! l). 
+    assert (Le: (V0 !! l : option positive) ⊑ V_l !! l).
       { rewrite -(Local' _). by apply :(Min0 (v', V')). }
     assert (V0 !! l ⊏ Some (mtime m)).
       { by move: Le Local ->. }
@@ -84,4 +84,4 @@ Proof.
       move: (proj2 HInv' l h' EqL) => [vV [? [? [_ [_ Ha]]]]].
       exists vV. repeat (split; auto).
       move => [? ?]. apply Ha.
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/base/atomic.v b/theories/base/atomic.v
index d8a02a7d..bb7138a7 100644
--- a/theories/base/atomic.v
+++ b/theories/base/atomic.v
@@ -1 +1 @@
-From igps.base Require Export at_read at_write at_cas at_fai.
\ No newline at end of file
+From igps.base Require Export at_read at_write at_cas at_fai.
diff --git a/theories/base/base.v b/theories/base/base.v
index f1742508..11091601 100644
--- a/theories/base/base.v
+++ b/theories/base/base.v
@@ -1 +1 @@
-From igps.base Require Export atomic alloc dealloc accessors fork.
\ No newline at end of file
+From igps.base Require Export atomic alloc dealloc accessors fork.
diff --git a/theories/base/dealloc.v b/theories/base/dealloc.v
index 9a540bbd..8db23e57 100644
--- a/theories/base/dealloc.v
+++ b/theories/base/dealloc.v
@@ -4,7 +4,7 @@ From igps Require Import lang.
 Import lang.base.
 
 Lemma f_dealloc `{fG : !foundationG Σ} π V_l l h v :
-  {{{ PSCtx ∗ ▷ Seen π V_l 
+  {{{ PSCtx ∗ ▷ Seen π V_l
       ∗ ▷ Hist l h ∗ ▷ Info l 1 v ∗ ⌜alloc_local h V_l⌝ }}}
     (Dealloc #l, V_l) @ ↑physN
   {{{ V', RET (#(), V');
@@ -31,7 +31,7 @@ Proof.
   iApply (wp_dealloc_pst with "oPS"); [eauto|auto|..].
   - (* drf_pre ς π l *)
     econstructor; eauto. rewrite NA0. by move : (Min0 _ In') => /= ->.
-  - (* alloc_pre ς π l *) 
+  - (* alloc_pre ς π l *)
     eapply alloc_pre_helper; eauto.
     + eapply alloc_local_proper; first by auto. exists v', V'. by erewrite <- Eqh2.
     + exists (v0, V0). repeat setoid_rewrite <-Eqh2. do 2 (split; first auto).
@@ -58,4 +58,4 @@ Proof.
 
     (* show post-condition *)
       iFrame "Seen'". iPureIntro. subst. rewrite VUpd. solve_jsl.
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/base/ghosts.v b/theories/base/ghosts.v
index d089720e..10fecce7 100644
--- a/theories/base/ghosts.v
+++ b/theories/base/ghosts.v
@@ -30,7 +30,7 @@ Class foundationG Σ :=
 Definition AuthType {Σ} {T} (X : inG Σ (authR T)) := T.
 
 (* TODO: figure out why this is needed and get rid of it *)
-Global Instance foundationG_hist_insert `{foundationG Σ} : 
+Global Instance foundationG_hist_insert `{foundationG Σ} :
   Insert loc (excl (gset (Val * View))) (AuthType foundationG_hist) := _.
 Global Instance foundationG_hist_singleton `{foundationG Σ} :
   SingletonM loc (excl (gset (Val * View))) (AuthType foundationG_hist) := _.
@@ -184,7 +184,7 @@ Section GhostDefs.
   Definition Info : loc -> frac -> dec_agree InfoT -> iProp Σ :=
     λ l q v, own IN (◯ ({[l := (q, v)]})).
 
-  Definition LInv Ï‚ l h := 
+  Definition LInv Ï‚ l h :=
            ∃min_[rel_of (hTime l) (⊑)] vV ∈ h,
              (nats ς !! l : option positive) ⊑ snd vV !! l
              ∧ h ≡ map_vt (hist_from_opt (mem ς) l (snd vV !! l))
@@ -302,12 +302,19 @@ Section Properties.
     exists m; repeat split; auto.
     by rewrite -LE.
   Qed.
+  Global Instance na_safe_proper_flip l h :
+    Proper (flip (⊑) ==> flip impl) (na_safe l h).
+  Proof. solve_proper. Qed.
 
   Global Instance at_safe_proper h : Proper ((⊑) ==> impl) (at_safe h).
   Proof.
     intros. intros ? ? LE.
     move => [? [? [? ?]]]. do 2 eexists. split; first eauto. etrans; eauto.
   Qed.
+  Global Instance at_safe_proper_flip h :
+    Proper (flip (⊑) ==> flip impl) (at_safe h).
+  Proof. solve_proper. Qed.
+
 
   Global Instance init_local_proper h : Proper ((⊑) ==> impl) (init_local h).
   Proof.
@@ -315,6 +322,9 @@ Section Properties.
     move => [? [? [? [? ?]]]]. do 2 eexists.
     split; first eauto. split; [etrans; eauto|auto].
   Qed.
+  Global Instance init_local_proper_flip h :
+    Proper (flip (⊑) ==> flip impl) (init_local h).
+  Proof. solve_proper. Qed.
 
   Lemma alloc_init_local h V:
     init_local h V → alloc_local h V.
@@ -330,6 +340,9 @@ Section Properties.
     move => [v [V [In [Le ND]]]].
     exists v, V. repeat split; [auto| |auto]. etrans; eauto.
   Qed.
+  Global Instance alloc_local_proper_flip h :
+    Proper (flip (⊑) ==> flip impl) (alloc_local h).
+  Proof. solve_proper. Qed.
 
   Lemma init_local_at_safe h V:
     init_local h V → at_safe h V.
@@ -496,7 +509,7 @@ Section UpdateGhosts.
     iMod (own_update with "[$Own]"); last iAssumption.
     apply auth_update.
     apply: singleton_local_update; first by exact VEq.
-    by apply exclusive_local_update. 
+    by apply exclusive_local_update.
   Qed.
 
   Lemma Views_alloc (VIEW: Views) ρ V'
@@ -570,7 +583,7 @@ Section UpdateGhosts.
 
   Lemma DInv_update_return l h (h': History) (HIST: Hists)
     (Old: HIST !! l = Excl' h)
-    : (∃ v, Info l 1 v) ∗ Hist l h' ∗ DInv HIST 
+    : (∃ v, Info l 1 v) ∗ Hist l h' ∗ DInv HIST
               -∗ DInv (<[l := Excl h']>HIST).
   Proof.
     assert (HDel: delete l (<[l:=Excl h']> HIST) = delete l HIST).
@@ -610,7 +623,7 @@ Section UpdateGhosts.
   Proof.
     cbn in *.
     have H : l ∈ dom (gset loc) HIST by apply: elem_of_dom_2.
-    rewrite -HInv_IInv_agree //=; case: (HI2) => [E2 _]; 
+    rewrite -HInv_IInv_agree //=; case: (HI2) => [E2 _];
       case: (HI1) => [E1 _] //.
     rewrite -E2 II1 dom_insert E1. abstract set_solver+H E1.
   Qed.
diff --git a/theories/base/helpers.v b/theories/base/helpers.v
index 61f24570..79910d65 100644
--- a/theories/base/helpers.v
+++ b/theories/base/helpers.v
@@ -70,7 +70,7 @@ Proof.
           by inversion IS. }
       inversion EqvV. subst. clear EqvV.
       assert (HD: ∀ m',
-                  (mval m', mview m') ∈ map_vt (hist_from (mem ς) (mloc m) t) 
+                  (mval m', mview m') ∈ map_vt (hist_from (mem ς) (mloc m) t)
                   → mval m' ≠ D).
         { move => m0.
           case (decide ((mval m0, mview m0) = (A, mview m))) => [[->] //|NEq].
diff --git a/theories/base/na_read.v b/theories/base/na_read.v
index b18f4f72..5025e6ec 100644
--- a/theories/base/na_read.v
+++ b/theories/base/na_read.v
@@ -106,10 +106,7 @@ Proof.
     + iExists (mview m0). destruct (IMOk _ H1 _ In) as [EqT _]. subst.
       rewrite -EqLoc EqT. repeat iSplit; iPureIntro.
       * solve_jsl.
-      * rewrite lookup_join_max EqT.
-        apply (anti_symm sqsubseteq); first by apply join_ext_r.
-        apply join_least; last by reflexivity.
-        by rewrite EqLoc Local0.
+      * rewrite lookup_join_max EqT EqLoc Local0. solve_jsl.
       * exists (mval m0, mview m0).
         assert (EqVT: Some (mtime m0) = Some (mtime m)).
         { rewrite -(_ : msg_ok m) -?HVl; last by apply H1.
diff --git a/theories/base/na_write.v b/theories/base/na_write.v
index 08bbf314..c05b0ed8 100644
--- a/theories/base/na_write.v
+++ b/theories/base/na_write.v
@@ -7,7 +7,7 @@ Lemma f_write_na `{fG : !foundationG Σ} π V_l l h v:
   {{{ V' h', RET (#(), V');
         ⌜V_l ⊑ V'⌝ ∗ Seen π V' ∗ Hist l h'
         ∗ ⌜init_local h' V'⌝ ∗ ⌜na_safe l h' V'⌝ ∗ ⌜init l h'⌝
-        ∗ ⌜value_at_hd l h' V' (VInj v)⌝ 
+        ∗ ⌜value_at_hd l h' V' (VInj v)⌝
         ∗ ⌜h' = {[(VInj v),V']}⌝ }}}.
 Proof.
   iIntros (Φ) "(#I & Seen & Hist & %) Post".
@@ -32,7 +32,7 @@ Proof.
   iApply (wp_store_na_pst with "oPS"); [eauto|auto|..].
   - (* drf_pre ς π l *)
     econstructor; eauto. rewrite NA0. by move : (Min0 _ In') => /= ->.
-  - (* alloc_pre ς π l *) 
+  - (* alloc_pre ς π l *)
     eapply alloc_pre_helper; eauto.
     + eapply alloc_local_proper; first by auto. exists v', V'. by erewrite <- Eqh2.
     + exists (v0, V0). repeat setoid_rewrite <-Eqh2.
@@ -72,4 +72,4 @@ Proof.
     + eexists (VInj v), _. by rewrite elem_of_singleton.
     + exists (VInj v, V''). abstract set_solver+.
     + exists (VInj v, V''). abstract set_solver+.
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/bigop.v b/theories/bigop.v
index b2cf475e..dbe5f7fd 100644
--- a/theories/bigop.v
+++ b/theories/bigop.v
@@ -80,4 +80,4 @@ Qed.
 
 (* (* Notation "'[∗' 'mset' ] x ∈ X , P" := (big_opMS (M:=uPredUR _) X (λ x, P)) *) *)
 (* (*   (at level 200, X at level 10, x at level 1, right associativity, *) *)
-(* (*    format "[∗  mset ]  x  ∈  X ,  P") : uPred_scope. *) *)
\ No newline at end of file
+(* (*    format "[∗  mset ]  x  ∈  X ,  P") : uPred_scope. *) *)
diff --git a/theories/blocks.v b/theories/blocks.v
index 58833823..9b3226c1 100644
--- a/theories/blocks.v
+++ b/theories/blocks.v
@@ -45,7 +45,7 @@ Section GHist.
     - right => Eq. apply: NEq. by move : Eq => [t1 [[<-] [->]]].
     - right. move => [? [//]].
     - right. move => [? [//]].
-    - right. move => [? [//]]. 
+    - right. move => [? [//]].
   Qed.
 
   Definition gblock_ends x h := block_ends (gadj x) h.
diff --git a/theories/blocks_generic.v b/theories/blocks_generic.v
index d38ba39b..95e99ee2 100644
--- a/theories/blocks_generic.v
+++ b/theories/blocks_generic.v
@@ -23,12 +23,12 @@ Section Blocks.
   Lemma block_ends_singleton a : block_ends {[a]} ≡ {[a]}.
   Proof.
     move => ?.
-    rewrite /block_ends elem_of_filter 
+    rewrite /block_ends elem_of_filter
             set_Forall_singleton elem_of_singleton (comm (∧)).
     split => [[->]//|->]; last by split.
   Qed.
 
-  Lemma block_ends_ins_mono a S: 
+  Lemma block_ends_ins_mono a S:
     block_ends ({[ a ]} ∪ S) ⊆ {[ a ]} ∪ block_ends S.
   Proof.
     apply set_unfold_2. move => b [Hyp [Eq|In]].
@@ -110,8 +110,8 @@ Section Blocks.
             [abstract set_solver+In2|abstract set_solver+In2 NIn].
         * split.
           { split => [|//]. exact: set_Forall_union_inv_2. }
-          { move => ?; subst. 
-            apply: FA; last eauto; 
+          { move => ?; subst.
+            apply: FA; last eauto;
               [abstract set_solver+|abstract set_solver+In1 NIn]. }
       + move => [[FA In] NEq]; split; last auto.
         apply: set_Forall_union; last auto. apply set_Forall_singleton => ?.
@@ -189,8 +189,8 @@ Qed.
       (f : A -> B)
       (R : relation A) (R' : relation B)
       (S S1 S2 : gset A) a
-      (Sub1 : S1 ⊆ S) 
-      (Sub2 : S2 ⊆ {[a]} ∪ S) 
+      (Sub1 : S1 ⊆ S)
+      (Sub2 : S2 ⊆ {[a]} ∪ S)
       (RInj : ∀ b b', b ∈ {[a]} ∪ S → b' ∈ {[a]} ∪ S → R' (f b) (f b') ↔ R b b')
       (FInj : ∀ b b', b ∈ {[a]} ∪ S → b' ∈ {[a]} ∪ S → f b = f b' ↔ b = b')
       `{∀ a1 a2, Decision (R a1 a2)}
@@ -249,7 +249,7 @@ Proof.
         apply FInj in Eq; first by subst.
         abstract set_solver+In Sub1. abstract set_solver+In' Sub2.
     + move => ? In.
-      apply/RInj; [set_solver+In|set_solver+|apply:nL; set_solver+In]. 
+      apply/RInj; [set_solver+In|set_solver+|apply:nL; set_solver+In].
     + move => ? In.
       apply/RInj; [set_solver+|set_solver+In|apply:nR; set_solver+In].
   - apply elem_of_map_gset in In as [a' [? ?]]. subst.
@@ -267,4 +267,4 @@ Proof.
     + destruct sR as [b [Inb Rab]].
       move : Inb => /elem_of_map_gset [a2 [Fa2 In]]. subst b.
       exists a2. split => //. apply RInj; [set_solver+|set_solver+In|auto].
-Qed. *)
\ No newline at end of file
+Qed. *)
diff --git a/theories/derived.v b/theories/derived.v
index ca1de11b..acb47ffb 100644
--- a/theories/derived.v
+++ b/theories/derived.v
@@ -40,7 +40,7 @@ Proof. intros ??. by rewrite -wp_let. Qed.
 
 Lemma wp_skip E π Φ : ▷ Φ (LitV LitUnit, π) ⊢ WP (Skip, π) @ E {{ Φ }}.
 Proof.
-  rewrite -wp_seq; last eauto. by rewrite <-wp_value. 
+  rewrite -wp_seq; last eauto. by rewrite <-wp_value.
 Qed.
 
 
diff --git a/theories/derived_vProp.v b/theories/derived_vProp.v
index 1cdb80ab..106ee2a2 100644
--- a/theories/derived_vProp.v
+++ b/theories/derived_vProp.v
@@ -71,4 +71,4 @@ Implicit Types Φ : base.val → vProp Σ. (* TODO: fix imports *)
 (*   destruct (bool_decide_reflect (n1 = n2)); by eauto. *)
 (* Qed. *)
 
-End derived.
\ No newline at end of file
+End derived.
diff --git a/theories/escrows.v b/theories/escrows.v
index f841cc5d..f7a1e4df 100644
--- a/theories/escrows.v
+++ b/theories/escrows.v
@@ -80,14 +80,13 @@ Section proof.
     split => ?. cbn. by setoid_rewrite bi.or_comm at 1.
   Qed.
 
-  Lemma exchange_alloc P Q: (▷ (P ∨ Q)  ={↑escrowN}=∗ ∃ β, [ex P ↭ Q] β)%I.
+  Lemma exchange_alloc P Q: (▷ (P ∨ Q) ={↑escrowN}=∗ ∃ β, [ex P ↭ Q] β)%I.
   Proof.
     iStartProof (uPred _). iIntros (V) "P".
     iMod (absView_alloc with "[//]") as (β) "?".
     rewrite exchange_eq /exchange_def /=.
-    iExists β, _.
-    iMod (inv_alloc escrowN (↑escrowN) (monPred_car P _ ∨ monPred_car Q _)%I with "P") as "?".
-    by iFrame.
+    iExists β, V. iFrame. iSplitR; [done|].
+    iApply (inv_alloc escrowN (↑escrowN) (P _ ∨ Q _)%I with "P").
   Qed.
 
   Lemma exchange_alloc_l P Q: (▷ P ={↑escrowN}=∗ ∃ β, [ex P ↭ Q] β)%I.
@@ -116,8 +115,7 @@ Section proof.
     iAssert (▷ monPred_car Q V1 ∗ ▷ (monPred_car P V1 ∨ monPred_car Q V1))%I with "[P Inv ND]" as "[Q Inv]".
       { iNext.
         iDestruct "Inv" as "[Inv|$]".
-        - iExFalso. iApply ("ND" $! _ with "[//] [P Inv]").
-          rewrite -(_:V1⊑V0) //. iFrame.
+        - iExFalso. iApply ("ND" $! V0 with "[%//] [$P $Inv]").
         - by iLeft. }
     iMod ("HClose" with "Inv").
     iModIntro. iExists _. iFrame. by iNext.
@@ -173,10 +171,7 @@ Section proof.
           rewrite obj_pred_eq /=.
           iDestruct "oP" as (V1) "P1".
           iDestruct "Inv" as (V2) "P2".
-          iApply ("ND" $! (V1 ⊔ V2 ⊔ V) with "[%] [P1 P2]").
-          { rewrite a. apply join_ext_r. (* TODO why does solve_jsl not work anymore? *) }
-          iSplitL "P1"; unshelve iApply monPred_mono; try iAssumption;
-          [solve_jsl|]. rewrite -assoc comm. solve_jsl. (* TODO automate *)
+          iApply ("ND" $! (V1 ⊔ V2 ⊔ V0) with "[%] [$P1 $P2]"). solve_jsl.
         - iLeft. by rewrite obj_pred_eq /=. }
     iMod ("HClose" with "Inv").
     iModIntro. iNext. by iApply monPred_mono.
@@ -186,8 +181,7 @@ Section proof.
     (P ∗ P → False) ⊢ ([es P ⇝ Q] ∗ ▷ P ={↑escrowN}=∗ ▷ Q).
   Proof.
     iIntros "ND (Esc & oP)".
-    iApply (escrow_apply_obj with "ND [-$Esc]").
-    iNext. done. (* TODO: why do we need iNext here? *) 
+    iApply (escrow_apply_obj with "ND [-$Esc]"). auto.
   Qed.
 
 End proof.
diff --git a/theories/examples/circ_buffer.v b/theories/examples/circ_buffer.v
index 61df6a95..c8a0da05 100644
--- a/theories/examples/circ_buffer.v
+++ b/theories/examples/circ_buffer.v
@@ -266,7 +266,7 @@ Section CircBuffer.
       wp_value. iApply ("Post" with "[%] kS [Prods Cons cW pW]"); [done|].
 
       (* TODO: fix implicits or make it work with proofmode *)
-      destruct (GPS_SW_Writer_Reader (l := ZplusPos 1 q) (s := 0%nat) 
+      destruct (GPS_SW_Writer_Reader (l := ZplusPos 1 q) (s := 0%nat)
                   (IP := consInt γ q)) as [cW].
       iDestruct (cW with "cW") as "[cW cR]"; [done|].
 
@@ -364,7 +364,7 @@ Section CircBuffer.
                         Px (λ s', [XP ZplusPos 0 q in s' | prodInt P γ q ]_W)%VP
                   with "[%] kI kS [$pW]"); [|done|done|done|..].
         { rewrite Nat.add_1_r; apply Nat.le_succ_diag_r. }
-        { iSplitR "". 
+        { iSplitR "".
           - iNext. iIntros (v5). iViewUp. iIntros "(#Px & _ & $)".
             iModIntro. iNext. iSplitL ""; first done.
             iSplitL ""; iIntros "!#"; [by rewrite Nat2Z.inj_add|].
@@ -451,7 +451,7 @@ Section CircBuffer.
 
         iNext. iViewUp. iIntros (x) "kS (% & oL)".
 
-        destruct (escrow_alloc (☐ pTok γ (j + Z.to_nat n)%nat) 
+        destruct (escrow_alloc (☐ pTok γ (j + Z.to_nat n)%nat)
                                (own_loc (loc_at q (j + Z.to_nat n)%nat)))
                   as [EA].
         iDestruct (EA with "[oL]") as "PEj"; [done| |].
diff --git a/theories/examples/for_loop.v b/theories/examples/for_loop.v
index 82545eae..78d99c0c 100644
--- a/theories/examples/for_loop.v
+++ b/theories/examples/for_loop.v
@@ -60,7 +60,7 @@ Proof.
     iApply (Hbody with "[$HP] H").
     iPureIntro; auto.
   - iIntros (??) "_ H".
-    assert ((Closed (<> :b: i :b: nil) (i + #1))).
+    assert (Closed (<> :b: i :b: nil) (i + #1)).
     { unfold Closed, is_closed; simpl.
       unfold bool_decide, decide_rel.
       destruct (elem_of_list_dec _); auto.
@@ -71,7 +71,7 @@ Proof.
     wp_op.
     iApply "H"; auto.
   - iIntros (j?) "_ H".
-    assert ((Closed (<> :b: i :b: nil) (i < #N))).
+    assert (Closed (<> :b: i :b: nil) (i < #N)).
     { unfold Closed, is_closed; simpl.
       unfold bool_decide, decide_rel.
       destruct (elem_of_list_dec _); auto.
diff --git a/theories/examples/hashtable.v b/theories/examples/hashtable.v
index a11e0e1d..36551a88 100644
--- a/theories/examples/hashtable.v
+++ b/theories/examples/hashtable.v
@@ -5,7 +5,7 @@ From igps Require Export malloc repeat notation escrows.
 From igps Require Import persistor viewpred proofmode weakestpre protocols logically_atomic
   invariants fork.
 From igps.gps Require Import shared plain recursive.
-From igps.examples Require Import abs_hashtable snapshot repeat for_loop.
+From igps.examples Require Import abs_hashtable snapshot hist_protocol repeat for_loop.
 Import uPred lang.base.
 
 Lemma NoDup_remove: forall {A} (l l' : list A) a, NoDup (l ++ a :: l') ->
@@ -166,135 +166,6 @@ Qed.
 
 End code.
 
-Canonical Structure histProtocol : protocolT := ProtocolT _ (map_subseteq(M := gmap nat)(A := Z)).
-
-Instance histPrt_facts : protocol_facts histProtocol.
-Proof. esplit; try apply _.
-  - intro. apply map_included_preorder, _.
-  - intros ???. apply map_included_preorder, _.
-Qed.
-
-Section gmaps.
-
-Context {K} {E : EqDecision K} {C : Countable K} {A : ofeT} {L : LeibnizEquiv A}.
-
-(* The existing version of this requires A to be a cmra, which is completely unnecessary. *)
-Global Instance gmap_unit : Unit (gmap K A) := (∅ : gmap K A).
-
-Definition compatible (m1 m2 : gmap K A) :=
-  map_relation eq (fun _ => True) (fun _ => True) m1 m2.
-
-Lemma compatible_union (m1 m2 : gmap K A) :
-  compatible m1 m2 -> forall i x, (m1 ∪ m2) !! i = Some x ↔
-    m1 !! i = Some x ∨ (m1 !! i = None ∧ m2 !! i = Some x).
-Proof.
-  intros; unfold union, map_union, union_with, map_union_with. rewrite (lookup_merge _).
-  destruct (m1 !! i), (m2 !! i); compute; intuition congruence.
-Qed.
-
-Lemma compatible_union_assoc (m1 m2 m3 : gmap K A) :
-  compatible m1 m2 -> compatible (m1 ∪ m2) m3 -> compatible m1 (m2 ∪ m3).
-Proof.
-  repeat intro.
-  specialize (H i); specialize (H0 i).
-  unfold union, map_union in *; rewrite !lookup_union_with in H H0 |- *.
-  destruct (m1 !! i), (m2 !! i), (m3 !! i); done.
-Qed.
-
-Lemma compatible_comm1 (m1 m2 : gmap K A) : compatible m1 m2 -> compatible m2 m1.
-Proof.
-  repeat intro.
-  specialize (H i).
-  destruct (m1 !! i), (m2 !! i); auto.
-  inversion H; subst; auto.
-Qed.
-
-Global Instance compatible_comm : Comm iff compatible.
-Proof.
-  split; apply compatible_comm1.
-Qed.
-
-Global Instance compatible_refl : Reflexive compatible.
-Proof.
-  repeat intro.
-  destruct (x !! i); simpl; auto.
-Qed.
-
-Lemma compatible_subset : forall m1 m2 m (Hcompat : compatible m2 m) (Hsub : m1 ⊆ m2),
-  compatible m1 m.
-Proof.
-  repeat intro.
-  specialize (Hsub i); specialize (Hcompat i).
-  destruct (m1 !! i), (m2 !! i), (m !! i); auto; simpl in *; subst; done.
-Qed.
-
-Corollary subset_compatible : forall m1 m2 (Hsub : m1 ⊆ m2), compatible m1 m2.
-Proof.
-  intros; eapply compatible_subset; eauto; reflexivity.
-Qed.
-
-Global Instance gmapC_union : Union (gmapC K A) := map_union.
-
-Lemma compatible_union_comm : forall m1 m2 (Hcompat : compatible m1 m2),
-  m1 ∪ m2 = m2 ∪ m1.
-Proof.
-  intros; apply map_leibniz; intro.
-  specialize (Hcompat i); unfold union, gmapC_union, map_union, union_with, map_union_with.
-  rewrite !lookup_merge.
-  destruct (m1 !! i), (m2 !! i); auto; inversion Hcompat; auto.
-Qed.
-
-Lemma map_subseteq_equiv (m1 m2 : gmap K A) : m1 ⊆ m2 -> m2 ⊆ m1 -> map_equiv m1 m2.
-Proof.
-  repeat intro.
-  destruct (m1 !! i) eqn: Hm1.
-  - setoid_rewrite map_subseteq_spec in H; rewrite (H _ _ Hm1); auto.
-  - destruct (m2 !! i) eqn: Hm2; auto.
-    setoid_rewrite map_subseteq_spec in H0; rewrite (H0 _ _ Hm2) in Hm1; discriminate.
-Qed.
-
-Global Instance map_subset_order : PartialOrder (map_subseteq(M := gmap K)(A := A)).
-Proof.
-  constructor; first apply _.
-  repeat intro.
-  apply gmapC_leibniz; first apply _.
-  apply map_subseteq_equiv; auto.
-Qed.
-
-End gmaps.
-
-Instance hist_join : Join
-  (fun a b => map_subseteq(M := gmap nat)(A := Z) a b) :=
-  { join (m1 m2 : gmapC nat ZC) := if decide (compatible m1 m2)
-      then Some (union m1 m2) else None }.
-Proof.
-  - repeat intro.
-    rewrite lookup_empty; destruct (a !! i) eqn: Ha; rewrite Ha; done.
-  - intros ?????; destruct (decide _); last discriminate.
-    destruct (decide _); last discriminate.
-    intros Ha Hb; inversion Ha; inversion Hb; subst.
-    exists (a ∪ d).
-    destruct (decide _).
-    split; auto.
-    assert (compatible b a) by (rewrite comm; done).
-    destruct (decide _).
-    rewrite fin_maps.Assoc_instance_0 (compatible_union_comm b a); auto.
-    { contradiction n; apply compatible_union_assoc; auto.
-      rewrite compatible_union_comm; auto. }
-    { contradiction n.
-      eapply compatible_subset; eauto.
-      apply (map_union_subseteq_l a b). }
-  - repeat intro.
-    destruct (decide _), (decide _); auto; first (rewrite compatible_union_comm; auto);
-      contradiction n; rewrite compatible_comm; auto.
-  - intros.
-    destruct (decide _); inversion H; subst.
-    apply (map_union_subseteq_l a b).
-  - intros.
-    destruct (decide _); [|contradiction n; apply subset_compatible; auto].
-    apply f_equal, (map_subseteq_union a b); auto.
-Qed.
-
 Definition value_of (m : gmap nat Z) v := exists n, m !! n = Some v /\
   forall n', (n < n')%nat -> m !! n' = None.
 
@@ -337,7 +208,7 @@ Proof.
 Qed.
 
 Section Interpretation.
-  Context `{foundationG Σ} `{histG : inG Σ (snapR hist_join)}
+  Context `{foundationG Σ} `{histG : inG Σ (snapR (hist_join ZC))}
     {ZG : inG Σ (snapR zero_join)} `{!absViewG Σ} `{inG Σ (exclR unitC)}.
 
   Definition vP (γ : gname) : interpC Σ histProtocol := (fun b _ s v =>
@@ -380,34 +251,10 @@ Proof.
   subst; rewrite nth_lookup H; auto.
 Qed.
 
-Lemma own_snap_ord : forall `{J : Join} Σ {sG : inG Σ (snapR J)} g q a b,
-  ⎡own g (Master q a)⎤ -∗ ⎡own g (Snapshot b)⎤ -∗ ⌜R b a⌝ : vProp Σ.
-Proof.
-  intros.
-  iIntros "H1 H2".
-  iPoseProof (own_valid_2 with "H1 H2") as "Hvalid".
-  rewrite snap_validI; simpl.
-  rewrite join_unit.
-  by iDestruct "Hvalid" as "(? & % & ?)".
-Qed.
-
-Lemma own_master_eq : forall `{J : Join} Σ {sG : inG Σ (snapR J)} g q r a b,
-  ⎡own g (Master q a)⎤ -∗ ⎡own g (Master r b)⎤ -∗ ⌜a = b⌝ : vProp Σ.
-Proof.
-  intros.
-  iIntros "H1 H2".
-  iPoseProof (own_valid_2 with "H1 H2") as "Hvalid".
-  rewrite snap_validI; simpl.
-  rewrite join_unit.
-  iDestruct "Hvalid" as "(? & % & Hagree)".
-  rewrite agree_validI agree_equivI.
-  iDestruct "Hagree" as %?; auto.
-Qed.
-
 Section proof.
   Context `{fG : foundationG Σ}
           `{persG : persistorG Σ}
-          {histG : inG Σ (snapR hist_join)}
+          {histG : inG Σ (snapR (hist_join ZC))}
           {ZG : inG Σ (snapR zero_join)}
           {hashG : inG Σ (exclR (Z -c> optionC (gmapC nat ZC)))}
           `{gpGk: !gpsG Σ zeroProtocol _}
@@ -677,7 +524,7 @@ Section proof.
         (* public *) 
         ⌜value_of lv v /\ (v ≠ 0 -> exists lv', H k = Some lv' /\ lv ⊆ lv')⌝ ∗
           hashtable H g lgk lgv)%I
-      base_mask top (get_item #k).
+      base_mask ⊤ (get_item #k).
   Proof.
     unfold atomic_wp, atomic_shift.
     iIntros (Φ) "kI".
diff --git a/theories/examples/hist_protocol.v b/theories/examples/hist_protocol.v
new file mode 100644
index 00000000..192f05b4
--- /dev/null
+++ b/theories/examples/hist_protocol.v
@@ -0,0 +1,131 @@
+From igps.gps Require Export shared.
+From igps.examples Require Import snapshot.
+
+Canonical Structure histProtocol : protocolT := ProtocolT _ (map_subseteq(M := gmap nat)(A := Z)).
+
+Instance histPrt_facts : protocol_facts histProtocol.
+Proof. esplit; try apply _.
+  - intro. apply map_included_preorder, _.
+  - intros ???. apply map_included_preorder, _.
+Qed.
+
+Section gmaps.
+
+Context {K} {E : EqDecision K} {C : Countable K} {A : ofeT} {L : LeibnizEquiv A}.
+
+(* The existing version of this requires A to be a cmra, which is completely unnecessary. *)
+Global Instance gmap_unit : Unit (gmap K A) := (∅ : gmap K A).
+
+Definition compatible (m1 m2 : gmap K A) :=
+  map_relation eq (fun _ => True) (fun _ => True) m1 m2.
+
+Lemma compatible_union (m1 m2 : gmap K A) :
+  compatible m1 m2 -> forall i x, (m1 ∪ m2) !! i = Some x ↔
+    m1 !! i = Some x ∨ (m1 !! i = None ∧ m2 !! i = Some x).
+Proof.
+  intros; unfold union, map_union, union_with, map_union_with. rewrite (lookup_merge _).
+  destruct (m1 !! i), (m2 !! i); compute; intuition congruence.
+Qed.
+
+Lemma compatible_union_assoc (m1 m2 m3 : gmap K A) :
+  compatible m1 m2 -> compatible (m1 ∪ m2) m3 -> compatible m1 (m2 ∪ m3).
+Proof.
+  repeat intro.
+  specialize (H i); specialize (H0 i).
+  unfold union, map_union in *; rewrite !lookup_union_with in H H0 |- *.
+  destruct (m1 !! i), (m2 !! i), (m3 !! i); done.
+Qed.
+
+Lemma compatible_comm1 (m1 m2 : gmap K A) : compatible m1 m2 -> compatible m2 m1.
+Proof.
+  repeat intro.
+  specialize (H i).
+  destruct (m1 !! i), (m2 !! i); auto.
+  inversion H; subst; auto.
+Qed.
+
+Global Instance compatible_comm : Comm iff compatible.
+Proof.
+  split; apply compatible_comm1.
+Qed.
+
+Global Instance compatible_refl : Reflexive compatible.
+Proof.
+  repeat intro.
+  destruct (x !! i); simpl; auto.
+Qed.
+
+Lemma compatible_subset : forall m1 m2 m (Hcompat : compatible m2 m) (Hsub : m1 ⊆ m2),
+  compatible m1 m.
+Proof.
+  repeat intro.
+  specialize (Hsub i); specialize (Hcompat i).
+  destruct (m1 !! i), (m2 !! i), (m !! i); auto; simpl in *; subst; done.
+Qed.
+
+Corollary subset_compatible : forall m1 m2 (Hsub : m1 ⊆ m2), compatible m1 m2.
+Proof.
+  intros; eapply compatible_subset; eauto; reflexivity.
+Qed.
+
+Global Instance gmapC_union : Union (gmapC K A) := map_union.
+
+Lemma compatible_union_comm : forall m1 m2 (Hcompat : compatible m1 m2),
+  m1 ∪ m2 = m2 ∪ m1.
+Proof.
+  intros; apply map_leibniz; intro.
+  specialize (Hcompat i); unfold union, gmapC_union, map_union, union_with, map_union_with.
+  rewrite !lookup_merge.
+  destruct (m1 !! i), (m2 !! i); auto; inversion Hcompat; auto.
+Qed.
+
+Lemma map_subseteq_equiv (m1 m2 : gmap K A) : m1 ⊆ m2 -> m2 ⊆ m1 -> map_equiv m1 m2.
+Proof.
+  repeat intro.
+  destruct (m1 !! i) eqn: Hm1.
+  - setoid_rewrite map_subseteq_spec in H; rewrite (H _ _ Hm1); auto.
+  - destruct (m2 !! i) eqn: Hm2; auto.
+    setoid_rewrite map_subseteq_spec in H0; rewrite (H0 _ _ Hm2) in Hm1; discriminate.
+Qed.
+
+Global Instance map_subset_order : PartialOrder (map_subseteq(M := gmap K)(A := A)).
+Proof.
+  constructor; first apply _.
+  repeat intro.
+  apply gmapC_leibniz; first apply _.
+  apply map_subseteq_equiv; auto.
+Qed.
+
+End gmaps.
+
+Instance hist_join (A : ofeT) {D : OfeDiscrete A} {L : LeibnizEquiv A} {U : Unit A} `{EqDecision A} : Join
+  (fun (a b : gmapC nat A) => map_subseteq(M := gmap nat)(A := A) a b) :=
+  { join (m1 m2 : gmapC nat A) := if decide (compatible m1 m2)
+      then Some (union m1 m2) else None }.
+Proof.
+  - repeat intro.
+    rewrite lookup_empty; destruct (a !! i) eqn: Ha; rewrite Ha; done.
+  - intros ?????; destruct (decide _); last discriminate.
+    destruct (decide _); last discriminate.
+    intros Ha Hb; inversion Ha; inversion Hb; subst.
+    exists (a ∪ d).
+    destruct (decide _).
+    split; auto.
+    assert (compatible b a) by (rewrite comm; done).
+    destruct (decide _).
+    rewrite fin_maps.Assoc_instance_0 (compatible_union_comm b a); auto.
+    { contradiction n; apply compatible_union_assoc; auto.
+      rewrite compatible_union_comm; auto. }
+    { contradiction n.
+      eapply compatible_subset; eauto.
+      apply (map_union_subseteq_l a b). }
+  - repeat intro.
+    destruct (decide _), (decide _); auto; first (rewrite compatible_union_comm; auto);
+      contradiction n; rewrite compatible_comm; auto.
+  - intros.
+    destruct (decide _); inversion H; subst.
+    apply (map_union_subseteq_l a b).
+  - intros.
+    destruct (decide _); [|contradiction n; apply subset_compatible; auto].
+    apply f_equal, (map_subseteq_union a b); auto.
+Qed.
diff --git a/theories/examples/kvnode.v b/theories/examples/kvnode.v
new file mode 100644
index 00000000..78f141bc
--- /dev/null
+++ b/theories/examples/kvnode.v
@@ -0,0 +1,332 @@
+From iris.program_logic Require Export weakestpre.
+From iris.proofmode Require Import tactics.
+
+From igps Require Export malloc repeat notation escrows.
+From igps Require Import persistor viewpred proofmode weakestpre protocols singlewriter logically_atomic
+  invariants fork.
+From igps.gps Require Import shared plain recursive.
+From igps.examples Require Import abs_hashtable snapshot hist_protocol repeat for_loop.
+Import uPred lang.base.
+
+Section kvnode.
+
+(*Context (m_entries : loc) (results : loc).*)
+
+Local Notation version := 0.
+Local Notation data := 1.
+
+(* There's no default location, so make sure i is in-bounds. *)
+Local Notation undef := 1%positive.
+
+Section code.
+
+Program Definition read : val := @RecV <> "n" (λ: "out",
+  (rec: "f" <> :=
+     let: "snap" := ["n" + #version]_at in
+     if: "snap" `mod` #2 = #1 then "f" #()
+     else (for: ("i" := #0; "i" < #8; "i" + #1)
+                  ["out" + "i"]_na <- ["n" + #data + "i"]_at);;
+          let: "v" := ["n" + #version]_at in
+          if: "v" = "snap" then #() else "f" #()) #()) _.
+Next Obligation.
+Proof.
+  apply I.
+Qed.
+
+Program Definition write : val := @RecV <> "n" (λ: "in",
+  let: "v" := ["n" + #version]_at in
+  ["n" + #version]_at <- "v" + #1;;
+  (for: ("i" := #0; "i" < #8; "i" + #1)
+          ["n" + #data + "i"]_at <- ["in" + "i"]_na);;
+  ["n" + #version]_at <- "v" + #2) _.
+Next Obligation.
+Proof.
+  apply I.
+Qed.
+
+Program Definition make_node : val := @RecV <> <>
+  (let: "n" := malloc #9 in
+   ["n" + #version]_na <- #0;;
+   (for: ("i" := #0; "i" < #8; "i" + #1)
+        ["n" + #data + "i"]_na <- #0);;
+   "n") _.
+Next Obligation.
+Proof.
+  apply I.
+Qed.
+
+(*void *writer(void *n){
+  int data[8] = {0, 0, 0, 0, 0, 0, 0, 0};
+  for(int i = 0; i < 3; i++){
+    for(int j = 0; j < 8; j++){
+      int v = data[j];
+      data[j] = v + 1;
+    }
+    write(n, data);
+  }
+  return NULL;
+}
+
+void *reader(void *n){
+  int data[8];
+  read(n, data);
+  return NULL;
+}*)
+
+End code.
+
+Canonical Structure versionProtocol : protocolT := ProtocolT _ le.
+
+Instance versionPrt_facts : protocol_facts versionProtocol.
+Proof. esplit; apply _.
+Qed.
+
+Instance version_join : Join le := { join a b := Some (max a b) }.
+Proof.
+  - apply Nat.le_0_l.
+  - intros ????? Hc He; inversion Hc; inversion He; subst.
+    do 2 eexists; eauto.
+    by rewrite Max.max_assoc (Max.max_comm a b).
+  - intros ??.
+    by rewrite Max.max_comm.
+  - intros ??? X; inversion X.
+    apply Nat.le_max_l.
+  - intros.
+    by rewrite Max.max_r.
+Qed.
+
+Definition log_latest {A} (m : gmap nat A) n v := m !! n = Some v /\ forall n', (n < n')%nat -> m !! n' = None.
+
+(* The existing version of this requires A to be a cmra, which is completely unnecessary. *)
+Global Instance list_unit {A} : Unit (list A) := nil.
+
+Section Interpretation.
+  Context `{foundationG Σ} `{histG : inG Σ (snapR (hist_join (list.listC ZC)))}
+          `{persG : persistorG Σ}
+          `{gpGv: !gpsG Σ versionProtocol _}
+          `{gpGd: !gpsG Σ histProtocol _}.
+
+  (*
+  Fixpoint versionP version locs (γ : gname) : interpC Σ versionProtocol := (fun _ _ s v =>
+      ⌜(Z.of_nat s = v)%Z⌝ ∗ ∃ L : gmap nat (list Z), ⌜∃ vs, log_latest L (if even s then s else (s - 1)%nat) vs⌝ ∗
+      ⎡@own _ _ histG γ (Snapshot L)⎤ ∗
+      [∗ list] i↦x ∈ locs, [XP x in (fun l => nth i l 0) <$> L | dataP version locs γ]_R)%I
+  with dataP version locs (γ : gname) : interpC Σ histProtocol := (fun _ _ s v =>
+    ∃ ver, ⌜even ver = true /\ log_latest s ver v⌝ ∗ [XP version in (ver - 1)%nat | versionP version locs γ]_R)%I.
+  *)
+
+  Definition dataP' n versionP : interpC Σ histProtocol := (fun _ _ s v =>
+    ∃ ver, ⌜even ver = true /\ log_latest s ver v⌝ ∗ [XP (ZplusPos version n) in (ver - 1)%nat | versionP]_R)%I.
+
+(*  Definition versionP_fun version locs γ : (nat -> Z -> vProp Σ) -> (nat -> Z -> vProp Σ) := (λ R s v,
+      ⌜(Z.of_nat s = v)%Z⌝ ∗ ∃ L : gmap nat (list Z), ⌜∃ vs, log_latest L (if even s then s else (s - 1)%nat) vs⌝ ∗
+      ⎡@own _ _ histG γ (Snapshot L)⎤ ∗
+      [∗ list] i↦x ∈ locs, [XP x in (fun l => nth i l 0) <$> L | dataP' version (fun _ _ => R)]_R)%I.*)
+
+  Definition versionP_fun n γ : interpC Σ versionProtocol -> interpC Σ versionProtocol := (λ R _ _ s v,
+      ⌜(Z.of_nat s = v)%Z⌝ ∗ ∃ L : gmap nat (list Z), ⌜∃ vs, log_latest L (if even s then s else (s - 1)%nat) vs⌝ ∗
+      ⎡@own _ _ histG γ (Snapshot L)⎤ ∗
+      [∗ list] i ∈ seq 0 8,
+        [XP (ZplusPos (Z.of_nat i) (ZplusPos data n)) in (fun l => nth i l 0) <$> L | dataP' n R]_R)%I.
+
+  Instance versionP_contractive n γ : Contractive (versionP_fun n γ).
+  Proof.
+    repeat intro.
+    apply sep_ne; first done.
+    apply exist_ne; intro.
+    apply sep_ne; first done.
+    apply sep_ne; first done.
+    apply big_opL_ne; [repeat intro | done].
+    apply vGPS_RSP_contractive.
+    apply dist_dist_later.
+    repeat intro.
+    apply exist_ne; intro.
+    apply sep_ne; first done.
+    apply vGPS_RSP_contractive; done.
+  Qed.
+  
+  Definition versionP n γ : interpC Σ versionProtocol := fixpoint (versionP_fun n γ).
+
+  Definition dataP n γ := dataP' n (versionP n γ).
+  
+  Lemma versionP_eq' n γ: versionP n γ ≡ (fun _ _ s v =>
+      ⌜(Z.of_nat s = v)%Z⌝ ∗ ∃ L : gmap nat (list Z), ⌜∃ vs, log_latest L (if even s then s else (s - 1)%nat) vs⌝ ∗
+      ⎡@own _ _ histG γ (Snapshot L)⎤ ∗
+      [∗ list] i ∈ seq 0 8,
+        [XP (ZplusPos (Z.of_nat i) (ZplusPos data n)) in (fun l => nth i l 0) <$> L | dataP n γ]_R)%I.
+  Proof.
+    by rewrite /versionP fixpoint_unfold.
+  Qed.
+
+  Lemma versionP_eq n γ b l s v: versionP n γ b l s v ≡
+     (⌜(Z.of_nat s = v)%Z⌝ ∗ ∃ L : gmap nat (list Z), ⌜∃ vs, log_latest L (if even s then s else (s - 1)%nat) vs⌝ ∗
+      ⎡@own _ _ histG γ (Snapshot L)⎤ ∗
+      [∗ list] i ∈ seq 0 8,
+        [XP (ZplusPos (Z.of_nat i) (ZplusPos data n)) in (fun l => nth i l 0) <$> L | dataP n γ]_R)%I.
+  Proof.
+    apply versionP_eq'.
+  Qed.
+
+  (*(* adapted from message_passing *)
+  Definition XE (γ: gname) (P : vProp Σ) := [es ⎡own γ (Excl ())⎤ ⇝ P ]%I.
+
+  Definition onceP (γ : gname) (P : vProp Σ) : interpC Σ boolProtocol := (fun b _ s v =>
+    ⌜(if s then v = 1 else v = 0)%Z⌝ ∗ if s then XE γ P else True)%I.*)
+
+End Interpretation.
+
+Section proof.
+  Context `{fG : foundationG Σ}
+          `{persG : persistorG Σ}
+          `{histG : inG Σ (snapR (hist_join (list.listC ZC)))}
+          `{gpGv: !gpsG Σ versionProtocol _}
+          `{gpGd: !gpsG Σ histProtocol _}
+          `{agreeGk : !gps_agreeG Σ versionProtocol}
+          `{agreeGv : !gps_agreeG Σ histProtocol}.
+
+  Set Default Proof Using "Type".
+
+  (* The collection of protocol assertions that describes a stable state of the node. *)
+  Definition node_state_R L v n g :=
+    (⌜even v = true /\ exists vs, log_latest L v vs⌝ ∗
+     [XP (ZplusPos version n) in v | versionP n g]_R ∗ ⎡@own _ _ histG g (Snapshot L)⎤ ∗
+     [∗ list] i ∈ seq 0 8,
+       [XP (ZplusPos (Z.of_nat i) (ZplusPos data n)) in (fun l => nth i l 0) <$> L | dataP n g]_R)%I.
+
+  Definition node_state_W L v n g :=
+    (⌜even v = true /\ exists vs, log_latest L v vs⌝ ∗
+     [XP (ZplusPos version n) in v | versionP n g]_W ∗ ⎡@own _ _ histG g (Master (1/2) L)⎤ ∗
+     [∗ list] i ∈ seq 0 8,
+       [XP (ZplusPos (Z.of_nat i) (ZplusPos data n)) in (fun l => nth i l 0) <$> L | dataP n g]_W)%I.
+
+  Ltac wp_rec' := lazymatch goal with |-context[Rec ?f ?x ?e] => let H := fresh "Hclosed" in
+    assert (Closed (f :b: x :b: nil) e) as H by apply I; wp_rec; clear H;
+    unfold subst; fold subst; try (rewrite !subst_for; try done); unfold subst; fold subst; simpl end.
+
+  Definition base_mask : coPset := ↑physN ∪ (*↑fracN ∪*) ↑persist_locN.
+
+  Lemma phys_base : ↑physN ⊆ base_mask.
+  Proof.
+    (*etransitivity; *)apply union_subseteq_l.
+  Qed.
+
+(*  Lemma frac_base {A} {_ : EqDecision A} {_ : Countable A} (i : A) :
+    ↑namespaces.ndot fracN i ⊆ base_mask.
+  Proof.
+    etransitivity; last apply union_subseteq_l.
+    etransitivity; last apply union_subseteq_r.
+    apply namespaces.nclose_subseteq.
+  Qed.*)
+
+  Lemma persist_base {A} {_ : EqDecision A} {_ : Countable A} (i : A) :
+    ↑namespaces.ndot persist_locN i ⊆ base_mask.
+  Proof.
+    etransitivity; last apply union_subseteq_r.
+    apply namespaces.nclose_subseteq.
+  Qed.
+
+  Hint Resolve phys_base (*frac_base*) persist_base.
+
+  Instance versionP_persistent n g b l s v: Persistent (versionP n g b l s v).
+  Proof.
+    intros.
+    rewrite versionP_eq.
+    apply sep_persistent; first apply _.
+    apply exist_persistent; intro.
+    apply sep_persistent; first apply _.
+    apply sep_persistent; first apply _.
+    apply big_sepL_persistent; intros.
+    apply _.
+  Qed.
+  
+  Instance node_state_R_persistent L v n g: Persistent (node_state_R L v n g).
+  Proof.
+    apply sep_persistent; first apply _.
+    apply sep_persistent; first apply _.
+    apply sep_persistent; first apply _.
+    apply big_sepL_persistent; intros.
+    apply _.
+  Qed.
+
+  Opaque seq.
+  
+  Lemma read_spec n out L0 v0 g: atomic_wp
+    (λ L, ⎡PSCtx⎤ ∗ node_state_R L0 v0 n g ∗ ([∗ list] i ∈ seq 0 8, ZplusPos (Z.of_nat i) out ↦ ?) ∗
+                       ⎡own g (Master (1/2) L)⎤)%I
+    (λ L val, ⌜val = #()⌝ ∗ ∃ v vals, ⌜length vals = 8%nat /\ (v0 <= v)%nat /\ L0 ⊆ L /\ L !! v = Some vals⌝ ∗
+          node_state_R (<[v := vals]>L0) v n g ∗ ([∗ list] i↦x ∈ vals, ZplusPos (Z.of_nat i) out ↦ x) ∗
+                       ⎡own g (Master (1/2) L)⎤)%I base_mask ⊤ (read #n #out).
+  Proof.
+    iIntros (Φ) "shift"; iDestruct "shift" as (P) "[HP #Hshift]".
+    match goal with |-context[(▷ P ={_,_}=∗ ?R)%I] => set (Q := R) end.
+
+    unfold read.
+    unfold of_val.
+    wp_rec'. wp_rec'.
+    iLöb as "IH" (*forall (L0 v0)*).
+    wp_rec'.
+    match goal with |-context[App (Rec _ ?b ?e) _] => assert (Closed (b :b: nil) e) as Hclosed by apply I end.
+    wp_bind ([_]_at)%E; clear Hclosed.
+    wp_op.
+    iMod ("Hshift" with "HP") as (L) "((#kI & #node_state & out & oL) & Hclose)".
+    iApply (GPS_SW_Read (versionP n g) emp%I
+      (fun s v => |={base_mask,⊤}=> versionP n g false (ZplusPos 0 n) s v ∗ ▷ P)%I with "[-]");
+      [auto | auto | ..].
+    { iSplit; first auto.
+      iSplitR ""; [|iSplit].
+      - iNext; iIntros (s Hs).
+        iLeft; iIntros (v) "[vP _]".
+        iFrame.
+        iApply "Hclose".
+        iFrame; auto.
+      - iNext; do 2 (rewrite all_forall; iIntros (?)).
+        iApply all_entails.
+        iIntros "[? #?]"; iModIntro; iFrame "#".
+      - iSplit; first auto.
+        by iDestruct "node_state" as "(? & ? & ?)". }
+    iNext; iIntros (v ?) "(% & #kV & Hrest)".
+    rewrite versionP_eq.
+    iMod "Hrest" as "((% & HL1) & HP)"; iModIntro.
+    iDestruct "HL1" as (L1) "(% & oL1 & Hdata)"; subst.
+    match goal with |-context[App (Rec _ ?b ?e) _] => assert (Closed (b :b: nil) e) as Hclosed by apply I end.
+    wp_rec'; clear Hclosed.
+    wp_op; wp_op.
+    destruct (decide (v mod 2 = 1)).
+    { rewrite bool_decide_true; last done.
+      wp_if_true.
+      by iApply "IH". }
+    rewrite bool_decide_false; last done.
+    wp_if_false.
+    match goal with |-context[App (Rec _ ?b ?e) _] => assert (Closed (b :b: nil) e) as Hclosed by apply I end.
+    wp_bind (for_loop _ _ _ _); clear Hclosed.
+    iApply fupd_wp.
+    iMod ("Hshift" with "HP") as (L2) "((_ & _ & out & oL2) & Hclose)".
+    iApply (wp_for_simple with "[-]").
+    - instantiate (1 := (fun i => ∃ vals, ⌜i = Z.of_nat (length vals)⌝ ∗
+        ([∗ list] i↦x ∈ vals, ZplusPos (Z.of_nat i) out ↦ x) ∗
+        ([∗ list] i ∈ seq (Z.to_nat i) (8 - Z.to_nat i), ZplusPos (Z.of_nat i) out ↦ ?) ∗
+        ∃ v', [XP ZplusPos version n in v' | versionP n g]_R ∗ ⎡@own _ _ histG g (Snapshot L1)⎤ ∗
+        ([∗ list] i↦x ∈ vals, ∃ vi, ⌜(v <= vi <= v')%nat⌝ ∗
+           [XP ZplusPos (Z.of_nat i) (ZplusPos data n) in <[vi := x]>((fun l => nth i l 0) <$> L1) | dataP n g]_R) ∗
+        ([∗ list] i ∈ seq (Z.to_nat i) (8 - Z.to_nat i), 
+           [XP ZplusPos (Z.of_nat i) (ZplusPos data n) in (fun l => nth i l 0) <$> L1 | dataP n g]_R))%I).
+      iIntros (i ?) "(% & H) Hpost".
+      iDestruct "H" as (vals) "(% & Hout & Hrest & H)".
+      iDestruct "H" as (v') "(#kV & oL1 & Hvals & Hdata)".
+      admit.
+    - iExists nil; simpl.
+      iFrame; eauto.
+    - iNext; iIntros (?) "H"; iDestruct "H" as (i) "(% & H)".
+      iDestruct "H" as (vals) "(% & Hout & _ & H)".
+      iDestruct "H" as (v') "(#kV' & oL1 & Hvals & _)".
+      match goal with |-context[App (Rec _ ?b ?e) _] => assert (Closed (b :b: nil) e) as Hclosed by apply I end.
+      wp_seq; clear Hclosed.
+      match goal with |-context[App (Rec _ ?b ?e) _] => assert (Closed (b :b: nil) e) as Hclosed by apply I end.
+      wp_bind ([_]_at)%E; clear Hclosed.      
+      wp_op.
+      admit.
+  Admitted.
+      
+End proof.
+
+End kvnode.
diff --git a/theories/examples/message_passing.v b/theories/examples/message_passing.v
index 21f2794c..648589a9 100644
--- a/theories/examples/message_passing.v
+++ b/theories/examples/message_passing.v
@@ -1,15 +1,13 @@
-From iris.program_logic Require Export weakestpre.
 From iris.proofmode Require Import tactics.
+From iris.base_logic Require Import iprop own.
 From iris.algebra Require Import excl.
 
-From igps Require Export malloc repeat notation wp_tactics escrows.
+From igps Require Export malloc repeat notation wp_tactics_vProp escrows.
 From igps Require Import persistor viewpred proofmode weakestpre protocols.
-From igps.base Require Import fork.
+From igps Require Import fork.
 From igps.gps Require Import plain.
 From igps.examples Require Import unit_token.
 
-Import uPred.
-
 Definition message_passing : base.expr :=
   let: "x" := Alloc in
   let: "y" := Alloc in
@@ -24,22 +22,20 @@ Section Interpretation.
 
   Set Default Proof Using "Type".
 
-  Definition XE (x : loc) (γ: gname) : @vPred Σ
-    := [es Tok γ ⇝ x ↦ 37 ].
-
-  Definition YP x (γ: gname) (s: pr_state boolProtocol) (v: Z) : @vPred Σ :=
-    match s with
-    | false => â–  (v = 0)
-    | true => ■ (v = 1) ∗ XE x γ
-    end.
+  Definition XE (x : loc) (γ: gname) : vProp Σ
+    := [es Tok γ ⇝ x ↦ 37 ]%I.
 
-  Global Instance YP_persistent x γ s v V: Persistent (YP x γ s v V).
-  Proof. destruct s; by iIntros "#YP". Qed.
+  Definition YP x (γ: gname) (s: pr_state boolProtocol) (v: Z) : vProp Σ :=
+    (match s with
+    | false => ⌜v = 0%nat⌝
+    | true => ⌜v = 1%nat⌝ ∗ XE x γ
+    end)%I.
 
-  Local Open Scope VP.
+  Global Instance YP_persistent x γ s v: Persistent (YP x γ s v).
+  Proof. rewrite /Persistent. destruct s; by iIntros "#YP". Qed.
 
   Definition mpInt x γ : interpC Σ boolProtocol
-    := λ b _ s v, if b then True else YP x γ s v.
+    := (λ b _ s v, if b then True else YP x γ s v)%I.
 
   Global Instance interp_persistent : Persistent (mpInt x γ b l s v V).
   Proof. destruct b; intros; apply _. Qed.
@@ -55,100 +51,103 @@ Section proof.
 
   Set Default Proof Using "Type*".
   Local Notation iProp := (iProp Σ).
-  Local Notation vPred := (@vPred Σ).
-
-  Open Scope vPred_scope.
+  Local Notation vPred := (vProp Σ).
 
   Lemma message_passing_spec:
-    {{{{ ☐ PersistorCtx }}}}
+    {{{⎡PSCtx⎤ ∗ ⎡PersistorCtx⎤ }}}
         (message_passing)
-    {{{{ (v: Z), RET #v; â–  (v = 37) }}}}%stdpp.
+    {{{ (v: Z), RET #v; ⌜v = 37%nat⌝ }}}.
   Proof.
-    intros. iViewUp; iIntros "#kI kS #kIp Post". rewrite /message_passing.
+    intros. iIntros "[#kI #kIp] Post". rewrite /message_passing.
     wp_bind Alloc.
-    iApply (alloc with "[%] kI kS []"); [done|done|done|].
-    iNext. iViewUp. iIntros (x) "kS Nax".
-    wp_seq. iNext.
+    iApply (alloc with "kI"); [done|].
+    iNext. iIntros (x) "Nax".
+    wp_seq.
     wp_bind Alloc.
-    iApply (alloc with "[%] kI kS []"); [done|done|done|].
-    iNext. iViewUp. iIntros (y) "kS Nay".
-    wp_seq. iNext.
+    iApply (alloc with "kI"); [done|].
+    iNext. iIntros (y) "Nay".
+    wp_seq.
     wp_bind ([_]_na <- _)%E.
 
-    iApply (na_write with "[%] kI kS [$Nax]"); [done|done|].
-    iNext. iViewUp; iIntros "kS Nax".
-    wp_seq. iNext.
+    iApply (na_write with "[$kI $Nax]"); [done|].
+    iNext. iIntros "Nax".
+    wp_seq.
     wp_bind ([_]_na <- _)%E.
 
     iMod (own_alloc (Excl ())) as (γ) "Tok"; first done.
 
-    iApply (GPS_PP_Init_default (mpInt x γ) false with "[%] kI kS [$kIp $Nay]");
+    iApply (GPS_PP_Init_default (mpInt x γ) false with "[kI $kIp] [$Nay]");
     [done|done|done|..].
     { iNext. by iSplitL "". }
 
-    iNext. iViewUp; iIntros "kS #yPRT".
-    wp_seq. iNext.
+    iNext. iIntros "#yPRT".
+    wp_seq.
 
     wp_bind (Fork _)%E.
     iApply (wp_mask_mono (↑physN)); first done.
     iCombine "Nax" "yPRT" as "SndT".
-    iApply (f_fork _ ((x ↦ 0) V_l ∗ [PP y in false @ () | mpInt x γ ] V_l)%I
-      with "[$kI $kS $SndT Post Tok]").
+    iApply (wp_fork with "[$kI SndT Post Tok]"); [done|].
     iSplitL "Post Tok".
-    - iNext. iViewUp. iIntros "kS".
-      wp_seq. iNext. wp_bind ((rec: "f" <> := _ ) _)%E.
-      (iLöb as "IH" forall (V_l) "kIp yPRT ∗").
-      (* TODO: kIp is a spurious dependency *)
+    - iNext.
+      wp_seq. wp_bind ((rec: "f" <> := _ ) _)%E.
+      (iLöb as "IH" forall "yPRT ∗").
       iApply wp_repeat; first auto.
-      iApply (GPS_PP_Read (p:=()) (mpInt x γ) True (YP x γ)
-        with "[%] kI kS [$yPRT]"); [done|done|done|..].
+      iApply (GPS_PP_Read (p:=()) (mpInt x γ) True%I (YP x γ)
+        with "kI [$yPRT]"); [done|done|..].
       + iSplitL.
-        * iNext. iIntros (?). iViewUp. iIntros "_".
-          iLeft. iIntros (??) "_ [YP  _]". by simpl.
-        * iNext. iIntros (? s v). iViewUp. by iIntros "[? #$]".
-      + iNext. iViewUp; iIntros (s v) "kS (_ & #yPRT5 & #YP)".
+        * iNext. iIntros (?). iIntros "_".
+          iLeft. iIntros (?) "[YP  _]". by simpl.
+        * iNext.
+          (* iIntros (? s v). iViewUp. by iIntros "[? #$]". *)
+          admit. (* FIXME *)
+      + iNext. iIntros (s v) "(_ & #yPRT5 & #YP)".
         destruct s.
         * iDestruct "YP" as "(% & XE)".
-          iExists true. iSplit; first by rewrite H.
-          iDestruct (escrow_apply (Tok γ) (x ↦ 37)) as "EA".
+          iExists (1). iSplit; first by rewrite H.
+          iDestruct (escrow_apply (Tok γ) (x ↦ 37))%I as "EA".
           iSpecialize ("EA" with "[]").
-          { iIntros (??) "_ ?". by iApply unit_tok_exclusive. }
-          erewrite decide_False; last auto.
-          iSpecialize ("EA" $! _ _ with "[%] [$XE $Tok]"); first done.
-          iDestruct (fupd_mask_mono with "EA") as "EA"; first auto.
-          iNext. iNext. iMod "EA" as "Nax".
-          wp_seq. iNext.
-          iApply (na_read with "[%] kI kS Nax"); [done|done|].
-          iNext. iViewUp; iIntros (z) "kS (% & Nax)".
-          subst z. by iApply ("Post" with "[%] kS").
+          { iIntros "?". by iApply unit_tok_exclusive. }
+          erewrite bool_decide_false; last auto.
+          iSpecialize ("EA" with "[$XE $Tok]").
+          iNext. iNext.
+          iApply fupd_wp.
+          Fail iDestruct (fupd_mask_mono with "EA") as "EA"; first auto.  (* TODO: why??? *)
+          iApply (fupd_mask_mono (I:=View_bi_index) (↑escrowN)); [done|].
+          iMod "EA" as "Nax".
+          iModIntro.
+          wp_seq.
+          iApply (na_read with "[$kI $Nax]"); [done|].
+          iNext. iIntros (z) "(% & Nax)".
+          subst z. by iApply ("Post" with "[%]").
         * iDestruct "YP" as "%". 
-          iExists false. iSplit; first by rewrite H.
-          erewrite decide_True; last done.
+          iExists 0. iSplit; first by rewrite H.
+          erewrite bool_decide_true; last done.
           iNext. iNext.
-          iApply ("IH" $! _ with "[kIp] [yPRT5] Post Tok kS").
-          { iIntros "!#"; iFrame "kIp". }
-          { iIntros "!#"; iFrame "yPRT5". }
-    - iNext. iIntros (ρ). iViewUp. iIntros "_ kS (Nax & #yPRT')". move: HV => ->.
+          iApply ("IH" with "[yPRT5] Post Tok").
+          iIntros "!#"; iFrame "yPRT5".
+    - iNext. iDestruct "SndT" as "(Nax & #yPRT')".
       wp_bind ([_]_na <- _)%E.
-      iApply (na_write with "[%] kI kS [$Nax]"); [done|done|].
-      iNext. iViewUp; iIntros "kS Nax".
-      wp_seq. iNext.
+      iApply (na_write with "[$kI $Nax]"); [done|].
+      iNext. iIntros "Nax".
+      wp_seq.
 
       (* allocate escrow *)
-      destruct (escrow_alloc (Tok γ) (x ↦ 37)) as [EA].
-      iDestruct (EA _ with "[$Nax]") as "XE"; [done|].
-
-      iDestruct (fupd_mask_mono _ ⊤ with "XE") as ">XE"; first auto.
+      iDestruct (escrow_alloc (Tok γ) (x ↦ 37)%I with "[$Nax]") as "XE".
 
+      iApply fupd_wp. iApply (fupd_mask_mono (I:=View_bi_index) (↑escrowN)); [done|]. (* TODO: why??? *)
+      (* iDestruct (fupd_mask_mono _ ⊤ with "XE") as ">XE"; first auto. *)
+      iMod "XE".
+ 
       (* write release to final state *)
-      iApply (GPS_PP_Write (p:=()) (mpInt x γ) true (XE x γ) (λ _, True)
-              with "[%] kI kS [yPRT' XE]"); [done|done|move=>[|]; done|done|..].
+      iApply (GPS_PP_Write (p:=()) (mpInt x γ) true (XE x γ) (λ _, True%I)
+              with "[$kI yPRT' XE]"); [done|done|move=>[|]; done|..].
       { iSplitL ""; [|by iFrame "XE"].
-        iNext. iViewUp; iIntros "[oL _]". iModIntro.
+        iNext. iIntros "[oL _]". iModIntro.
         iSplitL ""; [done|]. iSplit; iNext; auto. by iFrame "oL". }
 
       (* trivial post cond *)
-      iNext. iViewUp. auto.
-  Qed.
+      iModIntro.
+      iNext. auto.
+  Admitted.
 
 End proof.
diff --git a/theories/examples/message_passing_base.v b/theories/examples/message_passing_base.v
index c94c353a..2a53de87 100644
--- a/theories/examples/message_passing_base.v
+++ b/theories/examples/message_passing_base.v
@@ -156,4 +156,4 @@ Section proof.
         * by iApply ("Invxs" $! _ v').
   Qed.
 
-End proof.
\ No newline at end of file
+End proof.
diff --git a/theories/examples/msqueue.v b/theories/examples/msqueue.v
index a36cb386..d49072b0 100644
--- a/theories/examples/msqueue.v
+++ b/theories/examples/msqueue.v
@@ -45,7 +45,7 @@ Section MSQueue.
       else #0.
 
   Definition tryDeq : base.val :=
-    λ: "q", 
+    λ: "q",
       let: "s" := ["q" + #head]_at in
       let: "n" := [((loc)"s") + #link]_at in
       if: "n" = #0
@@ -100,7 +100,7 @@ Section MSQueue.
         match s with
         | inl _ => ⌜z = 0⌝
         | inr l' => ⌜z = Z.pos l'⌝
-                    ∗ ∃ γ', DEQ (ZplusPos data l') γ γ' 
+                    ∗ ∃ γ', DEQ (ZplusPos data l') γ γ'
                             ∗ [PP (ZplusPos link l') in Null @ γ' | F γ']
         end)%I.
 
@@ -137,7 +137,7 @@ Section MSQueue.
     Definition Head : interpC Σ unitProtocol :=
       (λ b _ _ x,
         if b
-        then 
+        then
           ∃ γ, [PP (Z.to_pos (x + link)%Z) in Null @ γ | Link γ] ∗ Tok γ
         else
           ⌜0 < x⌝ ∗ ∃ γ, [PP (Z.to_pos (x + link)%Z) in Null @ γ | Link γ])%I.
@@ -218,7 +218,7 @@ Section MSQueue.
     Lemma findTail_spec q:
       {{{ ☐ PSCtx ∗ Queue q }}}
         findTail #q
-      {{{ (n: Z), RET #n; ⌜n = 0⌝ ∨ 
+      {{{ (n: Z), RET #n; ⌜n = 0⌝ ∨
             ⌜n > 0⌝ ∗ ∃ γ, [PP (Z.to_pos (n + link)%Z) in Null @ γ | Link P γ] }}}.
     Proof.
       intros. iIntros "[#kI #(oH & oT)] Post".
@@ -226,7 +226,7 @@ Section MSQueue.
       apply: (coq_tactics.tac_next ?[Δ] ?[Δ'] ?[n]). (* TODO: why doesn't iNext/eapply/simple apply work *)
       wp_bind ([_]_at)%E. wp_op. iNext.
 
-      set Q : pr_state unitProtocol → Z → vProp Σ 
+      set Q : pr_state unitProtocol → Z → vProp Σ
         := λ _ x, (⌜0 < x⌝
               ∗ (∃ γ : gname, [PP Z.to_pos (x + 0) in Null @ γ | Link P γ ]))%I.
       iApply (GPS_PP_Read (p:=()) (Tail P) True%I Q with "[$oT]");
@@ -251,7 +251,7 @@ Section MSQueue.
           iIntros (v). iIntros "(LinkInt & _)".
           rewrite (unf_int (Link P)).
           iModIntro. iIntros "%". destruct s' as [|l'].
-          + by iDestruct "LinkInt" as "%". 
+          + by iDestruct "LinkInt" as "%".
           + iDestruct "LinkInt" as "(% & LinkInt)".
             iDestruct "LinkInt" as (γ') "(_ & Link)".
             iSplitL ""; first (iPureIntro; lia).
@@ -403,7 +403,7 @@ Section MSQueue.
         last first. { rewrite /ZplusPos. f_equal. rewrite Z2Pos.id; omega. }
 
       set Q' : pr_state LinkProtocol → Z → vPred
-        := λ s' n, 
+        := λ s' n,
           (■ (0 ≠ n)
               → (■ (0 < n) ∗ ∃ γ', DEQ P (Z.to_pos (n + data)) γ γ'
                       ∗ [PP Z.to_pos (n + link) in Null @ γ' | Link P γ']))%VP.
@@ -475,4 +475,4 @@ Section MSQueue.
           iApply ("Post" with "[%] kS []"); [done|by iLeft].
     Qed.
   End proof.
-End MSQueue.
\ No newline at end of file
+End MSQueue.
diff --git a/theories/examples/protocols.v b/theories/examples/protocols.v
index f83f7538..3b97cc7e 100644
--- a/theories/examples/protocols.v
+++ b/theories/examples/protocols.v
@@ -13,4 +13,4 @@ Proof. esplit; try apply _; [by intros []|by intros [] [] []]. Qed.
 Definition unitProtocol : protocolT := ProtocolT unit (λ _ _, True).
 
 Instance unitPrt_facts : protocol_facts (unitProtocol).
-Proof. esplit; apply _. Qed.
\ No newline at end of file
+Proof. esplit; apply _. Qed.
diff --git a/theories/examples/rcu.v b/theories/examples/rcu.v
index 5663fef1..3a11d39e 100644
--- a/theories/examples/rcu.v
+++ b/theories/examples/rcu.v
@@ -97,7 +97,7 @@ Section RCU.
       ["q" + #link]_na <- #0 ;;
       ["q" + #del]_na <- ((int)(newStack #())) ;;
       ["q" + #wc  ]_na <- #0 ;;
-      let: "sc" := Alloc in 
+      let: "sc" := Alloc in
       ["sc"]_na <- #0 ;;
       repeat (
         let: "i" := ["sc"]_na in
@@ -404,7 +404,7 @@ Section RCU.
       master Γ D V ={E}=∗ master Γ D' V ∗ snapshot Γ D' V.
     Proof.
       iIntros "MS". iDestruct "MS" as (W) "MS".
-      iAssert (|==> own Γ (● DList D' W' ⋅ ◯ DList D' W'))%I 
+      iAssert (|==> own Γ (● DList D' W' ⋅ ◯ DList D' W'))%I
         with "[MS]" as ">(MS & SN)".
         { iApply (own_update with "MS").
           rewrite auth_both_op.
@@ -505,11 +505,11 @@ Section RCU.
       Persistent (SnapshotValid Γ D V) := _.
 
     Local Open Scope VP.
-    Definition WCP Γ : interpC Σ RCUCounterProtocol := 
+    Definition WCP Γ : interpC Σ RCUCounterProtocol :=
       λ b _ Dn x, if b then True
                   else ■ (x = Dn.2) ∗ SnapshotValid Γ (Dn.1).
 
-    Definition RCP Γ t : interpC Σ RCUCounterProtocol := 
+    Definition RCP Γ t : interpC Σ RCUCounterProtocol :=
       λ b _ Dn x,
         if b then True
              else ■ (x = Dn.2) ∗ [XP (ZplusPos wc (lhd Γ)) in Dn | WCP Γ]_R
@@ -602,7 +602,7 @@ Section RCU.
     Definition TraceSpine_elem Γ D L L' (i: nat) p j γ: vPred :=
       ■ Trace_link_fact D L L' (p.1) i j γ
       ∗ (   ■ (i = 0)%nat ∗ Trace_own_head Γ D j (p.1) γ
-          ∨ ■ (i > 0)%nat ∗ Trace_own_tail Γ D j p γ 
+          ∨ ■ (i > 0)%nat ∗ Trace_own_tail Γ D j p γ
               ∗ □ PermXList Γ (Z.to_pos (p.1)) j γ).
 
     Definition TraceSpineAbs Γ D (L: list (Z * Z)) (L': list (aloc * (Z * Z)))
@@ -623,7 +623,7 @@ Section RCU.
       (* [∗ list] t ∈ seq 0 NPosn, tnToks Γ t (dom _ D). *)
 
     Definition CurrentState Γ L D : vPred
-      := master Γ D ∗ SnapshotValid Γ D 
+      := master Γ D ∗ SnapshotValid Γ D
         ∗ TraceSpine Γ D L ∗ ■ SinglePtr D
         ∗ RemainingThreadTokens Γ D.
 
@@ -637,7 +637,7 @@ Section RCU.
     Lemma rcu_remain_txtokens_one Γ D D' V i
       (HD: dom (gset aloc) D' ≡ {[i]} ∪ dom (gset _) D)
       (NIn: i ∉ dom (gset _) D) :
-      RemainingThreadTokens Γ D V 
+      RemainingThreadTokens Γ D V
       ⊢ RemainingThreadTokens Γ D' V ∗
         ([∗ list] t ∈ seq 0 NPosn, tTok Γ t i V).
     Proof.
@@ -673,7 +673,7 @@ Section RCU.
       {{{{ ☐ PersistorCtx}}}}
           (rcuNew #())
       {{{{ (q: loc), RET #q;
-            WriterSafe P q ((Zpos q, 0) :: nil) 
+            WriterSafe P q ((Zpos q, 0) :: nil)
             ∗ vPred_big_opL (seq 0 NPosn) (λ _ t, ∃ D, ReaderSafe P q D t) }}}}.
             (* ∗ [∗ list] t ∈ seq 0 NPosn, ∃ D, ReaderSafe P q D t *)
     Proof.
@@ -700,7 +700,7 @@ Section RCU.
 
       set i := 4%positive.
       set Dnil : RCUData := (nil,nil).
-      set D : gname → RCUData 
+      set D : gname → RCUData
         := λ γ, ((i, Null) :: nil, (i,ZplusPos 0 q,γ) :: nil).
       iMod (own_alloc (● dlist_empty)) as (γh) "Hist"; first done.
       iMod (Tok_map_alloc NPosn) as (γw) "WX".
@@ -741,7 +741,7 @@ Section RCU.
       wp_bind ([_]_na <- _)%E. wp_op. iNext.
 
       (* getting SnapshotValid *)
-      iDestruct (GPS_nSW_Writer_Reader (Param:=RCUGhost) (IP:=LLP P Γ i) with "ol") as "[ol #qr]". 
+      iDestruct (GPS_nSW_Writer_Reader (Param:=RCUGhost) (IP:=LLP P Γ i) with "ol") as "[ol #qr]".
       iAssert ((SnapshotValid P Γ (D γ))%VP V_l) as "#SV".
       { iFrame "SN". iExists i. iSplitL ""; first iPureIntro; last iSplitR "".
         - by rewrite /basel RCUData_basel_singleton.
@@ -907,7 +907,7 @@ Section RCU.
           iApply (big_sepL_mono with "RS"). iIntros (???) "?". by iExists _.
       - iNext. wp_let. iNext. wp_op => [_|?//].
         iNext. iApply wp_if_true. iNext. rewrite -Nat.add_1_r.
-        iApply ("IH" $! _ (j+1)%nat (n'') with 
+        iApply ("IH" $! _ (j+1)%nat (n'') with
                       "[%] [%] [%] [kP] [] [] [] [] Post
                        WX RX TX MS SCS od ol ow kS [osc] oLs RS").
         + apply Z.le_neq. split; first lia.
@@ -972,7 +972,7 @@ Section RCU.
           - destruct SN as [[SN|[SN Eqb]] _]; first done.
             move: H0. rewrite /base -Eqb /base' /basel elem_of_of_list.
             apply hd_error_some_elem.
-          - destruct SN as [[SN|[SN Eqb]] _]; first done. 
+          - destruct SN as [[SN|[SN Eqb]] _]; first done.
             move: H0. rewrite /base Eqb /base' /basel elem_of_of_list.
             apply hd_error_some_elem. }
         destruct bj as [|j|];
@@ -1051,7 +1051,7 @@ Section RCU.
         { iDestruct (rcuSnapshotValid_extract with "SN1") as (b1 γ1) "nH1".
           iDestruct (rcuSnapshotValid_extract with "SN2") as (b2 γ2) "nH2".
           rewrite -H -H0.
-          by iDestruct (GPS_nSWs_param_agree 
+          by iDestruct (GPS_nSWs_param_agree
                          (IP1:=LLP P Γ' b2) (IP2:=LLP P Γ b1) γ2 γ1 Γ' Γ
                          (hist D b2) (hist D b1) V_l with "[$nH2 $nH1]")
                          as %[? ?]. }
@@ -1273,7 +1273,7 @@ Section RCU.
 
 
       (* (([∗ set] i ∈ (dead D3 ∖ dead D), ModX Γ t i)%VP V_l) *)
-      iAssert (|={⊤}=> ((vPred_big_opS (dead D3 ∖ dead D) (λ i, ModX Γ t i))%VP V_l))%I 
+      iAssert (|={⊤}=> ((vPred_big_opS (dead D3 ∖ dead D) (λ i, ModX Γ t i))%VP V_l))%I
           with "[RD]" as "RD".
       { iApply (fupd_big_sepS _ _ (dead D3 ∖ dead D)).
         iApply (big_sepS_mono (λ i, rTok Γ t i V_l)); first reflexivity.
@@ -1285,7 +1285,7 @@ Section RCU.
           + by rewrite /rToks rcuTok_own_big_sepS. }
 
       iMod "RD" as "#RD".
-      iAssert ((□ (∀ i : aloc, ■ (i ∈ dead D3) → ModX Γ t i))%VP V_l) 
+      iAssert ((□ (∀ i : aloc, ■ (i ∈ dead D3) → ModX Γ t i))%VP V_l)
         with "[]" as "#MD'".
       { iIntros "!#". iIntros (i). iViewUp. iIntros (Ini).
         case (decide (i ∈ dead D)) => [InD|NInD].
@@ -1293,7 +1293,7 @@ Section RCU.
         - iApply (big_sepS_elem_of _ ((dead D3 ∖ dead D)) with "RD").
           by apply elem_of_difference. }
 
-      iApply (GPS_SW_ExWrite (RCP P Γ t) (D3 ,n3) True 
+      iApply (GPS_SW_ExWrite (RCP P Γ t) (D3 ,n3) True
                   (λ s, [XP ZplusPos (rc + t) q in s | RCP P Γ t ]_W)%VP
                 with "[%] kI kS [$RW]"); [by etrans|done|done|reflexivity|..].
       { iNext. iIntros (v'). iViewUp. iIntros "(_ & $ & $)".
@@ -1411,7 +1411,7 @@ Section RCU.
       inversion H0. subst D'' m''.
 
       (* ([∗ set] i ∈ (dead D2 ∖ dead D1), rTok Γ j i)%VP *)
-      iAssert (|={⊤}=> (vPred_big_opS (dead D2 ∖ dead D1) (λ i, rTok Γ j i) V_l))%I 
+      iAssert (|={⊤}=> (vPred_big_opS (dead D2 ∖ dead D1) (λ i, rTok Γ j i) V_l))%I
           with "[wj]" as ">rT".
       { case (decide (dead D2 ∖ dead D1 = ∅)) => [->|?].
         - by rewrite -vPred_big_opS_fold big_sepS_empty.
@@ -1451,7 +1451,7 @@ Section RCU.
         iFrame "WCP RS".
       - iNext. wp_let. iNext. wp_op => [_|? //]. iNext.
         iApply wp_if_true. iNext. rewrite -Nat.add_1_r.
-        iApply ("IH" $! _ (j+1)%nat (n'') with 
+        iApply ("IH" $! _ (j+1)%nat (n'') with
                       "[%] [%] [%] WCP wT Post kS [osc] RS");
           [lia|omega|lia|by rewrite Nat2Z.inj_add].
     Qed.
@@ -1523,11 +1523,11 @@ Section RCU.
     Lemma rcuDealloc_spec q Γ D1 D2 D3 p:
       lhd Γ = q → D1 ⊑ D2 → D2 ⊑ D3 →
       {{{{  RevokedUpTo P Γ D1 ∗ DeallocBetween P Γ D1 D2
-          ∗ SnapshotValid P Γ D3 ∗ master Γ D3 
+          ∗ SnapshotValid P Γ D3 ∗ master Γ D3
           ∗ ∃ i γ, ■ (i ∈ dead D3 ∖ dead D2
                           ∧ info D3 i = Some (p,γ))
                    ∗ □ PermXList P Γ p i γ
-                   ∗ ∃ s v, [nFXP (ZplusPos link p) in s@γ | LLP P Γ i]_W_nf 
+                   ∗ ∃ s v, [nFXP (ZplusPos link p) in s@γ | LLP P Γ i]_W_nf
                           ∗ (ZplusPos data p) ↦{nf} v }}}}
         rcuDealloc #q #(Zpos p)
       {{{{ (r: Z), RET #r;
@@ -1787,7 +1787,7 @@ Section RCU.
             -∗ ∃ Γ D La j j' γ γ', master Γ D
               ∗ TraceSpine_elem P Γ D Lc La i (p,v) j γ
               ∗ TraceSpine_elem P Γ D Lc La (i + 1)%nat (p',v') j' γ'
-              ∗ (master Γ D 
+              ∗ (master Γ D
                   ∗ TraceSpine_elem P Γ D Lc La i (p,v) j γ
                   ∗ TraceSpine_elem P Γ D Lc La (i + 1)%nat (p',v') j' γ'
                  -∗ WriterSafe P q Lc).
@@ -1826,7 +1826,7 @@ Section RCU.
         rewrite -Zplus_0_r_reverse.
         iApply ("Post" with "[] kS [XP $LLP]"); first done.
         iFrame (H). iLeft. rewrite /Trace_own_head -Zplus_0_r_reverse.
-        iFrame (H0) "XP". 
+        iFrame (H0) "XP".
       - iApply (GPS_nFWP_ExRead (LLP P Γ j) γ with "[%] kI kS [$FXP]");
             [auto|auto|reflexivity|..].
         { iNext. iIntros (? p'). iViewUp. iIntros "#?". by iFrame "#". }
@@ -2662,7 +2662,7 @@ Section RCU.
             iDestruct "OS" as (j γj) "[(%&_&%) _]". iPureIntro.
             destruct (HHj j H0) as [γj' TFj].
             assert (γj' = γj).
-            { destruct TFj as (_&_&HLk&_). 
+            { destruct TFj as (_&_&HLk&_).
               rewrite (RCUData_info_lookup _ _ _ _ UW DLe HLk) in H1.
               by inversion H1. } subst γj'.
             exists j. rewrite H0. split; last auto.
@@ -3140,7 +3140,7 @@ Section RCU.
           rewrite app_length /= in EqLen.
           rewrite Hi -plus_n_O -EqLen.
           split; [by apply list_lookup_middle|auto]. }
-      iAssert ((□ (∀ t, ■ (t < NPosn)%nat → PermX P Γ x t j γ))%VP V_l) 
+      iAssert ((□ (∀ t, ■ (t < NPosn)%nat → PermX P Γ x t j γ))%VP V_l)
         with "[]" as "#Perm2".
       { iIntros "!#". iIntros (t). iViewUp.
         iIntros (Lt). rewrite -vPred_big_opL_fold.
@@ -3223,7 +3223,7 @@ Section RCU.
 
       (* setting up new history *)
       set j : aloc := fresh (dom (gset aloc) D3).
-      set D3' : gname -> RCUData 
+      set D3' : gname -> RCUData
         := λ γ, ((i, Abs j) :: (j, Null) :: D3.1, (j,x,γ) :: D3.2).
       set L1 := (L ++ (p, v) :: (Z.pos x, v') :: nil)%list.
       set L1' := (L' ++ (j, (Z.pos x, v')) :: nil)%list.
diff --git a/theories/examples/rcu_data.v b/theories/examples/rcu_data.v
index 4068ec41..290c9bfa 100644
--- a/theories/examples/rcu_data.v
+++ b/theories/examples/rcu_data.v
@@ -92,10 +92,10 @@ Section RCUData.
     - intros [|]; by [left|by right].
     - intros ? ? ? [?|[? ?]]; first by left.
       intros [?|[? ?]]; subst.
-      + left. by apply suffix_nil_inv. 
+      + left. by apply suffix_nil_inv.
       + right. split; by etrans.
     - intros ?? [?|[? _]] [?|[? _]]; subst; auto.
-      + symmetry. by apply suffix_nil_inv. 
+      + symmetry. by apply suffix_nil_inv.
       + by apply suffix_nil_inv.
       + by apply: anti_symm.
   Qed.
@@ -744,7 +744,7 @@ Section RCUData.
     rewrite elem_of_list_filter elem_of_app. split; auto.
   Qed.
 
-  Definition RCU_part_le D1 D2 := 
+  Definition RCU_part_le D1 D2 :=
      (D1.1) `suffix_of` (D2.1) ∧ base' (D1.1) ≡ base' (D2.1) ∧ (D1.2) ⊑ (D2.2).
 
   Lemma RCUData_ext_share_total_1 D1 D2 D3
@@ -1158,7 +1158,7 @@ Section RCUData.
           { rewrite rcu_hist'_cons_some.
             destruct bjE as [|[j']]; subst bj; first auto.
             apply (elem_of_rcu_hist_dom_mono
-                      ((i, Abs j0) :: (j0, Abs j') :: (j, Dead) :: nil) 
+                      ((i, Abs j0) :: (j0, Abs j') :: (j, Dead) :: nil)
                       _ _ (InDj' j' eq_refl)). }
           { rewrite rcu_hist'_cons_next; last exact NEq2.
             case (decide (j = k)) => [->|NEq3];
@@ -1219,7 +1219,7 @@ Section RCUData.
             destruct (NEqi2 j') as (?&?&?); first auto.
             apply elem_of_rcu_live'_further; first auto.
             apply elem_of_rcu_live'_further.
-            - apply HN. by apply elem_of_live'_dom. 
+            - apply HN. by apply elem_of_live'_dom.
             - apply elem_of_rcu_live'_further; auto. }
           { rewrite rcu_hist'_cons_next; last exact NEqj0.
             case (decide (j = k)) => [?|NEj].
diff --git a/theories/examples/repeat.v b/theories/examples/repeat.v
index 2ce5b8bb..5a067886 100644
--- a/theories/examples/repeat.v
+++ b/theories/examples/repeat.v
@@ -28,9 +28,9 @@ Notation repeat e :=
 (* Qed. *)
 
 Lemma wp_repeat `{ghosts.foundationG Σ} E e Φ `{He : Closed nil e} :
-  WP e @ E 
-    {{ v, ∃ z, ⌜v = LitV $ LitInt z⌝ 
-          ∧ if bool_decide (z = 0) 
+  WP e @ E
+    {{ v, ∃ z, ⌜v = LitV $ LitInt z⌝
+          ∧ if bool_decide (z = 0)
             then ▷ ▷ (WP (repeat e) @ E {{v, Φ v }}) else ▷ ▷ Φ #z }}
      ⊢ WP (repeat e) @ E {{ Φ }}.
 Proof.
@@ -39,7 +39,7 @@ Proof.
   wp_bind e. iApply wp_mono; [|iExact "H"].
   iIntros (v) "H". iDestruct "H" as (l) "(% & R)".
   rewrite /= H0 /= {H0}.
-  wp_lam. 
+  wp_lam.
   case HD: (bool_decide _);
   wp_op; rewrite HD; [by wp_if_true|by wp_if_false].
 Qed.
diff --git a/theories/examples/sc_stack.v b/theories/examples/sc_stack.v
index 546dbe13..198fca03 100644
--- a/theories/examples/sc_stack.v
+++ b/theories/examples/sc_stack.v
@@ -27,7 +27,7 @@ Section SCStack.
       ["s"]_na <- ((int)"n").
 
   Definition pop : base.val :=
-    λ: "s", 
+    λ: "s",
       let: "h" := ["s"]_na in
       if: "h" = #0
       then #0
@@ -48,13 +48,13 @@ Section SCStack.
 
     Local Open Scope VP.
 
-    Definition SCStack': (loc -c> list Z -c> @vPred_ofe Σ) 
+    Definition SCStack': (loc -c> list Z -c> @vPred_ofe Σ)
                              -c> loc -c> list Z -c> @vPred_ofe Σ
     := (fun F l A =>
         ∃ l', ZplusPos next l ↦ l' ∗
         match A with
         | nil => l' = 0
-        | v::A' => ■ (0 < l') ∗ Z.to_pos (l' + data) ↦ v ∗ P v 
+        | v::A' => ■ (0 < l') ∗ Z.to_pos (l' + data) ↦ v ∗ P v
                               ∗ ▷ F (Z.to_pos l') A'
         end).
 
diff --git a/theories/examples/snapshot.v b/theories/examples/snapshot.v
index c443458f..e73d1822 100644
--- a/theories/examples/snapshot.v
+++ b/theories/examples/snapshot.v
@@ -343,6 +343,7 @@ Proof.
   rewrite -!snap_both_op snap_validN_eq /=.
   split_and!; [by apply cmra_included_includedN|by apply cmra_valid_validN|done].
 Qed.*)
+
 End cmra.
 
 Arguments snapR {_} {_} {_} {_} {_} J.
@@ -387,3 +388,31 @@ Proof. intros n f f' Hf [[[]|] [|]]; do 2 constructor; try naive_solver; simpl.
   constructor; auto. apply agreeC_map_ne; auto.
   constructor; auto. apply agreeC_map_ne; auto.
 Qed.
+
+From iris.proofmode Require Import tactics.
+From iris.base_logic Require Import own.
+From igps Require Import weakestpre.
+
+Lemma own_snap_ord : forall `{J : Join} Σ {sG : inG Σ (snapR J)} g q a b,
+  ⎡own g (Master q a)⎤ -∗ ⎡own g (Snapshot b)⎤ -∗ ⌜R b a⌝ : vProp Σ.
+Proof.
+  intros.
+  iIntros "H1 H2".
+  iPoseProof (own_valid_2 with "H1 H2") as "Hvalid".
+  rewrite snap_validI; simpl.
+  rewrite join_unit.
+  by iDestruct "Hvalid" as "(? & % & ?)".
+Qed.
+
+Lemma own_master_eq : forall `{J : Join} Σ {sG : inG Σ (snapR J)} g q r a b,
+  ⎡own g (Master q a)⎤ -∗ ⎡own g (Master r b)⎤ -∗ ⌜a = b⌝ : vProp Σ.
+Proof.
+  intros.
+  iIntros "H1 H2".
+  iPoseProof (own_valid_2 with "H1 H2") as "Hvalid".
+  rewrite snap_validI; simpl.
+  rewrite join_unit.
+  iDestruct "Hvalid" as "(? & % & Hagree)".
+  rewrite agree_validI agree_equivI.
+  iDestruct "Hagree" as %?; auto.
+Qed.
diff --git a/theories/examples/spin_lock.v b/theories/examples/spin_lock.v
index 2b50b6d7..9370a21d 100644
--- a/theories/examples/spin_lock.v
+++ b/theories/examples/spin_lock.v
@@ -23,7 +23,7 @@ Definition lock : base.val :=
 
 
 Section proof.
-  Context `{fG : !foundationG Σ, 
+  Context `{fG : !foundationG Σ,
             rG : rslG Σ, perG : persistorG Σ}.
   Set Default Proof Using "Type".
   Local Notation iProp := (iProp Σ).
@@ -56,7 +56,7 @@ Section proof.
     iApply (alloc with "[%] kI kS [%]"); [done|done|done|].
     iNext. iViewUp; iIntros (l) "kS oL".
     wp_seq. iNext. wp_bind ([_]_na <- _)%E.
-    iApply (RMWAcq_init l 0 (Q P) Q_dup 
+    iApply (RMWAcq_init l 0 (Q P) Q_dup
             with "[%] kI kS [$kP $oL P]"); [auto|auto|reflexivity|..].
     { iNext; iSplit. by iRight; iFrame. by iRight. }
     iNext. iViewUp. iIntros "kS (Rel & RMW & Init)".
@@ -108,4 +108,4 @@ Section proof.
       iApply ("IH" with "kS' Rel Acq Init Post").
   Qed.
 
-End proof.
\ No newline at end of file
+End proof.
diff --git a/theories/examples/ticket_lock.v b/theories/examples/ticket_lock.v
index f706a7e5..81be9060 100644
--- a/theories/examples/ticket_lock.v
+++ b/theories/examples/ticket_lock.v
@@ -9,7 +9,7 @@ From igps.examples Require Import nat_tokens protocols.
 From igps.gps Require Import singlewriter plain recursive.
 
 (** Formalization of bounded ticket-lock examples.
-    The proof is simplified by using single-writer protocols and 
+    The proof is simplified by using single-writer protocols and
       a simpler ticket tracker. **)
 
 Definition ticketLockR
@@ -70,7 +70,7 @@ Section TicketLock.
 
     Definition Ws M n : nat := size (Waiting M n).
 
-    Definition ns_tkt_bound M (t n: nat) : Prop := 
+    Definition ns_tkt_bound M (t n: nat) : Prop :=
       n ≤ t ∧ t ≤ n + Ws M n
       ∧ ∀ (i k: nat), M !! i = Excl2 k →
           (k < t) ∧ ∀ (j: nat), M !! j = Excl2 k → i = j.
@@ -282,7 +282,7 @@ Section TicketLock.
           split; [|split; first lia; split].
           * eapply Z.le_lt_trans; first by apply LE2.
             apply Z.add_lt_le_mono; [lia|]. by apply Waiting_size_le.
-          * rewrite Nat2Z.inj_succ Nat2Z.inj_add Zplus_assoc_reverse 
+          * rewrite Nat2Z.inj_succ Nat2Z.inj_add Zplus_assoc_reverse
                     (Z.add_comm 1) Zplus_assoc.
             apply Zplus_le_compat_r. etrans; first apply LE2.
             assert (Eqs :=
@@ -311,7 +311,7 @@ Section TicketLock.
               { rewrite disjoint_singleton_l /Waiting elem_of_filter.
                 rewrite Ini. move => [[? [[<-] ?]] _]. omega. }
               rewrite Nat2Z.inj_add size_singleton Nat2Z.inj_add
-                      (Z.add_comm 1) Zplus_assoc. 
+                      (Z.add_comm 1) Zplus_assoc.
               by apply Zplus_le_compat_r. }
           * by apply ns_inj_insert.
       - exists n. split; [done|split]; first by right.
@@ -328,7 +328,7 @@ Section TicketLock.
             { rewrite disjoint_singleton_l /Waiting elem_of_filter.
               rewrite Ini. move => [[? [//]]]. }
             rewrite Nat2Z.inj_add size_singleton Nat2Z.inj_add
-                    (Z.add_comm 1) Zplus_assoc. 
+                    (Z.add_comm 1) Zplus_assoc.
             by apply Zplus_le_compat_r. }
         * by apply ns_inj_insert.
     Qed.
@@ -385,7 +385,7 @@ Section TicketLock.
     Proof.
       iIntros "Perms". rewrite -own_op pair_op.
       iDestruct (own_valid with "Perms") as %[_ Valid].
-      iPureIntro. 
+      iPureIntro.
       move: Valid => /= /auth_valid_discrete_2 [/lookup_included /(_ i)].
       rewrite lookup_singleton option_included.
       move => [//|[x [y [-> [Leq [/leibniz_equiv_iff -> //|Incl]]]]]] Valid.
@@ -517,14 +517,14 @@ Section TicketLock.
     Local Notation vPred := (@vPred Σ).
     Context (P : vPred).
 
-    Definition MayAcquire (x: loc): vPred := 
+    Definition MayAcquire (x: loc): vPred :=
       ∃ γ i ot, ☐ MyTkt γ i ot
         ∗ [PP (ZplusPos tc x) in () | TCP P γ x ]
         ∗ □ ∀ (t: nat), ■ (ot = Some t) → [XP (ZplusPos ns x) in S t | NSP P γ ]_R.
 
     Definition MayRelease (x: loc): vPred :=
       ∃ γ i t, ☐ MyTkt γ i (Some t)
-        ∗ [PP (ZplusPos tc x) in () | TCP P γ x ] 
+        ∗ [PP (ZplusPos tc x) in () | TCP P γ x ]
         ∗ [XP (ZplusPos ns x) in t | NSP P γ ]_W.
 
     Lemma newLock_spec:
@@ -555,7 +555,7 @@ Section TicketLock.
         iDestruct (GPS_SW_Writer_Reader (IP := NSP P γ) _ V_l with "oW")
           as "[oW $]"; [done|].
         iApply fupd_mask_mono;
-        last iMod (escrow_alloc (☐ Perm γ 0) 
+        last iMod (escrow_alloc (☐ Perm γ 0)
                                 (P ∗ [XP ZplusPos 0 x in O |NSP P γ ]_W) V_l
                      with "[$P $oW]") as "#?";
         [solve_ndisj|done|].
@@ -573,7 +573,7 @@ Section TicketLock.
 
       wp_seq. iNext. wp_bind ([_]_na <- _)%E. wp_op. iNext.
 
-      iApply (GPS_PP_Init_default (TCP P γ x) () 
+      iApply (GPS_PP_Init_default (TCP P γ x) ()
                 with "[%] kI kS [$kP $oTC Perms AllTkts NSR]");
                 [done|done|done| |].
       { iNext. iSplitR ""; last done.
@@ -612,7 +612,7 @@ Section TicketLock.
       wp_lam. iNext. wp_bind (FAI _ _)%E. wp_op. iNext.
 
       set P' : vPred := (☐ MyTkt γ i ot)%VP.
-      set Q' : pr_state unitProtocol → Z → vPred 
+      set Q' : pr_state unitProtocol → Z → vPred
         := (λ _ v, ∃ (t n0: nat), ■ (v = t `mod` Z.pos C ∧ t < n0 + Z.pos C)
                      ∗ ☐ MyTkt γ i (Some t) ∗ ☐ Perm γ t
                      ∗ ▷ [XP x in n0 | NSP P γ ]_R)%VP.
@@ -634,7 +634,7 @@ Section TicketLock.
           iModIntro. iNext.
           destruct ot as [t0|]; first (move: (nEq2 t0) => [//|->|-> //]).
           + by iApply ("NS" $! t0 with "[%]").
-          + by rewrite nEq1. 
+          + by rewrite nEq1.
         - iModIntro. iNext. iSplitR ""; last done.
           iExists (S t), n', (<[i := Excl (Some t)]> M).
           iFrame "All' Perms". iSplitR "NSR".
@@ -657,7 +657,7 @@ Section TicketLock.
       iLöb as "IH" forall (V_l) "NS TCP ∗".
       iApply wp_repeat; first auto. wp_bind ([_]_at)%E. wp_op. iNext.
 
-      set Q2 : pr_state natProtocol → Z → vPred 
+      set Q2 : pr_state natProtocol → Z → vPred
         := (λ n z, ■ (z = n `mod` Z.pos C) ∗ ES P (NSP P γ) x γ n)%VP.
 
       iApply (GPS_SW_Read (NSP P γ) True Q2 with "[%] kI kS [$NS]");
@@ -673,7 +673,7 @@ Section TicketLock.
       iNext. iViewUp. iIntros (n z) "kS (% & #NSR & Q2)".
       iDestruct "Q2" as "(% & #ES)".
       wp_lam. iNext. wp_op => [Eq|NEq].
-      - iExists 1. iNext. iSplit; first auto. 
+      - iExists 1. iNext. iSplit; first auto.
         case (decide (1 = 0)) => [//|_]. do 2 iNext.
         rewrite yEq H0 in Eq.
 
diff --git a/theories/examples/tstack.v b/theories/examples/tstack.v
index a6f287f7..b506384c 100644
--- a/theories/examples/tstack.v
+++ b/theories/examples/tstack.v
@@ -27,7 +27,7 @@ Section TreiberStack.
       CAS "s" "h" ((int)"n").
 
   Definition tryPop : base.val :=
-    λ: "s", 
+    λ: "s",
       let: "h" := ["s"]_at in
       if: "h" = #0
       then #0
@@ -47,13 +47,13 @@ Section TreiberStack.
 
     Local Open Scope VP.
 
-    Definition Reachable' P: (Z -c> list Z -c> @vPred_ofe Σ) 
+    Definition Reachable' P: (Z -c> list Z -c> @vPred_ofe Σ)
                              -c> Z -c> list Z -c> @vPred_ofe Σ
     := (fun F n A =>
         match A with
         | nil => n = 0
         | v::A' => ■ (0 < n) ∗ Z.to_pos (n + data) ↦ v ∗ P v ∗
-                    ∃ q n', Z.to_pos (n + next) ↦{ q } n' ∗ ▷ F n' A' 
+                    ∃ q n', Z.to_pos (n + next) ↦{ q } n' ∗ ▷ F n' A'
         end).
 
     Instance Reachable'_inhabited: Inhabited (Z -c> list Z -c> @vPred_ofe Σ)
@@ -72,12 +72,12 @@ Section TreiberStack.
 
     Definition Reachable P := fixpoint (Reachable' P).
 
-    Definition ReachableD': (Z -c> list Z -c> @vPred_ofe Σ) 
+    Definition ReachableD': (Z -c> list Z -c> @vPred_ofe Σ)
                              -c> Z -c> list Z -c> @vPred_ofe Σ
     := (fun F n A =>
         match A with
         | nil => n = 0
-        | n'::A' => ■ (0 < n) ∗ ∃ q, Z.to_pos (n + next) ↦{ q } n' ∗ ▷ F n' A' 
+        | n'::A' => ■ (0 < n) ∗ ∃ q, Z.to_pos (n + next) ↦{ q } n' ∗ ▷ F n' A'
         end).
 
     Instance ReachableD'_contractive:
@@ -267,7 +267,7 @@ Section TreiberStack.
           rewrite (fixpoint_unfold (ReachableD') _ _ _ ).
           destruct A as [|h' A'].
           + by iDestruct "Head" as "%".
-          + iDestruct "Head" as "[$ Head]". 
+          + iDestruct "Head" as "[$ Head]".
             iDestruct "Head" as (q) "[Head _]". by iExists _, _.
         - iIntros (?[]??) "_ (_ & Head)". by iApply Head_dup. }
 
@@ -331,4 +331,4 @@ Section TreiberStack.
     Qed.
 
   End proof.
-End TreiberStack.
\ No newline at end of file
+End TreiberStack.
diff --git a/theories/examples/unit_token.v b/theories/examples/unit_token.v
index a7fe45db..feb687cf 100644
--- a/theories/examples/unit_token.v
+++ b/theories/examples/unit_token.v
@@ -25,4 +25,4 @@ Section UnitToken.
   Qed.
 End UnitToken.
 
-Global Notation Tok γ := (⎡own γ (Excl ())⎤%I : vProp _).
\ No newline at end of file
+Global Notation Tok γ := (⎡own γ (Excl ())⎤%I : vProp _).
diff --git a/theories/fork.v b/theories/fork.v
new file mode 100644
index 00000000..e3dccac1
--- /dev/null
+++ b/theories/fork.v
@@ -0,0 +1,24 @@
+From iris.program_logic Require Import weakestpre.
+From iris.base_logic Require Import lib.invariants lib.cancelable_invariants.
+
+From igps Require Import funbi viewpred proofmode weakestpre fractor abs_view.
+From igps Require Export notation bigop.
+From igps.base Require Export ghosts fork.
+
+Lemma wp_fork `{fG : foundationG Σ} E e Φ :
+  ↑physN ⊆ E →
+  ⎡PSCtx⎤ ∗ ▷ Φ #() ∗ ▷ WP e {{ _, True }} ⊢ WP base.Fork e @ E {{ Φ }}.
+Proof.
+  intros.
+  rewrite WP_Def.wp_eq /WP_Def.wp_def.
+  iStartProof (uPred _); iIntros (?) "/=".
+  iIntros "(? & ? & ?)" (???) "?".
+  iApply program_logic.weakestpre.wp_mask_mono; first done.
+  iApply f_fork; iFrame.
+  iSplitL "".
+  - iNext; iIntros (??) "?"; iFrame.
+    iPureIntro; etransitivity; eauto.
+  - iNext; iIntros (? v' ?) "? Ht H".
+    iSpecialize ("H" $! _ with "[%] Ht"); [by etrans|].
+    iApply (program_logic.weakestpre.wp_mono with "H"); auto.
+Qed.
diff --git a/theories/fractor.v b/theories/fractor.v
index 3973e934..87bb07bc 100644
--- a/theories/fractor.v
+++ b/theories/fractor.v
@@ -136,7 +136,7 @@ Section Fractor.
     Qed.
 
   Lemma fractor_alloc (E : coPset) v X l φ Ψ:
-    ↑physN ⊆ E → 
+    ↑physN ⊆ E →
     PSCtx ∗ Info l 1 v ∗ ▷ φ l X ∗ Ψ l 1%Qp X
       ={E}=∗ fractored l φ Ψ 1.
   Proof.
diff --git a/theories/gps/cas.v b/theories/gps/cas.v
index fcb49d67..f4c52496 100644
--- a/theories/gps/cas.v
+++ b/theories/gps/cas.v
@@ -60,7 +60,7 @@ Section CAS.
     rewrite /Consistent in H2.
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
 
     iApply (f_CAS with "[$kI $Seen $Hist]").
     - iPureIntro. exists (VInj v_r), V_r. repeat split; auto.
@@ -90,10 +90,8 @@ Section CAS.
       assert (s_r ⊑ s').
         { by apply (H0  _ _ H6 InVr). }
 
-      assert (VrVrV : Vr ⊑ Vr ⊔ V) by apply join_ext_l.
-
-      move : CSF => [   [-> [Eqv [InV' [Hh' [HDisj [Hi' [HAdj HNAdj]]]]]]]
-                      | [-> [NEqv [Hh' EqVrV]]]].
+      move : CSF => [   [-> [<- [InV' [Hh' [HDisj [Hi' [HAdj HNAdj]]]]]]]
+                      | [-> [NEqv [-> EqVrV]]]].
       + assert (Heb: (s', (v, Vr)) ∈ sEx l (sBE l ζ) (s_x, (v_x, V_x))).
           { rewrite !elem_of_filter. repeat split; last auto.
             - etrans; eauto.
@@ -104,10 +102,7 @@ Section CAS.
 
         rewrite /BE_Inv (bi.big_sepS_delete _ _ _ Heb).
         iDestruct "oBEI" as "[F oBEI]".
-        iSpecialize ("VS1" $! s' V' with "[%]"); [done|].
-
-        rewrite -Eqv monPred_fupd_eq.
-        iMod ("VS1" with "[$F $P]") as (s'') "(% & T & oW)"; first auto.
+        iMod ("VS1" $! _ V' with "[//] [$F $P //]") as (s'') "(% & T & oW)".
 
         iDestruct "oW" as (ζ') "oW".
         iDestruct (Writer_exact with "[$oA $oW]") as %?. subst ζ'.
@@ -119,19 +114,17 @@ Section CAS.
           as "[oW' #oR']".
           { apply union_subseteq_l. }
 
-        iAssert (⌜s' ⊑ s''⌝)%I as "#Hs''". { by done. }
-        iAssert (⌜V' ⊑ V'⌝)%I as "#HV'". { by done. }
-        iSpecialize ("VS2" $! s' s'' V' with "[%]"); [done|].
-        iMod ("VS2" with "[$T oW']") as "(F & Fp & Q)".
-        { iFrame (H10 H11). by iExists _. }
+        iMod ("VS2" $! s' s'' V' with "[//] [$T oW']") as "(F & Fp & Q)".
+        { do 2 iSplit=>//. auto 10. }
+
+        iMod ("Q" $! V' with "[oA' oAX Hist' oBEI oAI F Fp ofX]"); last first.
+        { iApply ("Post" $! s'' true v_n V' V').
+          iAssert ⌜s_r ⊑ s''⌝%I as "$" ; [by iPureIntro; etrans|].
+          auto with iFrame. }
 
-        iMod ("Q" $! V' V' with "[%] [oA' oAX Hist' oBEI oAI F Fp ofX]") as "Q";
-          first done.
-        { iFrame.
-          repeat (iSplit; auto).
-          iNext.
-          iExists ({[s'', (v_n, V')]} ∪ ζ), h', (s_x, (v_x, V_x)).
-          iFrame "oA' oAX Hist'".
+        iFrame. repeat (iSplit; auto). iNext.
+        iExists ({[s'', (v_n, V')]} ∪ ζ), h', (s_x, (v_x, V_x)).
+        iFrame "oA' oAX Hist'".
 
         assert (HC': Consistent ({[s'', (v_n, V')]} ∪ ζ) h').
         { unshelve eapply (Consistent_proper _ ({[s'', (v_n, V')]} ∪ ζ) _ _ _ Hh');
@@ -251,31 +244,22 @@ Section CAS.
           by apply disjoint_singleton_r in HDisj.
         * apply gblock_ends_ins_update; last done.
           apply (hTime_ok_mono _ _ h'); last done. by rewrite Hh'.
-        * apply (hTime_pw_disj_mono _ _ h');
-          [by rewrite Hh'|by apply H8]. }
-
-        iApply ("Post" $! s'' true v_n V' V').
-        iFrame (H5) "Seen' Q oR' HV'".
-        iModIntro; iSplit; by iPureIntro; etrans; eauto.
+        * apply (hTime_pw_disj_mono _  _ h');
+          [by rewrite Hh'|by apply H8].
 
       (* Case: CAS fails *)
       + rewrite /All_Inv (bi.big_sepS_delete _ _ _ InVr).
         iDestruct "oAI" as "[Fp oAI]".
-        rewrite monPred_fupd_eq.
-        iMod ("VSDup" $! _ _ _ with "[%] [$Fp]") as "[Fp Fp']"; [done|done|].
+        iMod ("VSDup" $! s' v Vr with "[//] [$Fp //]") as "[Fp Fp']".
         iCombine "Fp'" "oAI" as "oAI". rewrite -(bi.big_sepS_delete _ _ _ InVr).
-        iSpecialize ("VSF" $! s' v Vr V' with "[%]"); [done|].
-
         iMod (Reader_fork_Auth_2 _ _ {[s', (v, Vr)]} with "[$oA]") as "[oA #oR']".
           { by apply elem_of_subseteq_singleton. }
 
-        iMod ("VSF" with "[$Fp $P oA oAX Hist' oBEI oAI ofX]") as "R".
-        { iFrame; repeat iSplit; auto.
-          iExists ζ, h, (s_x, (v_x, V_x)). subst h'.
-          iFrame "oA oAX Hist' oBEI oAI". repeat (iSplit; first auto); auto. }
+        iMod ("VSF" $! s' v Vr V' with "[//] [- $Fp $P Post Seen']") as "R".
+          { iFrame "ofX oR'". repeat iSplit; auto.
+            iExists ζ, h, (s_x, (v_x, V_x)). auto with iFrame. }
 
-        iApply ("Post" $! s' false v Vr V' with "[- $Seen' $R $oR']").
-        iFrame(H5 H10 VrV' VVr).
+        iApply ("Post" $! s' false v Vr V' with "[- $Seen' $R $oR' //]").
   Qed.
 
   Lemma RawProto_CAS_non_ex_weak
@@ -323,21 +307,16 @@ Section CAS.
       iFrame "VSDup P".
       iSplitL "VS"; last iSplitR "VSF oW"; last iSplitR "oW"; last done; iNext.
       - iIntros (s' V1 VV1) "(Hsr & F & P & oW)".
-        rewrite monPred_fupd_eq.
-        iMod ("VS" $! s' V1 with "[%] [$Hsr $F $P]") as (s'') "?"; first done.
+        iMod ("VS" $! s' V1 with "[//] [$F $P $Hsr]") as (s'') "?".
         iExists s''. by iFrame.
       - iIntros (? ? ?) "_ (_ & _ & (? & ? & Q) & ?)".
-        rewrite monPred_fupd_eq.
         iModIntro; iFrame.
         iIntros (?? HV) "(? & ? & ? & ? & ?)".
         iApply ("Q" $! _ _ with "[//] [-]"). iFrame.
       - iIntros (???? HV) "(? & ? & ? & (? & ?) & ? & ? & ? & ? & ?)".
-        iApply ("VSF" $! _ _ _ _); iFrame; done. }
-
+        iApply ("VSF" $! _ _ _ _ with "[//] [-]"). iFrame. }
     iNext.
     iIntros (s'' b v Vr V') "(VV' & kS' & Hs'' & IF & oR' & VrV & VVr)".
-    iApply ("Post" $! s'' b v Vr V'
-            with "[$VV' $kS' $Hs'' $oR' $VrV $VVr IF]").
-    destruct b; by iFrame.
+    iApply ("Post" $! s'' b v Vr V' with "[$VV' $kS' $Hs'' $oR' $VrV $VVr $IF]").
   Qed.
 End CAS.
diff --git a/theories/gps/fai.v b/theories/gps/fai.v
index 98cfc479..a06b01c5 100644
--- a/theories/gps/fai.v
+++ b/theories/gps/fai.v
@@ -49,7 +49,7 @@ Section FAI.
     rewrite /Consistent in H2.
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
 
     iApply (f_FAI C with "[$kI $kS $Hist]").
     - iPureIntro. exists (VInj v_r), V_r. repeat split; auto.
@@ -90,8 +90,7 @@ Section FAI.
       rewrite /BE_Inv (bi.big_sepS_delete _ _ _ Heb).
       iDestruct "oBEI" as "[F oBEI]".
 
-      iMod ("VS" $! v s' V' with "[%] [$F $P]") as (s'') "(% & Q & F & Fp)";
-        [done|done|].
+      iMod ("VS" $! v s' V' with "[//] [$F $P //]") as (s'') "(% & Q & F & Fp)".
 
       iDestruct "oW" as (ζ') "oW".
       iDestruct (Writer_exact with "[$oA $oW]") as %?. subst ζ'.
diff --git a/theories/gps/fractional.v b/theories/gps/fractional.v
index 5c3b773d..e7d879bd 100644
--- a/theories/gps/fractional.v
+++ b/theories/gps/fractional.v
@@ -59,7 +59,7 @@ Section Fractional.
   Local Close Scope I.
 
   Lemma GPS_FPs_agree l s1 s2 q1 q2 (E : coPset) (HE: ↑fracN .@ l ⊆ E):
-    (vGPS_FP l s1 q1 ∗ vGPS_FP l s2 q2 
+    (vGPS_FP l s1 q1 ∗ vGPS_FP l s2 q2
      ={E}=∗  ⌜s1 ⊑ s2 ∨ s2 ⊑ s1⌝ ∗ vGPS_FP l s1 q1 ∗ vGPS_FP l s2 q2)%I.
   Proof.
     iStartProof (uPred _).
@@ -125,7 +125,7 @@ Section Fractional.
 
   Lemma GPS_FP_Read' l q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
         s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-      {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ → 
+      {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ →
                   (∀ v, F_past l s' v ∗ P ∗ vGPS_FP l s' q ={E}=∗ Q s' v)
                   ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
            ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
@@ -153,11 +153,11 @@ Section Fractional.
             with "[VS $oI $kI $kS $oP $oR VSDup Writer Hex HClose]"); [solve_ndisj|auto|..].
     { etransitivity; eauto. }
     { iSplitR "VSDup"; iNext; last by iSpecialize ("VSDup" $! _).
-      iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
-      { etrans; eauto. }
+      iIntros (s' j ? Hs').
+      iDestruct ("VS" $! _ Hs') as "[VS|VS]"; last first.
+      { iRight. iIntros (?????). iApply ("VS" $! _ _ _ with "[%] [%//]").
+        repeat etrans=>//. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (RFRaw s' j0) with "[Writer oI oR' Hex]").
       { iSplitL "oI Writer".
         + iNext. iExists γ, γ_x, ζ. iFrame; auto.
@@ -165,20 +165,20 @@ Section Fractional.
           iExists Vr, v. rewrite monPred_in_eq /=.
           iSplit; last iFrame "oR'"; iPureIntro; auto.
           do 4 (etrans; eauto). }
-      iMod ("VS" $! _ _ with "[%] [-]") as "VS"; [done | iFrame |].
+      iMod ("VS" $! _ j0 with "[% ] [-]") as "VS"; [|by iFrame|].
+      { repeat (etrans; eauto). }
       iApply fancy_updates.fupd_intro_mask; [solve_ndisj | done]. }
     iNext.
     iIntros (s' v Vr V') "(% & kS' & #Hs' & oQ & #oR' & % & %)".
     iFrame.
-    rewrite monPred_fupd_eq /=.
     iApply ("Post" $! _ _ _ with "[%]").
     { etrans; eauto. }
-    iFrame "Hs' oQ".
-  Qed.
+    by iFrame "Hs'".
+Qed.
 
   Lemma GPS_FP_Read l q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
         s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-    ⎡PSCtx⎤ ⊢ {{{ (▷ ∀ s', ⌜s ⊑ s'⌝ → 
+    ⎡PSCtx⎤ ⊢ {{{ (▷ ∀ s', ⌜s ⊑ s'⌝ →
                   (∀ v, F_past l s' v ∗ P ={E ∖ ↑fracN .@ l}=∗ Q s' v)
                   ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
            ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
@@ -233,27 +233,24 @@ Section Fractional.
     iApply (RawProto_Read IP l P (fun s v => |={E ∖ ↑fracN .@ l, E}=> Q s v)%I _ _ _ V_r v_r γ γ_x
             with "[VS $oI $kI $kS $oP $oR VSDup Writer Hex HClose abs]"); [solve_ndisj|done|..].
     { iNext. iSplitR "VSDup"; last by iSpecialize ("VSDup" $! _).
-      iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
+      iIntros (s' j ? Hs').
+      iDestruct ("VS" $! _ j with "[%] [%//]") as "[VS|VS]"; last by iRight.
       { etrans; eauto. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (RFRaw s V0) with "[Writer oI oR' Hex]").
       { iSplitL "oI Writer".
         + iNext. iExists γ, γ_x, ζ. iFrame; auto.
         + iExists γ, γ_x, e_x. iFrame "Hex". iSplitL ""; first auto.
           iExists V_r, v_r. rewrite monPred_in_eq /=; auto. }
-      iMod ("VS" $! _ _ with "[%] [-]") as "VS"; first done.
-      { iFrame.
-        iExists V0; iFrame. }
+      iMod ("VS" $! _ j0 with "[//] [-]") as "VS".
+      { iFrame. iExists V0; iFrame. }
       iApply fancy_updates.fupd_intro_mask; [solve_ndisj | done]. }
     iNext.
     iIntros (s' v Vr V') "(% & kS' & #Hs' & oQ & #oR' & % & %)".
     iFrame.
-    rewrite monPred_fupd_eq /=.
     iApply ("Post" $! _ _ _ with "[%]").
     { etrans; eauto. }
-    iFrame "Hs' oQ".
+    by iFrame "Hs'".
   Qed.
 
   Lemma GPS_FP_Read_abs `{absG : absViewG Σ}
@@ -279,7 +276,7 @@ Section Fractional.
   Qed.
 
   Lemma GPS_FP_Write l q s s' v (E : coPset)
-        (P: vPred) (Q : pr_state Prtcl → vPred) 
+        (P: vPred) (Q : pr_state Prtcl → vPred)
         (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E) (HFinal : final_st s'):
       {{{ ⎡PSCtx⎤ ∗ ▷ (P ∗ vGPS_FP l s' q ={E ∖ ↑fracN .@ l}=∗
                                                  Q s' ∗ ▷ (F l s' v) ∗ ▷ F_past l s' v)
@@ -318,7 +315,6 @@ Section Fractional.
       iExists V', v. iFrame "oR'". rewrite monPred_in_eq /=. iSplit; first done.
       iPureIntro. do 4 (etrans; eauto).
     }
-    rewrite monPred_fupd_eq.
     assert (V ⊑ V') by (do 2 (etrans; eauto)).
     iMod ("VS" $! _ with "[%] [oP oW]") as "($ & oF & oFp)"; first done.
     { rewrite vGPS_FP_eq /=; iFrame. }
@@ -333,7 +329,7 @@ Section Fractional.
         (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E) :
     let mod1C (v: Z) := ((v + 1) `mod` Z.pos C)%Z in
       {{{ ⎡PSCtx⎤ ∗ ▷ (∀ v s', ⌜s ⊑ s'⌝ ∗ F l s' v ∗ P
-                     ={E ∖ ↑fracN .@ l}=∗ ∃ s'', ⌜s' ⊑ s''⌝ 
+                     ={E ∖ ↑fracN .@ l}=∗ ∃ s'', ⌜s' ⊑ s''⌝
                                                    ∗ Q s'' v
                                                    ∗ ▷ (F l s'' (mod1C v)) ∗ ▷ F_past l s'' (mod1C v))
             ∗ P ∗ vGPS_FP l s q }}}
@@ -365,7 +361,7 @@ Section Fractional.
     { iSplitR "wTok"; last by iExists _.
       iNext.
       iIntros (?? j ?) "Hs'".
-      iApply ("VS" $! _ _ _ with "[%]"); last done.
+      iApply ("VS" $! _ _ _ with "[%] [-//]").
       etrans; eauto. }
     iNext.
     iIntros (s'' v Vr V')
@@ -384,7 +380,7 @@ Section Fractional.
   Qed.
 
   Lemma GPS_FP_CAS' l q s v_o v_n (E : coPset)
-        (P : vPred) (Q: pr_state Prtcl → vPred) (R : pr_state Prtcl → Z → vPred) 
+        (P : vPred) (Q: pr_state Prtcl → vPred) (R : pr_state Prtcl → Z → vPred)
         (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E) :
       {{{ ⎡PSCtx⎤ ∗ ▷ (∀ s', ⌜s ⊑ s'⌝ ∗ F l s' v_o ∗ P
                   ={E ∖ ↑fracN .@ l}=∗ ∃ s'', ⌜s' ⊑ s''⌝ ∗ ▷ (F l s'' v_n) ∗ ▷ F_past l s'' v_n
@@ -424,7 +420,6 @@ Section Fractional.
       [solve_ndisj|solve_ndisj|solve_ndisj| |].
     { iSplitL "VS"; last iSplitL "VSDup"; last iSplitL "VSF"; last (by iFrame; iExists _); iNext.
       - iIntros (s' j ?) "(Hs' & F & oP & HClose)".
-        rewrite monPred_fupd_eq.
         iMod ("VS" $! _ _ with "[%] [$Hs' $F $oP]") as (s'') "(? & ? & ? & Q)".
         { etrans; eauto. }
         iExists s''; iFrame.
@@ -437,12 +432,11 @@ Section Fractional.
             iExists Vr, v_n. rewrite monPred_in_eq /=.
             iSplit; last iFrame "oR'"; iPureIntro; auto.
             do 4 (etrans; eauto). }
-        iMod ("Q" $! _ with "[%] oW"); first done.
+        iMod ("Q" $! j0 with "[//] [$oW]").
         iApply fancy_updates.fupd_intro_mask; [solve_ndisj | done].
       - by iSpecialize ("VSDup" $! _).
       - iIntros (s' v Vr j ?) "(Hs' & Hv & Fp & (oP & HClose) & % & % & oI & ex & Writer & oR')".
         iDestruct "Writer" as (ζ') "Writer".
-        rewrite monPred_fupd_eq.
         iMod ("HClose" $! (RFRaw s' j) with "[oI ex Writer oR']") as "oW".
         { iSplitL "oI Writer".
           + iNext. iExists γ, γ_x, ζ'. iFrame; auto.
@@ -460,15 +454,14 @@ Section Fractional.
     iFrame.
     iSpecialize ("Post" $! s'' b v _ with "[%]").
     { etrans; eauto. }
-    rewrite monPred_fupd_eq.
-    destruct b; iMod "IF"; iApply "Post"; auto.
+    iApply "Post". iFrame "Hs'". by destruct b.
   Qed.
 
   Lemma GPS_FP_CAS l q s v_o v_n (E : coPset)
-        (P : vPred) (Q: pr_state Prtcl → vPred) (R : pr_state Prtcl → Z → vPred) 
+        (P : vPred) (Q: pr_state Prtcl → vPred) (R : pr_state Prtcl → Z → vPred)
         (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E) :
     ⎡PSCtx⎤ ⊢ {{{ ▷ (∀ s', ⌜s ⊑ s'⌝ ∗ F l s' v_o ∗ P
-                  ={E ∖ ↑fracN .@ l}=∗ ∃ s'', ⌜s' ⊑ s''⌝ 
+                  ={E ∖ ↑fracN .@ l}=∗ ∃ s'', ⌜s' ⊑ s''⌝
                                                 ∗ Q s'' ∗ ▷ (F l s'' v_n) ∗ ▷ F_past l s'' v_n)
            ∗ ▷ (∀ s' v, ⌜s ⊑ s'⌝ ∗ ⌜v ≠ v_o⌝ ∗ ▷ F_past l s' v ∗ P
                         ={E ∖ ↑fracN .@ l}=∗ R s' v)
@@ -560,7 +553,7 @@ Arguments GPS_FP_splitjoin [_ _ _ _ _ _ _ _ _ _ _].
 
 Notation "'[FP' l 'in' s | p ]_ q" :=
   (vGPS_FP p l s q)
-    (at level 0, 
+    (at level 0,
      format "[FP  l  in  s  |  p  ]_ q").
 
 (* TODO: fix instances to compile *)
@@ -579,4 +572,4 @@ Qed.
 
 Global Instance vGPS_FP_proper {Σ Prtcl fG PF gpsG persG}:
   Proper ((≡) ==> (=) ==> (=) ==> (≡)) (@vGPS_FP Σ Prtcl fG PF gpsG persG q).
-Proof. solve_proper. Qed.
\ No newline at end of file
+Proof. solve_proper. Qed.
diff --git a/theories/gps/init.v b/theories/gps/init.v
index 8c125b17..21573e3f 100644
--- a/theories/gps/init.v
+++ b/theories/gps/init.v
@@ -4,7 +4,7 @@ From igps.gps Require Import shared.
 From igps.base Require Import na_write.
 
 Section Init_Strong.
-  Context {Σ} `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF} 
+  Context {Σ} `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF}
           (IP : gname -> interpC Σ Prtcl) (l : loc).
   Set Default Proof Using "Type".
   Notation state_sort := (pr_state Prtcl * (Z * View))%type.
@@ -35,7 +35,7 @@ Section Init_Strong.
     iIntros (Φ) "(#kI & kS & oH & #Ha) Post".
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
 
     iApply (f_write_na with "[$kI $kS $oH $Ha]").
     iNext.
@@ -113,17 +113,13 @@ Section Init.
         ∗ Reader γ {[s, (v, V')]}
         ∗ Writer γ {[s, (v, V')]} }}}.
   Proof.
-    iIntros (Φ) "(#kI & kS & oH & #Ha & Fp) Post".
+    iIntros (Φ) "(#kI & kS & oH & #Ha & Fp & Fpast) Post".
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
     iApply (RawProto_init_strong (λ _, IP) l (λ _, s)
               with "[$kI $kS $oH $Ha]");[done|].
     iNext. iIntros (V' γ γ_x) "(% & kS' & Inv & ?)".
-    iApply ("Post" $! V' γ γ_x). iFrame (H) "∗ #". iApply "Inv".
-    rewrite ![monPred_car _ V]monPred_mono; first (iModIntro; by iNext); auto.
-    (* TODO: why doesn't this work anymore?
-    rewrite [V in (_ V ∗ _ V)%I]H. iModIntro. by iNext.
-     *)
+    iApply ("Post" $! V' γ γ_x). iFrame (H) "∗ #". iApply "Inv". by iFrame.
   Qed.
-End Init.
\ No newline at end of file
+End Init.
diff --git a/theories/gps/plain.v b/theories/gps/plain.v
index 1618a750..684f6e72 100644
--- a/theories/gps/plain.v
+++ b/theories/gps/plain.v
@@ -37,18 +37,18 @@ Section Plain.
   Definition RPRaw p s V l n : iProp :=
     (∃ (γ γ_x γx: gname) e_x,
       ⌜n = encode (p,γ,γ_x,γx)⌝ ∗ own.own γx (DecAgree e_x)
-      ∗ ∃ Vr v, monPred_car (monPred_in Vr) V
+      ∗ ∃ Vr v, ⌜Vr ⊑ V⌝
               ∗ ⌜st_view e_x !! l ⊑ (Vr !! l : option positive)⌝ ∗ Reader γ {[s, (v, Vr)]})%I.
 
   Program Definition vGPS_PP_def l p s : vProp Σ :=
     MonPred (λ V, persisted l (gpsPRaw p) (RPRaw p s V)) _.
   Next Obligation.
-  intros. intros V V' Mono. 
-  iIntros "RP". iDestruct "RP" as (X) "(RPV & ?)".
-  iDestruct "RPV" as (γ γ_x γx e_x) "(? & ? & RA)".
-  iDestruct "RA" as (Vr v) "(? & % & ?)".
-  iExists X. iFrame. iExists γ, γ_x, γx, e_x. iFrame.
-  iExists Vr, v. by iFrame.
+    intros. intros V V' Mono.
+    iIntros "RP". iDestruct "RP" as (X) "(RPV & ?)".
+    iDestruct "RPV" as (γ γ_x γx e_x) "(% & ? & RA)".
+    iDestruct "RA" as (Vr v) "(% & % & ?)".
+    iExists X. iFrame. iExists γ, γ_x, γx, e_x. iFrame "%∗".
+    iExists Vr, v. iFrame "%∗". iPureIntro. by etrans.
   Qed.
   Definition vGPS_PP_aux : seal vGPS_PP_def. by eexists. Qed.
   Definition vGPS_PP := unseal vGPS_PP_aux.
@@ -71,10 +71,8 @@ Section Plain.
   Lemma GPS_PPs_agree l p s1 s2 (E : coPset) (HE: ↑persist_locN .@ l ⊆ E):
     vGPS_PP l p s1 ∗ vGPS_PP l p s2 ={E}=∗ ⌜s1 ⊑ s2 ∨ s2 ⊑ s1⌝.
   Proof.
-    split => ?.
-    rewrite vGPS_PP_eq /vGPS_PP_def /=.
-    iIntros "RP". rewrite /bi_sep /= /sbi_sep /= monPred_sep_eq /=. rewrite persistor_splitjoin.
-    rewrite monPred_fupd_eq /=.
+    iStartProof (uPred _). rewrite vGPS_PP_eq /vGPS_PP_def /=.
+    iIntros (?) "RP". rewrite persistor_splitjoin.
     iMod (persistor_open with "RP") as (X) "(oI & (oR1 & oR2) & HClose)";
       first auto.
     iDestruct "oR1" as (γ1 γ_x1 γx1 e_x1) "(% & #oD1 & #RA1)".
@@ -127,7 +125,7 @@ Section Plain.
           with "[-Post $kIP]") as "PP"; [auto|auto|..].
       { iFrame "kI Info oI oW ex oEx".
         repeat iExists _. iFrame "oEx". iSplitL; [done|].
-        iExists _, _; iFrame "oR". rewrite monPred_in_eq /=. by iSplit. }
+        iExists _, _; iFrame "oR". by iSplit. }
       iModIntro; iApply ("Post" $! _ with "[%]").
       { etrans; eauto. }
     rewrite vGPS_PP_eq /=; done.
@@ -135,7 +133,7 @@ Section Plain.
 
   Lemma GPS_PP_Read' l p (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑persist_locN .@ l ⊆ E):
-  {{{ ⎡PSCtx⎤ ∗ ▷(∀ s', ⌜s ⊑ s'⌝ → 
+  {{{ ⎡PSCtx⎤ ∗ ▷(∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past l s' v ∗ vGPS_PP l p s' ∗ P ={E}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P
                             ={E ∖ ↑persist_locN .@ l}=∗ False))
@@ -152,7 +150,6 @@ Section Plain.
     iMod (persistor_open with "[$oR]")
       as (X) "(gpsRaw & RRaw & HClose)"; first auto.
     iDestruct "RRaw" as (γ γ_x γx e_x) "(% & #Hex & RA)".
-    rewrite monPred_in_eq /=.
     iDestruct "gpsRaw" as (γ' γ_x' γx' e_x' ζ)
                           "(>% & oI & Writer & >Ex & >#Hex')".
     subst X. apply encode_inj in H0. inversion H0. subst γ' γ_x' γx'.
@@ -168,31 +165,27 @@ Section Plain.
             with "[VS $oI $kI $kS $oP $oR VSDup Writer Ex HClose]"); [solve_ndisj|auto|..].
     { etransitivity; eauto. }
     { iSplitR "VSDup"; iNext; last by iSpecialize ("VSDup" $! _).
-      iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
-      { etrans; eauto. }
+      iIntros (s' j ? Hs'). iDestruct ("VS" $! s' Hs') as "[VS|VS]"; last first.
+      { iRight. iIntros (?????). iApply ("VS" $! _ _ _ with "[%] [//]").
+        repeat etrans=>//. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (RPRaw p s' j0) with "[Writer oI oR' Ex]").
       { iSplit.
         + iNext. iExists γ, γ_x, γx, e_x, ζ. iFrame; auto.
         + iExists γ, γ_x, γx, e_x. iFrame "Hex". iSplitL ""; first auto.
-          iExists Vr, v. rewrite monPred_in_eq /=.
+          iExists Vr, v.
           iSplit; last iFrame "oR'"; iPureIntro; auto.
           do 4 (etrans; eauto). }
-      iMod ("VS" $! _ _ with "[%] [-]") as "VS"; first done; first iFrame.
+      iMod ("VS" $! _ j0 with "[%] [-]") as "VS"; [solve_jsl|by iFrame|].
       iApply fancy_updates.fupd_intro_mask; [solve_ndisj | done]. }
     iNext.
-    iIntros (s' v Vr V') "(% & kS' & #Hs' & oQ & #oR' & % & %)".
-    assert (V2 ⊑ V') as Hle by (etransitivity; eauto).
-    iFrame.
-    rewrite monPred_fupd_eq /=.
-    iApply ("Post" $! _ _ _ with "[%]"); first done. iFrame "Hs' oQ".
+    iIntros (s' v Vr V') "(% & kS' & % & oQ & #oR' & % & %)".
+    iFrame. iApply ("Post" $! s' _ _ with "[//]").  by iFrame "%".
   Qed.
 
   Lemma GPS_PP_Read l p (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑persist_locN .@ l ⊆ E):
-  ⎡PSCtx⎤ ⊢ {{{ ▷(∀ s', ⌜s ⊑ s'⌝ → 
+  ⎡PSCtx⎤ ⊢ {{{ ▷(∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past l s' v ∗ P ={E ∖ ↑persist_locN .@ l}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P
                             ={E ∖ ↑persist_locN .@ l}=∗ False))
@@ -212,7 +205,7 @@ Section Plain.
   Qed.
 
   Lemma GPS_PP_Write l p s s' v (E : coPset)
-      (P: vPred) (Q : pr_state Prtcl → vPred) 
+      (P: vPred) (Q : pr_state Prtcl → vPred)
       (HN : ↑physN ⊆ E) (HNl: ↑persist_locN .@ l ⊆ E)
       (HFinal: final_st s'):
    {{{ ⎡PSCtx⎤ ∗ ▷ (P ∗ vGPS_PP l p s' ={E ∖ ↑persist_locN .@ l}=∗
@@ -236,10 +229,9 @@ Section Plain.
     iDestruct (own.own_valid with "He") as %He%dec_agree_op_inv.
     inversion He. subst e_x'. destruct e_x as [s_x [v_x V_x]].
 
-    rewrite monPred_in_eq /=.
     iDestruct "RA" as (V_r v_r) "(% & % & #oR)".
 
-    iApply (RawProto_Write_non_ex IP l 
+    iApply (RawProto_Write_non_ex IP l
               s s_x s' _ _ V_r V_x 1%Qp v v_r v_x γ γ_x _ ζ
             with "[$oI $kI $kS $ex $oR $wTok]"); [solve_ndisj|auto|auto|auto|].
     { do 2 (etransitivity; eauto). }
@@ -252,11 +244,9 @@ Section Plain.
     { iApply "HT".
       iExists γ, γ_x, γx, (s_x, (v_x, V_x)).
       repeat iSplit; [auto|auto|].
-      iExists V', v. rewrite monPred_in_eq /=. repeat iSplit; [auto| |auto].
+      iExists V', v. repeat iSplit; [auto| |auto].
       iPureIntro. do 4 (etransitivity; eauto). }
-    rewrite monPred_fupd_eq.
-    assert (V ⊑ V') by (do 2 (etrans; eauto)).
-    iMod ("VS" $! V' with "[%] [$oP $oPP]") as "($ & oF & oFp)"; first done.
+    iMod ("VS" $! V' with "[%] [$oP $oPP]") as "($ & oF & oFp)"; [solve_jsl|].
     iSpecialize ("oI" with "[$oF $oFp]").
     iFrame "oPP". iApply "HClose".
     iNext. iExists γ, γ_x, γx, (s_x, (v_x, V_x)), ({[s', (v, V')]} ∪ ζ).
@@ -271,7 +261,7 @@ Section Plain.
                   (vGPS_PP l p s'' ={E}=∗ Q s''))
       ∗ ▷(∀ s' v, ⌜s ⊑ s'⌝ ∗ ⌜v ≠ v_o⌝ ∗ ▷ F_past l s' v ∗ vGPS_PP l p s' ∗ P
                       ={E}=∗ R s' v)
-      ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v 
+      ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
                   ={E ∖ ↑persist_locN .@ l}=∗ F_past l s' v ∗ F_past l s' v )
       ∗ P ∗ vGPS_PP l p s }}}
     (CAS #l #v_o #v_n) @ E
@@ -293,7 +283,6 @@ Section Plain.
     iDestruct (own.own_valid with "He") as %He%dec_agree_op_inv.
     inversion He. subst e_x'. destruct e_x as [s_x [v_x V_x]].
 
-    rewrite monPred_in_eq /=.
     iDestruct "RA" as (V_r v_r) "(% & % & #oR)".
 
     assert (V_r ⊑ V3) by (do 2 (etransitivity; eauto)).
@@ -309,7 +298,6 @@ Section Plain.
       [solve_ndisj|solve_ndisj|solve_ndisj| |].
     { iSplitL "VS"; last iSplitL "VSDup"; last iSplitL "VSF"; last (by iFrame; iExists _); iNext.
       - iIntros (s' j ?) "(Hs' & F & oP & HClose)".
-        rewrite monPred_fupd_eq.
         iMod ("VS" $! _ _ with "[%] [$Hs' $F $oP]") as (s'') "(? & ? & ? & Q)".
         { etrans; eauto. }
         iExists s''; iFrame.
@@ -319,20 +307,19 @@ Section Plain.
         { iSplit.
           + iNext. iExists γ, γ_x, γx, (s_x, (v_x, V_x)), ζ'. iFrame; auto.
           + iExists γ, γ_x, γx, (s_x, (v_x, V_x)). iFrame "Hex". iSplitL ""; first auto.
-            iExists Vr, v_n. rewrite monPred_in_eq /=.
+            iExists Vr, v_n.
             iSplit; last iFrame "oR'"; iPureIntro; auto.
             do 4 (etrans; eauto). }
-        iMod ("Q" $! _ with "[%] oW"); first done.
+        iMod ("Q" $! _ with "[//] [oW //]").
         iApply fancy_updates.fupd_intro_mask; [solve_ndisj | done].
       - by iSpecialize ("VSDup" $! _).
       - iIntros (s' v Vr j ?) "(Hs' & Hv & Fp & (oP & HClose) & % & % & oI & ex & Writer & oR')".
         iDestruct "Writer" as (ζ') "Writer".
-        rewrite monPred_fupd_eq.
         iMod ("HClose" $! (RPRaw p s' j) with "[oI ex Writer oR']") as "oW".
         { iSplit.
           + iNext. iExists γ, γ_x, γx, (s_x, (v_x, V_x)), ζ'. iFrame; auto.
           + iExists γ, γ_x, γx, (s_x, (v_x, V_x)). iFrame "Hex". iSplitL ""; first auto.
-            iExists Vr, v. rewrite monPred_in_eq /=.
+            iExists Vr, v.
             iSplit; last iFrame "oR'"; iPureIntro; auto.
             do 4 (etrans; eauto). }
         iMod ("VSF" $! _ _ _ with "[%] [$Hs' $Hv $Fp $oW $oP]").
@@ -345,8 +332,7 @@ Section Plain.
     iFrame.
     iSpecialize ("Post" $! s'' b v _ with "[%]").
     { etrans; eauto. }
-    rewrite monPred_fupd_eq /=.
-    destruct b; iMod "IF"; iApply "Post"; auto.
+    iApply "Post". iFrame. by destruct b.
   Qed.
 
   Lemma GPS_PP_CAS l p s v_o v_n (E : coPset)
@@ -357,7 +343,7 @@ Section Plain.
                         ∗ Q s'' ∗ ▷ (F l s'' v_n) ∗ ▷ F_past l s'' v_n)
       ∗ ▷(∀ s' v, ⌜s ⊑ s'⌝ ∗ ⌜v ≠ v_o⌝ ∗ ▷ F_past l s' v ∗ P
                       ={E ∖ ↑persist_locN .@ l}=∗ R s' v)
-      ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v 
+      ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
                   ={E ∖ ↑persist_locN .@ l}=∗ F_past l s' v ∗ F_past l s' v )
       ∗ P ∗ vGPS_PP l p s }}}
     (CAS #l #v_o #v_n) @ E
@@ -408,7 +394,6 @@ Section Plain.
     iDestruct (own.own_valid with "He") as %He%dec_agree_op_inv.
     inversion He. subst e_x'. destruct e_x as [s_x [v_x V_x]].
 
-    rewrite monPred_in_eq /=.
     iDestruct "RA" as (V_r v_r) "(% & % & #oR)".
 
     assert (V ⊑ V3) by (do 2 (etrans; eauto)).
@@ -418,14 +403,12 @@ Section Plain.
       [auto|auto|solve_ndisj| |].
     { etrans; eauto. }
     { iSplitR "wTok"; last by iExists _.
-      iNext; iIntros (?? j ?).
-      iApply ("VS" $! _ _ _ with "[%]").
-      etrans; eauto. }
+      iNext; iIntros (?? j ?) "?".
+      iMod ("VS" $! _ _ j with "[%] [-//]"); [solve_jsl|done]. }
     iNext.
     iIntros (s'' v Vr V')
             "(% & kS' & Hs' & oI & Q & ex & #oR' & % & % & wTok)".
-    iFrame; iApply ("Post" $! _ _ _ with "[%]").
-    { etrans; eauto. }
+    iFrame; iApply ("Post" $! _ _ _ with "[%]"); [solve_jsl|].
     iFrame.
     iApply "HClose". iSplitL "oI ex wTok".
     + iNext. iDestruct "wTok" as (ζ') "wTok".
@@ -433,7 +416,7 @@ Section Plain.
       by iFrame "oI wTok ex Hex".
     + iExists γ, γ_x, γx, (s_x, (v_x, V_x)).
       repeat iSplit; [auto|auto|].
-      iExists Vr, _. rewrite monPred_in_eq /=. repeat iSplit; [auto| |auto].
+      iExists Vr, _. repeat iSplit; [auto| |auto].
       iPureIntro. do 4 (etrans; eauto).
   Qed.
 
@@ -535,6 +518,6 @@ Global Instance vGPS_PP_proper {Σ Prtcl fG PF gpsG persG agreeG PType EqPT Coun
 Proof.
   intros ? ? H1 ? ? H2 ? ? H3 ? ? H4. constructor => ?. subst.
   rewrite !vGPS_PP_eq /=. apply persisted_proper.
-  - repeat intro. exact: gpsPPRaw_proper. 
+  - repeat intro. exact: gpsPPRaw_proper.
   - repeat intro. by f_equiv.
-Qed.
\ No newline at end of file
+Qed.
diff --git a/theories/gps/raw.v b/theories/gps/raw.v
index c2b4a38c..ae1f1475 100644
--- a/theories/gps/raw.v
+++ b/theories/gps/raw.v
@@ -1 +1 @@
-From igps.gps Require Export shared init read write cas fai.
\ No newline at end of file
+From igps.gps Require Export shared init read write cas fai.
diff --git a/theories/gps/read.v b/theories/gps/read.v
index d3ba46e0..c8b001f1 100644
--- a/theories/gps/read.v
+++ b/theories/gps/read.v
@@ -41,7 +41,7 @@ Section Read.
 
     iDestruct (Reader_Auth_sub with "[$oA $oR]") as %?%elem_of_subseteq_singleton.
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
     iApply (f_read_at with "[$kI $Seen $Hist]").
     - iPureIntro. exists (VInj v_r), V_r. repeat split; last auto.
       rewrite -H2 elem_of_map_gset. by exists (s_r, (v_r, V_r)).
@@ -68,16 +68,14 @@ Section Read.
       assert (s_r ⊑ s').
         { by apply (H0  _ _ H6 In). }
 
-      rewrite monPred_fupd_eq.
-      iDestruct ("VS" $! s' V  with "[%] [%]") as "[VS|VS]"; [done|auto|..].
+      iDestruct ("VS" $! s' with "[//]") as "[VS|VS]".
       + rewrite /All_Inv (bi.big_sepS_delete _ _ _ In).
         iDestruct "oAI" as "[Fp oAI]".
-        iMod ("VSDup" $! _ _ _ with "[%] [$Fp]") as "[Fp Fp']"; [done|auto|].
+        iMod ("VSDup" $! _ _ Vr with "[//] [$Fp //]") as "[Fp Fp']".
         iCombine "Fp'" "oAI" as "oAI".
         rewrite -(bi.big_sepS_delete _ _ _ In).
-        iSpecialize ("VS" $! v Vr V' with "[%]"); [solve_jsl|].
         assert ((V !! l : option positive) ⊑ Vr !! l) by (rewrite H9; done).
-        iMod ("VS" with "[$Fp $P oA o1 Hist oBEI oAI]") as "Q".
+        iMod ("VS" $! v Vr V' with "[//] [$Fp $P oA o1 Hist oBEI oAI]") as "Q".
         { repeat (iSplit; auto).
           iExists ζ', h, e_x. iFrame. repeat (iSplit; first auto); auto. }
         iApply ("Post" $! s' v Vr V').
@@ -111,14 +109,8 @@ Section Read.
 
         rewrite /BE_Inv (bi.big_sepS_delete _ _ _ Heb).
         iDestruct "oBEI" as "[F oBEI]".
-
-        iSpecialize ("VS" $! (st_prst e) (st_val e) (st_view e ⊔ V ⊔ Vr) with "[%] Hse"); [|].
-        { etrans; [apply join_ext_r|apply join_ext_l].
-          (* TODO: make solve_jsl do this *) }
-
-        iMod ("VS" $! _ with "[%] [F P]") as "?"; [reflexivity| |by iExFalso].
-        iFrame "F P". iPureIntro.
-        etrans; last by apply join_ext_l. apply join_ext_r.
+        iMod ("VS" $! (st_prst e) (st_val e) (V ⊔ (st_view e))
+             with "[%] [$Hse] [$F $P]") as "[]". solve_jsl.
   Qed.
 
 
@@ -147,7 +139,7 @@ Section Read.
     iDestruct (Reader_Writer_sub with "[$oW $oR]") as %?%elem_of_subseteq_singleton.
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
     iApply (f_read_at with "[$kI $kS $Hist]").
     - iPureIntro. exists (VInj v), V. repeat split; last auto.
       rewrite -H2 elem_of_map_gset. by exists (s, (v, V)).
@@ -155,8 +147,7 @@ Section Read.
       iDestruct "Pure" as (V'') "(% & % & %)".
       rewrite /All_Inv (bi.big_sepS_delete _ _ _ H6).
       iDestruct "oAI" as "[Fp oAI]".
-      rewrite monPred_fupd_eq.
-      iMod ("VSDup" $! _ _ with "[%] [$Fp]") as "[Fp Fp']"; [done|].
+      iMod ("VSDup" $! v V with "[//] [$Fp]") as "[Fp Fp']".
       iCombine "Fp'" "oAI" as "oAI". rewrite -(bi.big_sepS_delete _ _ _ H6).
 
       iMod (Hist_hTime_ok _ _ _ HEN with "[$kI $Hist]") as "(Hist & %)".
@@ -170,10 +161,9 @@ Section Read.
       apply (HDisj (VInj v, V) (VInj v', V'')) in H12; last auto; last first.
         { rewrite -H2 elem_of_map_gset. by exists (s, (v,V)). }
       inversion H12. subst v' V''.
-      rewrite [V in monPred_car (F_past _ _) V]H8.
 
       iApply ("Post" $! v V' with "[$kS' $ofX $oR $oW $Fp oA oAX oAI Hist oBEI]").
       iFrame (H5).
       iExists ζ, h, (s, (v, V)). iFrame (H H0 H1 H2 H3 H4) "oA oAX oAI Hist oBEI".
   Qed.
-End Read.
\ No newline at end of file
+End Read.
diff --git a/theories/gps/recursive.v b/theories/gps/recursive.v
index 583f2a06..9dd08d05 100644
--- a/theories/gps/recursive.v
+++ b/theories/gps/recursive.v
@@ -9,7 +9,7 @@ From igps Require Export notation na.
 From igps.gps Require Import shared inst_shared.
 
 Section RecInterp.
-  Context {Σ : gFunctors} (Prtcl : protocolT). 
+  Context {Σ : gFunctors} (Prtcl : protocolT).
   Set Default Proof Using "Type".
 
   Global Instance IntType_cofe : Cofe (interpC Σ Prtcl).
@@ -61,14 +61,12 @@ Canonical funC_interpC Σ Prtcl :=
   mk_funC_of (type_of_interp Σ Prtcl)
              _
              _
-             (λ F HF, fixpoint (type_ofe_conv _ F) (Hf := HF))
-.
+             (λ F HF, fixpoint (type_ofe_conv _ F) (Hf := HF)).
+Local Existing Instances funC_Cofe funC_Inhabited.
 Canonical funC_param (f : funC_of) (X : ofeT) :=
   mk_funC_of (type_of_param X (funC_ofe f))
-             (@ofe_fun_cofe _ _ (funC_Cofe f))
-             (populate (λ _, @inhabitant _ (funC_Inhabited f))) 
-             (λ F HF,
-              @fixpoint _ (@ofe_fun_cofe _ _ (funC_Cofe f)) (populate (λ _, @inhabitant _ (funC_Inhabited f))) (type_ofe_conv _ F) HF).
+             _ _
+             (λ F, @fixpoint _ _ _ (type_ofe_conv _ F)).
 
 Tactic Notation "exvar" constr(type) tactic(cnt) :=
   let name := fresh in
@@ -120,7 +118,7 @@ Ltac unfold_lemma int :=
   let orig_int := int in
   let rec unfold_lemma int :=
   match int with
-  | funC_fix ?f ?F ?C => 
+  | funC_fix ?f ?F ?C =>
     let rec go f F unf int :=
         (* idtac f; idtac F; *)
         lazymatch f with
diff --git a/theories/gps/shared.v b/theories/gps/shared.v
index 39b249a5..953638d3 100644
--- a/theories/gps/shared.v
+++ b/theories/gps/shared.v
@@ -408,7 +408,7 @@ Qed.
 
 Global Instance ProtoInv_proper {Σ Prtcl PF}:
   Proper ((≡) ==> (=) ==> (≡) ==> (=) ==> (≡)) (@ProtoInv Σ Prtcl PF).
-Proof. rewrite /ProtoInv /All_Inv /BE_Inv /sEx /sBE. 
+Proof. rewrite /ProtoInv /All_Inv /BE_Inv /sEx /sBE.
   intros ? ? H1 ? ? H2 ? ? H3 ? ? H4.
   subst. apply bi.sep_proper.
   - rewrite big_opS_proper; last first.
diff --git a/theories/gps/singlewriter.v b/theories/gps/singlewriter.v
index 2230533d..1708f84b 100644
--- a/theories/gps/singlewriter.v
+++ b/theories/gps/singlewriter.v
@@ -5,7 +5,7 @@ From igps Require Import bigop abs_view.
 From iris.bi Require Import fractional.
 
 Section Gname_StrongSW.
-  Context {Σ : gFunctors} {Prtcl : protocolT} 
+  Context {Σ : gFunctors} {Prtcl : protocolT}
           `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF}
           `{CInvG : cinvG Σ}
           {IP : gname → interpC Σ Prtcl}.
@@ -17,7 +17,7 @@ Section Gname_StrongSW.
   Notation F_past γ := (IP γ false).
 
   Definition WAEx l γ γ_x s : vProp Σ
-    := (∃ v V_0 ζ, (monPred_in V_0) 
+    := (∃ v V_0 ζ, (monPred_in V_0)
         ∧ ⎡Writer γ ζ ∗ exwr γ_x 1%Qp (s, (v, V_0))
                         ∗ Reader γ {[s, (v, V_0)]}
                         ∗ ⌜maxTime l V_0 ζ⌝⎤)%I.
@@ -76,7 +76,7 @@ Section Gname_StrongSW.
           (gpsSWnRaw γ l (encode (γ, (γ_x)))) True.
   Proof. intros. rewrite /Frame /=. iIntros "[? _]". iExists _. by iFrame "∗". Qed.
 
-  Lemma GPS_nFWP_Init_strong l (s: gname -> pr_state Prtcl) v 
+  Lemma GPS_nFWP_Init_strong l (s: gname -> pr_state Prtcl) v
     (Q : gname -> vPred) (E : coPset) (HEN : ↑physN ⊆ E):
   {{{ ⎡PSCtx⎤ ∗ own_loc l ∗
         (∀ γ,
@@ -170,8 +170,8 @@ Section Gname_StrongSW.
     - rewrite /PrtSeen. iExists V0, v. by iFrame "oR".
   Qed.
 
-  Lemma GPS_nFWP_Writer_Reader_join γ l s q1 q2:
-    vGPS_nFWP γ l s q1 ∗ vGPS_nFRP γ l s q2 ⊢ vGPS_nFWP γ l s (q1 + q2).
+  Lemma GPS_nFWP_Writer_Reader_join γ l s1 s2 q1 q2:
+    vGPS_nFWP γ l s1 q1 ∗ vGPS_nFRP γ l s2 q2 ⊢ vGPS_nFWP γ l s1 (q1 + q2).
   Proof.
     constructor => V.
     iIntros "(o1 & o2)". iCombine "o1" "o2" as "oW".
@@ -211,12 +211,10 @@ Section Gname_StrongSW.
     by iIntros (X) "[R1 R2]".
   Qed.
 
-  Lemma GPS_nFWP_join_writer γ l s1 s2 q1 q2 V:
-    vGPS_nFWP γ l s1 q1 V ∗ vGPS_nFRP γ l s2 q2 V ⊢ vGPS_nFWP γ l s1 (q1 + q2) V.
+  Lemma GPS_nFWP_join_writer γ l s1 s2 q1 q2:
+    vGPS_nFWP γ l s1 q1 ∗ vGPS_nFRP γ l s2 q2 ⊢ vGPS_nFWP γ l s1 (q1 + q2).
   Proof.
-    iIntros "[oW oR]". iCombine "oW" "oR" as "oW".
-    rewrite vGPS_nFWP_eq vGPS_nFRP_eq fractor_join_l. iApply (fractor_mono with "[$oW]").
-    by iIntros (X) "[R1 R2]".
+    exact: GPS_nFWP_Writer_Reader_join.
   Qed.
 
   Lemma GPS_nFWP_writer_exclusive γ1 γ2 l s1 s2 q1 q2 V:
@@ -238,7 +236,7 @@ Section Gname_StrongSW.
     iIntros "oW". iDestruct (GPS_nFWP_gname_agree with "oW") as %?.
     subst γ2. iApply (GPS_nFWP_join_writer with "oW").
   Qed.
-  
+
   Global Instance GPS_nFRP_fractional' γ l s V:
     Fractional (λ q, vGPS_nFRP γ l s q V).
   Proof.
@@ -284,7 +282,7 @@ Section Gname_StrongSW.
     unfold bi_sep; simpl; unfold sbi_sep; simpl.
     rewrite monPred_sep_eq; done.
   Qed.
-  
+
   Lemma GPS_nFWP_dealloc γ l s V (E: coPset) (HNl: ↑fracN .@ l ⊆ E):
     vGPS_nFWP γ l s 1%Qp V ={E}=∗ own_loc l V.
   Proof.
@@ -327,7 +325,7 @@ Section Gname_StrongSW.
   Qed.
 
   Lemma GPS_nFRP_mult_splitL γ l s (n: positive) q:
-    vGPS_nFRP γ l s (Pos2Qp n * q) 
+    vGPS_nFRP γ l s (Pos2Qp n * q)
       ⊢ [∗ list] _ ∈ seq 0 (Pos.to_nat n), vGPS_nFRP γ l s q.
   Proof.
     rewrite -{1}(Pos2Nat.id n).
@@ -420,7 +418,7 @@ Section Gname_StrongSW.
 
   Lemma GPS_nFRP_Read' l γ q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ → 
+  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past γ l s' v ∗ P ∗ vGPS_nFRP γ l s' q ={E}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F γ l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
          ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past γ l s' v
@@ -448,10 +446,8 @@ Section Gname_StrongSW.
             with "[VS $oI $kI $kS $oP $oR VSDup HClose]"); [solve_ndisj|auto|..].
     { iNext; iSplitR "VSDup"; last by iSpecialize ("VSDup" $! _).
       iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
-      { done. }
+      iDestruct ("VS" $! _ j with "[//] [Hs' //]") as "[VS|VS]"; last by iRight.
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (fun l _ p => ReaderSWnRaw γ s' j0 l p) with "[oI oR']").
       { iSplit.
         + iNext. iExists γ_x. by iFrame "oI".
@@ -461,12 +457,12 @@ Section Gname_StrongSW.
     iNext.
     iIntros (s' v Vr V') "(% & kS' & Hs' & oQ & oR' & % & %)".
     iFrame; iApply ("Post" $! _ _ _ with "[%]"); first done.
-    rewrite monPred_fupd_eq. iFrame "Hs' oQ".
+    by iFrame "Hs'".
   Qed.
 
   Lemma GPS_nFRP_Read l γ q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ → 
+  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past γ l s' v ∗ P ={E ∖ ↑fracN .@ l}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F γ l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
          ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past γ l s' v
@@ -489,7 +485,7 @@ Section Gname_StrongSW.
   Lemma GPS_nFRP_Read_PrtSeen_abs' `{aG: absViewG Σ}
     l γ q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s1 s2 β (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s1 ⊑ s' ∧ s2 ⊑ s'⌝ → 
+  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s1 ⊑ s' ∧ s2 ⊑ s'⌝ →
               (∀ v, F_past γ l s' v ∗ P ∗ ⌞vGPS_nFRP γ l s1 q⌟ β ={E}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F γ l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
          ∗ ▷ ɐ (∀ s' v, ⌜s1 ⊑ s' ∧ s2 ⊑ s'⌝ ∗ F_past γ l s' v
@@ -528,7 +524,6 @@ Section Gname_StrongSW.
           { etrans; eauto. }
           { split; [by etrans|auto]. }
           iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-          rewrite monPred_fupd_eq.
           iMod ("HClose" $! (fun l _ => ReaderSWnRaw γ s1 Vβ l) with "[oI oR']").
           { iSplitL "oI".
             + iNext. iExists γ_x. iFrame; auto.
@@ -541,10 +536,9 @@ Section Gname_StrongSW.
           iApply ("VSDup" $! _ s' v _ with "[%] [$Fp]"); first reflexivity.
           iPureIntro. split; [by etrans|auto]. }
       iNext. iIntros (s' v Vr V') "(% & kS & % & Q & oR' & % & %)".
-      iFrame; iApply ("Post" $! _ _ V' with "[%]").
+      iFrame; iApply ("Post" $! s' _ V' with "[%]").
       { etrans; eauto. }
-      rewrite monPred_fupd_eq.
-      iMod "Q"; iFrame.
+      rewrite monPred_fupd_eq. iMod "Q" as "$".
       iPureIntro; split; [by etrans|auto].
     - iPoseProof "aS" as "aS1".
       iDestruct "aS1" as (Vβ') "[% aS1]".
@@ -558,7 +552,6 @@ Section Gname_StrongSW.
           { etrans; eauto. }
           { split; [auto|by etrans]. }
           iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-          rewrite monPred_fupd_eq.
           iMod ("HClose" $! (fun l _ => ReaderSWnRaw γ s1 Vβ l) with "[oI oR']").
           { iSplitL "oI".
             + iNext. iExists γ_x. iFrame; auto.
@@ -573,15 +566,14 @@ Section Gname_StrongSW.
       iNext. iIntros (s' v Vr V') "(% & kS & % & Q & oR' & % & %)".
       iFrame; iApply ("Post" $! _ _ V' with "[%]").
       { etrans; eauto. }
-      rewrite monPred_fupd_eq.
-      iMod "Q"; iFrame.
+      rewrite monPred_fupd_eq. iMod "Q"; iFrame.
       iPureIntro; split; [auto|by etrans].
   Qed.
 
   Lemma GPS_nFRP_Read_PrtSeen_abs `{aG: absViewG Σ}
     l γ q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s1 s2 β (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-  ⎡PSCtx⎤ ⊢ {{{ (▷ ∀ s', ⌜s1 ⊑ s' ∧ s2 ⊑ s'⌝ → 
+  ⎡PSCtx⎤ ⊢ {{{ (▷ ∀ s', ⌜s1 ⊑ s' ∧ s2 ⊑ s'⌝ →
               (∀ v, F_past γ l s' v ∗ P ={E ∖ ↑fracN .@ l}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F γ l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
          ∗ ▷ ɐ (∀ s' v, ⌜s1 ⊑ s' ∧ s2 ⊑ s'⌝ ∗ F_past γ l s' v
@@ -641,7 +633,6 @@ Section Gname_StrongSW.
       { etrans; eauto. }
       { done. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (fun l _ => ReaderSWnRaw γ s Vβ l) with "[oI oR']").
       { iSplitL "oI".
         + iNext. iExists γ_x. iFrame; auto.
@@ -682,7 +673,7 @@ Section Gname_StrongSW.
 End Gname_StrongSW.
 
 Section Gname_StrongSW_2.
-  Context {Σ : gFunctors} {Prtcl : protocolT} 
+  Context {Σ : gFunctors} {Prtcl : protocolT}
           `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF}
           `{CInvG : cinvG Σ}
           {IP1 IP2 : gname → interpC Σ Prtcl}.
@@ -707,7 +698,7 @@ Section Gname_StrongSW_2.
 End Gname_StrongSW_2.
 
 Section Gname_StrongSW_Param.
-  Context {Σ : gFunctors} {Prtcl : protocolT} 
+  Context {Σ : gFunctors} {Prtcl : protocolT}
           `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF}
           `{persG : persistorG Σ}
           `{agreeG : gps_agreeG Σ Prtcl}
@@ -749,7 +740,6 @@ Section Gname_StrongSW_Param.
   Definition GPS_nWSP γ l p s V : iProp :=
     persisted l (gpsSWpnRaw p γ) (WriterSWpnRaw p γ s V).
 
-Locate monPred.
   Definition vPred_GPS_nRSP_mono γ l p s:
     Proper ((⊑) ==> (⊢)) (GPS_nRSP γ l p s).
   Proof.
@@ -762,8 +752,13 @@ Locate monPred.
   Canonical Structure vGPS_nRSP γ l p s :=
     MonPred _ (vPred_GPS_nRSP_mono γ l p s).
 
-  Global Instance vGPS_nRSP_persistent γ l p s V
-    : Persistent (vGPS_nRSP γ l p s V) := _.
+  Global Instance vGPS_nRSP_persistent γ l p s
+    : Persistent (vGPS_nRSP γ l p s).
+  Proof.
+    rewrite /Persistent /bi_persistently. split => ? /=.
+    rewrite /sbi_persistently /= monPred_persistently_eq /=.
+    apply persistor_persistent => ?. apply _.
+  Qed.
 
   Definition vPred_GPS_nWSP_mono γ l p s:
     Proper ((⊑) ==> (⊢)) (GPS_nWSP γ l p s).
@@ -918,11 +913,10 @@ Locate monPred.
               with "[VS $oI $kI $kS $oP $oR VSDup HClose]"); [solve_ndisj|auto|..].
     { etrans; eauto. }
     { iNext; iSplitR "VSDup"; last by iSpecialize ("VSDup" $! _).
-      iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
+      iIntros (s' j ? Hs').
+      iDestruct ("VS" $! _ Hs' j with "[%]") as "[VS|VS]"; last by iRight.
       { etrans; eauto. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (ReaderSWpnRaw p γ s' j0) with "[oI oR']").
       { iSplit.
         + iNext. iExists γ_x. by iFrame "oI".
@@ -1008,7 +1002,7 @@ End Gname_StrongSW_Param.
 
 
 Section Gname_StrongSW_Param_Default.
-  Context {Σ : gFunctors} {Prtcl : protocolT} 
+  Context {Σ : gFunctors} {Prtcl : protocolT}
           `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF}
           `{persG : persistorG Σ}
           `{agreeG : gps_agreeG Σ Prtcl}
@@ -1033,7 +1027,7 @@ Section Gname_StrongSW_Param_Default.
 End Gname_StrongSW_Param_Default.
 
 Lemma GPS_nSWs_param_agree
-  {Σ : gFunctors} {Prtcl : protocolT} 
+  {Σ : gFunctors} {Prtcl : protocolT}
   `{fG : !foundationG Σ} `{GPSG : !gpsG Σ Prtcl PF}
   `{persG : persistorG Σ}
   `{agreeG : gps_agreeG Σ Prtcl}
@@ -1076,7 +1070,7 @@ Section SingleWriter.
 
   Local Open Scope I.
   Definition gpsSWRaw l n : iProp :=
-    ∃ (γ γ_x: gname), 
+    ∃ (γ γ_x: gname),
       ⌜n = encode (γ,γ_x)⌝ ∗ gps_inv IP l γ γ_x.
 
   Definition ReaderSWRaw s V l n : iProp :=
@@ -1113,7 +1107,12 @@ Section SingleWriter.
   Canonical Structure vGPS_RSP l s :=
     MonPred _ (vPred_GPS_RSP_mono l s).
 
-  Global Instance vGPS_RSP_persistent l s V : Persistent (vGPS_RSP l s V) := _.
+  Global Instance vGPS_RSP_persistent l s : Persistent (vGPS_RSP l s).
+  Proof.
+    rewrite /Persistent /bi_persistently. split => ? /=.
+    rewrite /sbi_persistently /= monPred_persistently_eq /=.
+    apply persistor_persistent => ?. apply _.
+  Qed.
 
   Definition vPred_GPS_WSP_mono l s:
     Proper ((⊑) ==> (⊢)) (GPS_WSP l s).
@@ -1151,7 +1150,7 @@ Section SingleWriter.
   Canonical Structure vGPS_FWP l s q :=
     MonPred _ (vPred_GPS_FWP_mono l s q).
 
-  Local Instance : 
+  Local Instance :
     Frame false (gps_inv IP l γ γ_x) (gpsSWRaw l (encode (γ, (γ_x)))) True.
   Proof. intros. unfold Frame, bi_affinely_if, bi_persistently_if.
          iIntros "[H _]". iExists _, _. by iFrame "∗". Qed.
@@ -1288,11 +1287,11 @@ Section SingleWriter.
 
   Lemma GPS_SW_Read' l (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑persist_locN .@ l ⊆ E):
-    {{{ ⎡PSCtx⎤ ∗ ▷ (∀ s', ⌜s ⊑ s'⌝ → 
+    {{{ ⎡PSCtx⎤ ∗ ▷ (∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past l s' v ∗ P ∗ vGPS_RSP l s' ={E}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P
                                        ={E ∖ ↑persist_locN .@ l}=∗ False))
-           ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v 
+           ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
               ={E ∖ ↑persist_locN .@ l}=∗ F_past l s' v ∗ F_past l s' v )
            ∗ P ∗ vGPS_RSP l s }}}
     ([ #l ]_at) @ E
@@ -1315,11 +1314,10 @@ Section SingleWriter.
               with "[VS $oI $kI $kS $oP $oR VSDup HClose]"); [solve_ndisj|auto|..].
     { etrans; eauto. }
     { iNext; iSplitR "VSDup"; last by iSpecialize ("VSDup" $! _).
-      iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
+      iIntros (s' j ? Hs').
+      iDestruct ("VS" $! _ Hs' j with "[%]") as "[VS|VS]"; last by iRight.
       { etrans; eauto. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (ReaderSWRaw s' j0) with "[oI oR']").
       { iSplit.
         + iNext. iExists γ, γ_x. by iFrame "oI".
@@ -1330,16 +1328,16 @@ Section SingleWriter.
     iIntros (s' v Vr V') "(% & kS' & Hs' & oQ & oR' & % & %)".
     iFrame; iApply ("Post" $! _ _ _ with "[%]").
     { etrans; eauto. }
-    rewrite monPred_fupd_eq. iFrame "Hs' oQ".
+    by iFrame "Hs'".
   Qed.
 
   Lemma GPS_SW_Read l (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑persist_locN .@ l ⊆ E):
-    {{{ ⎡PSCtx⎤ ∗ ▷ (∀ s', ⌜s ⊑ s'⌝ → 
+    {{{ ⎡PSCtx⎤ ∗ ▷ (∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past l s' v ∗ P ={E ∖ ↑ persist_locN .@ l}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P
                                        ={E ∖ ↑persist_locN .@ l}=∗ False))
-           ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v 
+           ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
               ={E ∖ ↑persist_locN .@ l}=∗ F_past l s' v ∗ F_past l s' v )
            ∗ P ∗ vGPS_RSP l s }}}
     ([ #l ]_at) @ E
@@ -1385,11 +1383,11 @@ Section SingleWriter.
     setoid_rewrite H1; setoid_rewrite H2; setoid_rewrite H4.
     iFrame; iApply ("Post" $! V' with "[%]"); [done|].
     iMod ("VS" $! _ _ with "[%] [$oP $oF wTok ex]") as "(oQ & oF & oFp)"; [done|..].
-    - iApply "HT". 
+    - iApply "HT".
       iExists γ, γ_x. iSplitL ""; first auto.
       iExists v, V',_. iFrame "wTok ex oR'". by iSplit; auto.
     - iFrame "oQ".
-      iMod ("HClose" with "[-]"). 
+      iMod ("HClose" with "[-]").
       iExists γ, γ_x. iSplitL ""; first auto.
       iApply ("oI" with "[$oF $oFp]").
       iModIntro. iApply "HT".
@@ -1412,13 +1410,12 @@ Section SingleWriter.
   Qed.
 
   Lemma GPS_FRPs_agree l s1 s2 q1 q2 (E : coPset) (HE: ↑fracN .@ l ⊆ E):
-    (vGPS_FRP l s1 q1 ∗ vGPS_FRP l s2 q2 
+    (vGPS_FRP l s1 q1 ∗ vGPS_FRP l s2 q2
       ={E}=∗ ⌜s1 ⊑ s2 ∨ s2 ⊑ s1⌝ ∗ vGPS_FRP l s1 q1 ∗ vGPS_FRP l s2 q2)%stdpp.
   Proof.
-    constructor => V.
-    iIntros "[oR1 oR2]"; iCombine "oR1 oR2" as "RP".
+    iStartProof (uPred _).
+    iIntros (V) "[oR1 oR2]"; iCombine "oR1 oR2" as "RP".
     iDestruct (fractor_splitjoin l with "RP") as "H".
-    rewrite monPred_fupd_eq.
     iMod (fractor_open with "H") as (X) "(oI & (oR1 & oR2) & HClose)";
       first auto.
     iDestruct "oR1" as (γ1 γ_x1) "(% & #RA1)".
@@ -1477,7 +1474,7 @@ Section SingleWriter.
 
   Lemma GPS_FRP_Read' l q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ → 
+  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past l s' v ∗ P ∗ vGPS_FRP l s' q ={E}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
          ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
@@ -1504,11 +1501,10 @@ Section SingleWriter.
               with "[VS $oI $kI $kS $oP $oR VSDup HClose]"); [solve_ndisj|auto|..].
     { etrans; eauto. }
     { iNext; iSplitR "VSDup"; last by iSpecialize ("VSDup" $! _).
-      iIntros (s' j ?) "Hs'".
-      iDestruct ("VS" $! _ _ with "[%] Hs'") as "[VS|VS]"; last by iRight.
+      iIntros (s' j ? Hs').
+      iDestruct ("VS" $! _ Hs' j with "[%]") as "[VS|VS]"; last by iRight.
       { etrans; eauto. }
       iLeft; iIntros (v Vr j0 Hle') "(? & ? & % & % & oI & oR')".
-      rewrite monPred_fupd_eq.
       iMod ("HClose" $! (fun l _ => ReaderSWRaw s' j0 l) with "[oI oR']").
       { iSplit.
         + iNext. iExists γ, γ_x. by iFrame "oI".
@@ -1519,12 +1515,12 @@ Section SingleWriter.
     iIntros (s' v Vr V') "(% & kS' & Hs' & oQ & oR' & % & %)".
     iFrame; iApply ("Post" $! _ _ _ with "[%]").
     { etrans; eauto. }
-    rewrite monPred_fupd_eq. iFrame "Hs' oQ".
+    by iFrame "Hs'".
   Qed.
 
   Lemma GPS_FRP_Read l q (P : vPred) (Q : pr_state Prtcl → Z → vPred)
     s (E : coPset) (HN : ↑physN ⊆ E) (HNl: ↑fracN .@ l ⊆ E):
-  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ → 
+  {{{ ⎡PSCtx⎤ ∗ (▷ ∀ s', ⌜s ⊑ s'⌝ →
               (∀ v, F_past l s' v ∗ P ={E ∖ ↑fracN .@ l}=∗ Q s' v)
             ∨ (∀ s'' v, ⌜s' ⊑ s''⌝ → F l s'' v ∗ P ={E ∖ ↑fracN .@ l}=∗ False))
          ∗ ▷ ɐ (∀ s' v, ⌜s ⊑ s'⌝ ∗ F_past l s' v
@@ -1819,9 +1815,6 @@ Notation "'[nFXP' l 'in' s '@' γ | p ']_R_' q" :=
      format "[nFXP  l  in  s @ γ  |  p  ]_R_ q").
 
 
-(* TODO: fix instances to compile *)
-Check (fun l F s1 => [XP l in s1 | F ]_R).
-
 Global Instance vGPS_RSP_contractive `{fG : foundationG Σ} `{pG : persistorG Σ} `{GPSG : gpsG Σ Prtcl}:
   Contractive (λ F : interpC _ _, [XP l in s1 | F ]_R).
 Proof.
@@ -1926,7 +1919,7 @@ Global Instance vGPS_WSP_proper {Σ Prtcl fG PF gpsG persG}:
 Proof.
   intros ? ? H1 ? ? H2 ? ? H3. constructor => ?. subst.
   unfold vGPS_WSP. rewrite /GPS_WSP /=.
-  f_equiv. intros ? ?. by rewrite (gpsSWRaw_proper _ _ H1). 
+  f_equiv. intros ? ?. by rewrite (gpsSWRaw_proper _ _ H1).
 Qed.
 
 Global Instance vGPS_RSP_proper {Σ Prtcl fG PF gpsG persG}:
@@ -1934,5 +1927,5 @@ Global Instance vGPS_RSP_proper {Σ Prtcl fG PF gpsG persG}:
 Proof.
   intros ? ? H1 ? ? H2 ? ? H3. constructor => ?. subst.
   unfold vGPS_RSP. rewrite /GPS_RSP /=.
-  f_equiv. intros ? ?. by rewrite (gpsSWRaw_proper _ _ H1). 
+  f_equiv. intros ? ?. by rewrite (gpsSWRaw_proper _ _ H1).
 Qed.
diff --git a/theories/gps/write.v b/theories/gps/write.v
index 37e5eef7..68916c2a 100644
--- a/theories/gps/write.v
+++ b/theories/gps/write.v
@@ -45,7 +45,7 @@ Section Write.
     iDestruct (Reader_Writer_sub with "[$oW $oR]") as %?%elem_of_subseteq_singleton.
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
     iApply (f_write_at with "[$kI $Seen $Hist]").
     - iPureIntro. exists (VInj v_r), V_r. repeat split => //.
       rewrite -H3 elem_of_map_gset. exists (s_r, (v_r, V_r)); split => //=.
@@ -186,7 +186,7 @@ Section Write.
     iDestruct (Reader_Writer_sub with "[$oW $oR]") as %?%elem_of_subseteq_singleton.
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
     iApply (f_write_at with "[$kI $kS $Hist]").
     - iPureIntro. exists (VInj v_x), V_x. repeat split; [|auto|auto].
       rewrite -H2 elem_of_map_gset. by exists (s_x, (v_x, V_x)).
diff --git a/theories/history.v b/theories/history.v
index f261669c..6bb9ac80 100644
--- a/theories/history.v
+++ b/theories/history.v
@@ -19,7 +19,7 @@ Section History.
   Proof.
     exists (fresh (locs M)).
     assert (NIn := is_fresh (locs M)).
-    move => m In Eq. apply NIn. 
+    move => m In Eq. apply NIn.
     rewrite elem_of_map_gset. eauto.
   Qed.
 
@@ -32,7 +32,7 @@ Section History.
 
   Lemma hist_subseteq M x : ∀ m, m ∈ hist M x → m ∈ M.
   Proof. move => m. by rewrite /=; set_solver. Qed.
-  
+
   Lemma hist_upto_subseteq M x t : hist_upto M x t ⊆ hist M x.
   Proof. rewrite /=. set_solver. Qed.
 
@@ -79,7 +79,7 @@ Section History.
     {disj: pairwise_disj M}
     (NEq : ∀ m, m ∈ M' → (mloc m ≠ x ∨ t ⊑ mtime m)):
     hist_upto (M ∪ M') x t ≡ hist_upto M x t.
-  Proof. 
+  Proof.
     apply set_unfold_2; move => m; intuition.
     edestruct NEq => //.
     assert (mtime m ⊏ mtime m).
@@ -87,11 +87,11 @@ Section History.
     exfalso. exact: (irreflexive_fw _ _ H3).
   Qed.
 
-  Lemma hist_complete {M : MessageSet} {x} {m} 
+  Lemma hist_complete {M : MessageSet} {x} {m}
     (In : m ∈ M) (EqLoc : mloc m = x)
     : m ∈ hist M x.
   Proof. by apply elem_of_hist. Qed.
-  
+
   Definition hist_from M x t :=
     {[ m <- hist M x | t ⊑ mtime m ]}.
 
@@ -130,7 +130,7 @@ Section History.
                       (mloc m) (mtime m))
     (Inm: m ∈ msgs M)
     (Valm: mval m = VInj v)
-    (ValGE: ∀ m', m' ∈ msgs M → mloc m' = mloc m 
+    (ValGE: ∀ m', m' ∈ msgs M → mloc m' = mloc m
                   → mtime m ⊏ mtime m' → isval (mval m'))
     : initialized M (mloc m) (mtime m).
   Proof.
diff --git a/theories/infrastructure.v b/theories/infrastructure.v
index 670eb532..f4cbbe04 100644
--- a/theories/infrastructure.v
+++ b/theories/infrastructure.v
@@ -45,7 +45,7 @@ Section suffixes.
   Qed.
 
   Lemma suffix_share_total l1 l2 l3:
-    l2 `suffix_of` l1 → l3 `suffix_of` l1 
+    l2 `suffix_of` l1 → l3 `suffix_of` l1
     → l2 `suffix_of` l3 ∨ l3 `suffix_of` l2.
   Proof.
     revert l2 l3. induction l1.
@@ -180,12 +180,12 @@ Section Max.
 
   Definition of_MaxDep' `(md : @MaxDep R B F c (Some b))
     : { a : A | a ∈ c ∧ ∀ a', a' ∈ c → R a' a} := ltac:(MD_tac; eauto).
-  Definition of_MaxDep `(md : @MaxDep R B F c (Some b)) 
+  Definition of_MaxDep `(md : @MaxDep R B F c (Some b))
     := let (m, _) := of_MaxDep' md in m.
   Global Arguments of_MaxDep' [_ _ _ _ _] _ /.
   Global Arguments of_MaxDep [_ _ _ _ _] _ /.
 
-  Lemma MaxDep_equiv {c2} `(md : @MaxDep R B F c1 ob) (Equiv : c1 ≡ c2) 
+  Lemma MaxDep_equiv {c2} `(md : @MaxDep R B F c1 ob) (Equiv : c1 ≡ c2)
     : @MaxDep R B F c2 ob.
   Proof.
     destruct md. pose (Max_equiv m Equiv).
@@ -221,7 +221,7 @@ Section Filter_GSet.
     X ⊆ Y → filter P X ⊆ filter P Y.
   Proof. set_solver. Qed.
 
-  Global Instance filter_gset_Proper : 
+  Global Instance filter_gset_Proper :
     Proper (((≡) : relation (gset A)) ==> ((≡) : relation (gset A))) (filter P).
   Proof. solve_proper. Qed.
 
@@ -232,7 +232,7 @@ Lemma gset_filter_comm `{Countable A} {P P2 : A -> Prop}
   filter P (filter P2 M) ≡ filter P2 (filter P M).
 Proof. set_solver. Qed.
 
-Lemma gset_filter_impl `{Countable A} {P P2 : A -> Prop} 
+Lemma gset_filter_impl `{Countable A} {P P2 : A -> Prop}
   `{DecP : ∀ m, Decision (P m)} `{DecP2 : ∀ m, Decision (P2 m)} {M : gset A} :
   (∀ a, P a -> P2 a) -> filter P (filter P2 M) ≡ filter P M.
 Proof. set_solver. Qed.
@@ -277,7 +277,7 @@ Section Filter_Help.
   Definition lType := (sigT (λ (P : A -> Prop), ∀ a, Decision (P a))).
   Definition lPred P {DecP : ∀ a : A, Decision (P a)} : lType := existT P DecP.
   Bind Scope lType_scope with lType.
-  Definition lfold := ( 
+  Definition lfold := (
                     fold_right
                       (λ s : lType, let (P, DecP) := s in @filter _ (gset _) _ P DecP)
                   ).
@@ -425,8 +425,8 @@ Class CompatibleExtensions A `{SqSubsetEq A} `{StrictExtension A} : Prop := {
   compat_ext_r {x y z} : x ⊑ y → y ⊏ z → x ⊏ z;
 }.
 
-Global Instance TE_Proper `{@TightExtensions A E SE} `{Transitive A SE} 
-  : Proper (flip (⊏) ==> (⊏) ==> impl) (⊑). 
+Global Instance TE_Proper `{@TightExtensions A E SE} `{Transitive A SE}
+  : Proper (flip (⊏) ==> (⊏) ==> impl) (⊑).
 Proof.
   move => ? ? /= H1 ? ? H2 /=. rewrite !tight_extension => ?. intuition; subst.
   - right. by etrans.
@@ -434,91 +434,74 @@ Proof.
 Qed.
 
 (* TODO: why both? *)
-Global Instance CE_Proper `{C : @CompatibleExtensions A E SE} 
+Global Instance CE_Proper `{C : @CompatibleExtensions A E SE}
   : Proper (flip (⊑) ==> (⊑) ==> impl) (⊏).
 Proof.
   repeat intro. apply: compat_ext_l; last eauto.
   apply: compat_ext_r => //.
 Qed.
 
-Global Instance CE_Proper_flip `{C : @CompatibleExtensions A E SE} 
-  : Proper ((⊑) ==> flip (⊑) ==> flip impl) (⊏). 
+Global Instance CE_Proper_flip `{C : @CompatibleExtensions A E SE}
+  : Proper ((⊑) ==> flip (⊑) ==> flip impl) (⊏).
 Proof.
   repeat intro. apply: compat_ext_l; last eauto.
   apply: compat_ext_r => //.
 Qed.
 
-Global Instance PreOrder_Proper `{@PreOrder A R}
-  : Proper (flip R ==> R ==> impl) R.
-Proof. move => ? ? /= H1 ? ? H2 ?. etrans; first etrans; eauto. Qed.
-
 Class JoinSemiLattice A `{SqSubsetEq A, Join A} : Prop := {
-  join_semi_lattice_pre :>> PreOrder (⊑);
+  join_semi_lattice_pre :> PreOrder (⊑);
   join_ext_l X Y : X ⊑ X ⊔ Y;
   join_ext_r X Y : Y ⊑ X ⊔ Y;
   join_least X Y Z : X ⊑ Z → Y ⊑ Z → X ⊔ Y ⊑ Z
 }.
 
-Global Instance JoinSemiLattice_Transitive `{JoinSemiLattice A} : Transitive (⊑) := _.
-Global Instance JoinSemiLattice_Reflexive `{JoinSemiLattice A} : Reflexive (⊑) := _.
+Lemma join_ext_trans_l `{JoinSemiLattice A} (X Y Z : A) :
+  Z ⊑ X → Z ⊑ X ⊔ Y.
+Proof. intros. etrans; [done|]. apply join_ext_l. Qed.
+Lemma join_ext_trans_r `{JoinSemiLattice A} (X Y Z : A) :
+  Z ⊑ Y → Z ⊑ X ⊔ Y.
+Proof. intros. etrans; [done|]. apply join_ext_r. Qed.
 
 Create HintDb jsl.
-Ltac solve_jsl :=
-  now match goal with
-  | |- _ ⊑ _ => auto with jsl
-  | |- _ = _ ⊔ _ => auto with jsl
-  | |- _ ⊔ _ = _ => auto with jsl
-  end.
-Hint Extern 0 (?x ⊑ ?x) => reflexivity : jsl.
-Hint Extern 0 (?x ⊑ ?x ⊔ _) => apply join_ext_l : jsl.
-Hint Extern 0 (?x ⊑ _ ⊔ ?x) => apply join_ext_r : jsl.
-Hint Extern 2 (_ ⊔ _ ⊑ _) => apply join_least : jsl.
-Hint Resolve join_ext_l : jsl.
-Hint Resolve join_ext_r : jsl.
+Ltac solve_jsl := by typeclasses eauto with jsl typeclass_instances.
 Hint Resolve join_least : jsl.
-
-Lemma join_idempotent `{JSL : @JoinSemiLattice A E J} `{AS : AntiSymm A R E} (x : A) :
+Hint Extern 0 (?a ⊑ ?b) =>
+  (* We first check whether a and b are unifiable, in order not to
+     trigger typeclass search for Reflexivity when this is not needed. *)
+  unify a b with jsl; reflexivity : jsl.
+Hint Extern 0 (_ = _) => apply (anti_symm (⊑)) : jsl.
+Hint Resolve join_ext_trans_l | 10 : jsl.
+Hint Resolve join_ext_trans_r | 10 : jsl.
+Hint Extern 100 (?a ⊑ ?c) =>
+  match goal with H : a ⊑ ?b |- _ => transitivity b; [exact H|] end
+  : jsl.
+Hint Extern 200 (?a ⊑ ?c) =>
+  match goal with H : ?b ⊑ c |- _ => transitivity b; [|exact H] end
+  : jsl.
+
+Lemma join_idempotent `{JSL : JoinSemiLattice A} `{AS : AntiSymm A R (⊑)} (x : A) :
   R x (x ⊔ x).
-Proof.
-  eapply anti_symm; eauto.
-  - by eapply join_ext_l. 
-  - by eapply join_least.
-Qed.
+Proof. eapply anti_symm; solve_jsl. Qed.
 
-Hint Extern 1 (?x = ?x ⊔ ?x) => apply join_idempotent : jsl.
-(* Hint Resolve (@join_idempotent _ _ _ _ (=)). *)
-Hint Extern 2 (_ = _ ⊔ _) => apply (anti_symm (⊑)) : jsl.
-Hint Extern 2 (_ ⊔ _ = _) => symmetry; apply (anti_symm (⊑)) : jsl.
-
-Instance JSL_Comm `{E : SqSubsetEq A} `{J : Join A} `{@PartialOrder A (⊑)} `{@JoinSemiLattice A E J} : @Comm A _ (=) (⊔).
+Instance JSL_Comm `{JoinSemiLattice A} `{AntiSymm A (=) (⊑)} : Comm (=) (⊔).
 Proof. move => ? ?. solve_jsl. Qed.
-  
-Instance JSL_Assoc `{E : SqSubsetEq A} `{J : Join A} `{@PartialOrder A (⊑)} `{@JoinSemiLattice A E J} : @Assoc A (=) (⊔).
-Proof.
-  move => ? ? ?.
-  apply: anti_symm; repeat apply: join_least; eauto using PreOrder_Transitive with jsl.
-Qed.
-
-Hint Extern 0 (?x ⊑ ?x ⊔ _ ⊔ _) => rewrite <- (assoc (⊔)) : jsl.
-Hint Extern 0 (?x ⊔ ?y ⊔ ?z = ?x ⊔ (?y ⊔ ?z)) => rewrite <- (assoc (⊔)) : jsl.
-Hint Extern 0 (?x ⊔ (?y ⊔ ?z) = ?x ⊔ ?y ⊔ ?z) => rewrite <- (assoc (⊔)) : jsl.
 
+Instance JSL_Assoc `{JoinSemiLattice A} `{AntiSymm A (=) (⊑)} : Assoc (=) (⊔).
+Proof. move => ? ? ?. solve_jsl. Qed.
 
 Instance JSL_join_Proper `{JoinSemiLattice A} : Proper ((⊑) ==> (⊑) ==> (⊑)) (⊔).
-Proof. intros ? ? ? ? ? ?. apply join_least; etrans; eauto with jsl. Qed.
+Proof. intros ? ? ? ? ? ?. solve_jsl. Qed.
 
 (* Generic Instances for SqSubsetEq, Join and JoinSemiLattice *)
 Instance prod_SqSubsetEq `{SqSubsetEq A} `{SqSubsetEq B} : SqSubsetEq (A * B)
   := λ p1 p2, p1.1 ⊑ p2.1 ∧ p1.2 ⊑ p2.2.
-Instance prod_PreOrder 
+Instance prod_PreOrder
   `{SqSubsetEq A} `{SqSubsetEq B} `{PreOrder A (⊑)} `{PreOrder B (⊑)}
   : @PreOrder (A * B) (⊑).
 Proof.
   econstructor.
-  - move => ?; cbv; split; by apply PreOrder_Reflexive.
-  - move => ? ? ? [H11 H12] [H21 H22]; split.
-    + by apply (PreOrder_Transitive _ _ _ H11 H21).
-    + by apply (PreOrder_Transitive _ _ _ H12 H22).
+  - move => ?; by split.
+  - move => ? ? ? [H11 H12] [H21 H22]; split; by etrans.
 Qed.
 
 Instance prod_PartialOrder `{SqSubsetEq A} `{SqSubsetEq B} `{PartialOrder A (⊑)} `{PartialOrder B (⊑)} : @PartialOrder (A * B) (⊑).
@@ -527,13 +510,13 @@ Proof.
   - by apply prod_PreOrder.
   - move => [? ?] [? ?] [? ?] [? ?]. f_equal; by apply: anti_symm.
 Qed.
-  
+
 Instance prod_Join `{Join A} `{Join B} : Join (A * B) :=
   λ p1 p2, (p1.1 ⊔ p2.1, p1.2 ⊔ p2.2).
 Instance prod_JSL `{JoinSemiLattice A} `{JoinSemiLattice B} : JoinSemiLattice (A * B).
 Proof.
   econstructor.
-  - exact: prod_PreOrder.
+  - apply _.
   - move => ? ?; split => /=; solve_jsl.
   - move => ? ?; split => /=; solve_jsl.
   - move => ? ? ? [? ?] [? ?]. split => /=; solve_jsl.
@@ -592,13 +575,13 @@ Proof.
 Qed.
 
 Instance option_strict_subrel
-  `{StrictExtension T} `{SqSubsetEq T} 
-  `{subrelation _ ((⊏) : relation T) ((⊑) : relation T)} 
+  `{StrictExtension T} `{SqSubsetEq T}
+  `{subrelation _ ((⊏) : relation T) ((⊑) : relation T)}
   : subrelation ((⊏) : relation (option T)) ((⊑) : relation (option T)).
 Proof. move => [?|] [?|] //=. cbn. eauto with typeclass_instances. Qed.
 
 Instance option_StrictOrder `{StrictExtension T} `{StrictOrder T (⊏)}
-  : @StrictOrder (option T) (⊏). 
+  : @StrictOrder (option T) (⊏).
 Proof.
   econstructor.
   - move => ?. cbv; repeat case_match; last auto.
@@ -624,7 +607,7 @@ Proof.
     exact: PreOrder_Transitive.
 Qed.
 
-Instance option_PartialOrder `{SqSubsetEq T} `{PartialOrder T (⊑)} 
+Instance option_PartialOrder `{SqSubsetEq T} `{PartialOrder T (⊑)}
   : @PartialOrder (option T) (⊑).
 Proof.
   econstructor.
@@ -637,7 +620,7 @@ Instance option_Reflexive `{SqSubsetEq T} `{Reflexive T (⊑)}
 Proof. move => ?. cbv. by case_match. Qed.
 
 Instance option_AntiSymm `{SqSubsetEq T} `{AntiSymm T (=) (⊑)}
-  : AntiSymm (=) ((⊑) : relation (option T)).
+  : @AntiSymm (option T) (=) (⊑).
 Proof.
   move => ? ?; cbv; repeat case_match => //; move => H1 H2. apply/f_equal.
   by apply (@anti_symm _ _ _ H0).
@@ -649,11 +632,7 @@ Proof.
 Qed.
 
 Instance option_Total `{SqSubsetEq T} `{Total T (⊑)}: (@Total (option T) (⊑)).
-Proof.
-  move => [?|] [?|]; eauto with jsl.
-  - by right.
-  - by left.
-Qed.
+Proof. move => [?|] [?|]; eauto with jsl. by right. by left. Qed.
 
 (* Instance option_Assoc `{Join T} `{@Assoc T (=) join} : @Assoc (option T) (=) (join). *)
 (* Proof. *)
@@ -666,15 +645,10 @@ Qed.
 Instance option_JSL `{JoinSemiLattice T} : JoinSemiLattice (option T).
 Proof.
   econstructor.
-  - exact: option_PreOrder.
-  - move => ? ?. cbv; repeat case_match => //; simplify_option_eq.
-    + by apply join_ext_l.
-    + by apply PreOrder_Reflexive.
-  - move => ? ?. cbv; repeat case_match => //; simplify_option_eq.
-    + by apply join_ext_r.
-    + by apply PreOrder_Reflexive.
-  - move => ? ? ?; cbv; repeat case_match => //; simplify_option_eq.
-    by apply join_least.
+  - apply _.
+  - move => [?|][?|] //. apply (join_ext_l (A:=T)).
+  - move => [?|][?|] //. apply (join_ext_r (A:=T)).
+  - move => [?|][?|][?|] //. apply (join_least (A:=T)).
 Qed.
 
 Instance option_extension_eq_dec
@@ -717,7 +691,7 @@ Lemma join_None `{Join T} (o1 : option T) : o1 ⊔ None = o1.
 Proof. by cbv; repeat case_match. Qed.
 Lemma join_Some_not_None `{Join T} t1 o1 : None = Some t1 ⊔ o1 -> False.
 Proof. by cbv; repeat case_match. Qed.
-Lemma option_subseteq_None_l `{SqSubsetEq T} (o1 o2 : option T) 
+Lemma option_subseteq_None_l `{SqSubsetEq T} (o1 o2 : option T)
   : o1 ⊑ None ↔ o1 = None.
 Proof. by cbv; repeat case_match. Qed.
 
@@ -806,14 +780,14 @@ Qed.
 Local Instance gmap_SqSubsetEq `{Countable K} `{SqSubsetEq A} : SqSubsetEq (gmap K A) :=
   λ f1 f2, ∀ k, f1 !! k ⊑ f2 !! k.
 
-Instance gmap_PreOrder `{Countable K} `{SqSubsetEq A} `{PreOrder A (⊑)} 
+Instance gmap_PreOrder `{Countable K} `{SqSubsetEq A} `{PreOrder A (⊑)}
   : @PreOrder (gmap K A) (⊑).
 Proof.
   econstructor.
   - move => ? ?; reflexivity.
   - move => ? ? ? ? ? ?. by etransitivity; eauto.
 Qed.
-Instance gmap_PartialOrder `{Countable K} `{SqSubsetEq A} `{PartialOrder A (⊑)} 
+Instance gmap_PartialOrder `{Countable K} `{SqSubsetEq A} `{PartialOrder A (⊑)}
   : @PartialOrder (gmap K A) (⊑).
 Proof.
   econstructor.
@@ -844,7 +818,7 @@ Proof.
   - unfold set_Exists, set_Forall. abstract naive_solver.
 Qed.
 
-Lemma gset_neg_Exists `{Countable A} (P : A -> Prop) 
+Lemma gset_neg_Exists `{Countable A} (P : A -> Prop)
   `{∀ a, Decision (P a)} (s : gset A)
   : ¬ set_Exists P s ↔ set_Forall (λ x, ¬ P x) s.
 Proof.
@@ -937,7 +911,7 @@ Proof.
   repeat intro; unfold rel_of in *.
     by eauto with typeclass_instances.
 Qed.
-Instance rel_of_Irreflexive 
+Instance rel_of_Irreflexive
   : (@Irreflexive T r) → (@Irreflexive X (@rel_of X T f r)) | 10.
 Proof.
   repeat intro; unfold rel_of in *.
@@ -954,13 +928,13 @@ Proof.
 Qed.
 
 (* Miscellaneous *)
-Lemma elem_of_contains (T : Type) (l1 l2 : list T) (t : T) 
+Lemma elem_of_contains (T : Type) (l1 l2 : list T) (t : T)
   : l2 ⊆+ l1 → t ∈ l2 → t ∈ l1.
 Proof.
   induction 1 => //.
   - move/elem_of_cons => [->|?]; apply/elem_of_cons; tauto.
   - move/elem_of_cons => [->|/elem_of_cons [->|?]]; apply/elem_of_cons.
-    + right; apply/elem_of_cons; tauto. 
+    + right; apply/elem_of_cons; tauto.
     + by left.
     + right; apply/elem_of_cons; tauto.
   - move/IHsubmseteq => ?. apply/elem_of_cons. tauto.
@@ -980,7 +954,7 @@ Notation "'∀max_[' R ']' m ∈ M , P" :=
 
 
 Lemma gset_nonempty_min_strict
-  `{Countable A} `{Transitive A R} `{Total A R} 
+  `{Countable A} `{Transitive A R} `{Total A R}
   `{∀ a a', Decision (R a a')} (M : gset A) (NE: M ≢ ∅)
   : ∃ m, m ∈ M ∧ (∀ m', m' ∈ M ∖ {[m]} → R m m').
 Proof.
@@ -991,7 +965,7 @@ Proof.
 Qed.
 
 Lemma gset_nonempty_min
-  `{Countable A} `{Transitive A R} `{Reflexive A R} `{Total A R} 
+  `{Countable A} `{Transitive A R} `{Reflexive A R} `{Total A R}
   `{∀ a a', Decision (R a a')} (M : gset A) (NE: M ≢ ∅)
   : ∃min_[R] m ∈ M, True : Prop.
 Proof.
diff --git a/theories/lang/base.v b/theories/lang/base.v
index 500d274e..7aa413ff 100644
--- a/theories/lang/base.v
+++ b/theories/lang/base.v
@@ -160,6 +160,7 @@ Definition subst' (mx : binder) (es : expr) : expr → expr :=
 (** The stepping relation *)
 Definition un_op_eval (op : un_op) (l : val) : option val :=
   match op, l with
+  | NegOp, LitV (LitInt n) => Some $ LitV $ LitInt (if decide (n = 0) then 1 else 0)
   | MinusUnOp, LitV (LitInt n) => Some $ LitV $ LitInt (- n)
   | CastIntOp, LitV (LitLoc n) => Some $ LitV $ LitInt (Z.pos n)
   | CastLocOp, LitV (LitInt n) => Some $ LitV $ LitLoc $ Z.to_pos n
@@ -275,7 +276,7 @@ Proof. intros. apply is_closed_subst with []; set_solver. Qed.
 Instance base_lit0_dec_eq (l1 l2 : base_lit) : Decision (l1 = l2).
 Proof. solve_decision. Defined.
 Instance base_lit_dec_eq (l1 l2 : base_lit) : Decision (l1 = l2).
-Proof. 
+Proof.
   destruct l1, l2; try (by right); try (by left);
   eauto with typeclass_instances.
 Defined.
@@ -310,4 +311,4 @@ Arguments Fork _%E.
 Arguments Load _ _%E.
 Arguments Store _ _%E _%E.
 Arguments CAS _%E _%E _%E.
-Arguments FAI _ _%E.
\ No newline at end of file
+Arguments FAI _ _%E.
diff --git a/theories/lang/ra_lang.v b/theories/lang/ra_lang.v
index 1c38e432..fd9dc005 100644
--- a/theories/lang/ra_lang.v
+++ b/theories/lang/ra_lang.v
@@ -1,4 +1,3 @@
-
 From iris.program_logic Require Import ectx_language ectxi_language language.
 From stdpp Require Export strings gmap.
 From iris.algebra Require Import ofe.
@@ -60,7 +59,7 @@ Inductive step (V : View)
           (base.of_val (base.LitV $ base.LitInt v))
           []
 | StoreS l e v acc :
-    base.to_val e = Some (base.LitV $ base.LitInt v) -> 
+    base.to_val e = Some (base.LitV $ base.LitInt v) ->
     step V
           (base.Store acc (base.Lit $ base.LitLoc l) e)
           (EWrite acc l (VInj v))
@@ -95,8 +94,8 @@ Inductive head_step :
 ∀ (e : expr) (σ : state) (e' : expr) (σ' : state) (ef : list expr), Prop :=
 | pure_step σ π e e' (BaseStep : base.head_step e e')
   : head_step (e, π) σ (e', π) σ []
-| impure_step σ σ' V V' e e' ef evt 
-              (ExprStep : step V e evt e' ef) 
+| impure_step σ σ' V V' e e' ef evt
+              (ExprStep : step V e evt e' ef)
               (MachStep : machine_red σ evt V σ' V')
   : head_step (e, V) σ (e', V') σ' ef
 .
@@ -107,7 +106,7 @@ Proof. destruct v. cbv -[base.to_val base.of_val]. by rewrite base.to_of_val. Qe
 Lemma of_to_val e v : to_val e = Some v → of_val v = e.
 Proof.
   destruct e, v. cbv -[base.to_val base.of_val]. case C : (base.to_val e) => //.
-  move => [<- <-]. f_equal. exact: base.of_to_val. 
+  move => [<- <-]. f_equal. exact: base.of_to_val.
 Qed.
 
 Lemma val_stuck σ1 e1 σ2 e2 ef : head_step e1 σ1 e2 σ2 ef → to_val e1 = None.
@@ -131,7 +130,7 @@ Proof.
   exact: base.fill_item_no_val_inj H1 H2 H3.
 Qed.
 
-Lemma head_ctx_step_val Ki e σ e2 σ2 ef : 
+Lemma head_ctx_step_val Ki e σ e2 σ2 ef :
   head_step (fill_item Ki e) σ e2 σ2 ef → is_Some (to_val e).
 Proof.
   inversion 1; subst; last inversion ExprStep; subst; simplify_option_eq;
@@ -202,4 +201,4 @@ Canonical Structure ra_ectxi_lang :=
   EctxiLanguage (expr:=prod _ _) ra_ectxi_language_mixin.
 
 Canonical Structure ra_ectx_lang := EctxLanguageOfEctxi ra_ectxi_lang.
-Canonical Structure ra_lang := LanguageOfEctx ra_ectx_lang.
\ No newline at end of file
+Canonical Structure ra_lang := LanguageOfEctx ra_ectx_lang.
diff --git a/theories/lang/surface.v b/theories/lang/surface.v
index 3855f171..192e7af8 100644
--- a/theories/lang/surface.v
+++ b/theories/lang/surface.v
@@ -1,4 +1,3 @@
-
 From iris.program_logic Require Import ectx_language ectxi_language language.
 From stdpp Require Export strings gmap.
 From iris.algebra Require Import ofe.
@@ -45,4 +44,4 @@ Canonical Structure surface_ectxi_lang :=
   EctxiLanguage surface_ectxi_language_mixin.
 
 Canonical Structure surface_ectx_lang := EctxLanguageOfEctxi surface_ectxi_lang.
-Canonical Structure surface_lang := LanguageOfEctx surface_ectx_lang.
\ No newline at end of file
+Canonical Structure surface_lang := LanguageOfEctx surface_ectx_lang.
diff --git a/theories/lifting.v b/theories/lifting.v
index b7cb2a9d..a1a8d69a 100644
--- a/theories/lifting.v
+++ b/theories/lifting.v
@@ -26,12 +26,12 @@ Lemma wp_lift_atomic_step2 {E Φ} e1 σ1 :
   reducible e1 σ1 →
   ▷ ownP σ1 ∗
   ▷ (∀ v2 σ2 efs,
-    ⌜prim_step e1 σ1 (of_val v2) σ2 efs⌝ ∧ ownP σ2 -∗ 
+    ⌜prim_step e1 σ1 (of_val v2) σ2 efs⌝ ∧ ownP σ2 -∗
       |==> (|={E}=> Φ v2) ∗ [∗ list] ef ∈ efs, WP ef {{ _, True }})
   ⊢ WP e1 @ E {{ Φ }}.
 Proof.
   iIntros (Hatomic ?) "[Hσ H]".
-  iApply (ownP_lift_step E _ e1).
+  iApply (ownP_lift_step _ E _ e1).
   iMod (fupd_intro_mask' E ∅) as "Hclose"; first set_solver+. iModIntro.
   iExists σ1. iFrame "Hσ"; iSplit; eauto.
   iNext; iIntros (e2 σ2 efs) "% oP".
@@ -44,7 +44,7 @@ Qed.
 Lemma wp_lift_atomic_head_step2 {E Φ} e1 σ1 :
   atomic' e1 →
   head_reducible e1 σ1 →
-  ▷ ownP σ1 ∗ 
+  ▷ ownP σ1 ∗
   ▷ (∀ v2 σ2 efs,
     ⌜ectx_language.head_step e1 σ1 (of_val v2) σ2 efs⌝ ∧ ownP σ2 -∗
       |==> (|={E}=> Φ v2) ∗ [∗ list] ef ∈ efs, WP ef {{ _, True }})
@@ -287,7 +287,7 @@ Proof.
     + apply/Pos.lt_nle : (Pos.lt_trans _ _ _ Alloc Local) => F. apply: F.
       rewrite -EqTime. apply: MT_Some_Lt => //.
       by apply elem_of_filter.
-    + apply (not_elem_of_empty (C:=gset message) m). 
+    + apply (not_elem_of_empty (C:=gset message) m).
       by rewrite -(MT_empty TD) !elem_of_filter.
 Qed.
 
@@ -328,8 +328,8 @@ Lemma wp_load_at_pst E V σ l :
   drf_pre σ V l →
   init_pre σ V l  →
     {{{ ▷ ownP σ }}}
-      (Load at_hack (Lit $ (LitLoc l)), V) @ E 
-    {{{ v σ' V', RET (LitV (LitInt v), V'); 
+      (Load at_hack (Lit $ (LitLoc l)), V) @ E
+    {{{ v σ' V', RET (LitV (LitInt v), V');
         ⌜phys_inv σ'⌝ ∧ ⌜rel_inv σ σ'⌝ ∧ ⌜view_ok (mem σ') V'⌝ ∧ ⌜read_at_post σ σ' V V' l v⌝ ∧ ownP σ' }}}.
 Proof.
   intros Inv VOk A B Φ.
@@ -437,7 +437,7 @@ Lemma wp_store_at_pst E V σ l v:
   alloc_pre σ V l  →
    {{{ ▷ ownP σ }}}
     (Store at_hack (Lit $ (LitLoc l)) (Lit $ LitInt v), V) @ E
-   {{{ σ' V', RET (LitV (LitUnit), V');  
+   {{{ σ' V', RET (LitV (LitUnit), V');
         ⌜phys_inv σ'⌝ ∧ ⌜rel_inv σ σ'⌝ ∧ ⌜view_ok (mem σ') V'⌝ ∧ ⌜write_at_post σ σ' V V' l v⌝ ∧ ownP σ' }}}.
 Proof.
   intros Inv VOk A B Φ.
@@ -463,9 +463,7 @@ Proof.
         by apply tMax, elem_of_hist. }
 
     eexists.
-    exists (mkState 
-                    (add_ins (mem σ) m Disj)
-                    (nats σ)).
+    exists (mkState (add_ins (mem σ) m Disj) (nats σ)).
     econstructor. eright; first by constructor.
     econstructor; try eauto => //.
     + move => ? [<-]. by econstructor.
@@ -686,7 +684,7 @@ Lemma wp_CAS_pst E V σ l v_r v_w:
       (CAS (Lit $ LitLoc l) (Lit $ LitInt v_r) (Lit $ LitInt v_w), V) @ E
     {{{ b σ' V', RET (LitV $ lit_of_bool b, V');
            ⌜phys_inv σ'⌝ ∧ ⌜rel_inv σ σ'⌝ ∧ ⌜view_ok (mem σ') V'⌝
-         ∧ ⌜  b = true  ∧ CAS_succ_post σ σ' V V' l v_r v_w 
+         ∧ ⌜  b = true  ∧ CAS_succ_post σ σ' V V' l v_r v_w
               ∨ b = false ∧ CAS_fail_post σ σ' V V' l v_r  ⌝
          ∧ ownP σ' }}}.
 Proof.
@@ -728,7 +726,7 @@ Proof.
         rewrite -{2}Eq. eapply Pos.le_lt_trans; last by apply Pos.lt_add_r.
         by apply tMax, elem_of_hist. }
       eexists.
-      exists (mkState (add_ins (mem σ) m Disj) 
+      exists (mkState (add_ins (mem σ) m Disj)
                  (nats σ)).
       econstructor. eright; first by constructor.
       econstructor; try eauto => //=.
@@ -841,8 +839,7 @@ Proof.
         - rewrite TEq. by apply tMax, elem_of_hist.
         - apply Pos.lt_add_r. }
     eexists.
-    exists (mkState (add_ins (mem σ) m Disj) 
-               (nats σ)).
+    exists (mkState (add_ins (mem σ) m Disj) (nats σ)).
     econstructor. eright; first by constructor.
     econstructor; try eauto => //=.
     * move => ? [<-]. by econstructor.
@@ -908,8 +905,7 @@ Proof.
         last by left. exfalso. apply (Fresh _ HIn). rewrite /m in LEq.
         by simpl in LEq. }
     eexists.
-    exists (mkState (add_ins (mem σ) m Disj) 
-               (<[l := mtime m]> (nats σ))).
+    exists (mkState (add_ins (mem σ) m Disj) (<[l := mtime m]> (nats σ))).
     econstructor. eright; first by constructor.
     econstructor; try eauto => //.
     + move => x. instantiate (1 := l). move => /= EqLoc.
@@ -941,7 +937,7 @@ Proof.
       econstructor; subst; eauto; simpl; eauto.
       case Vl: (V !! _) => //.
       destruct (VOk _ _ Vl) as [m [In [EqLoc _]]].
-      exfalso. apply (not_elem_of_empty (C:=gset message) m). 
+      exfalso. apply (not_elem_of_empty (C:=gset message) m).
       rewrite -Fresh. move: In EqLoc; set_solver+.
 Qed.
 
diff --git a/theories/lifting_vProp.v b/theories/lifting_vProp.v
index 7af42f4e..107bac4a 100644
--- a/theories/lifting_vProp.v
+++ b/theories/lifting_vProp.v
@@ -22,7 +22,7 @@ Proof.
       destruct e1' as [e1' V1'], e2'0 as [e2'0 V2'0]. cbn in *.
       exists K (e1') (e2'0).
       * rewrite ->fill_view in *. by simplify_eq.
-      * rewrite ->fill_view in *. by simplify_eq. 
+      * rewrite ->fill_view in *. by simplify_eq.
       * rewrite ->fill_view in *. simplify_eq.
         inversion H1.
         { simplify_eq. unfold surface.head_step. cbn. exact H1. }
@@ -116,4 +116,4 @@ Proof. solve_pure_exec. Qed.
 
 Global Instance pure_if_false e1 e2:
   PureExec True (base.If (base.Lit (base.LitInt 0)) e1 e2) (e2).
-Proof. intros. solve_pure_exec. Qed.
\ No newline at end of file
+Proof. intros. solve_pure_exec. Qed.
diff --git a/theories/logically_atomic.v b/theories/logically_atomic.v
index d47614b3..ff769642 100644
--- a/theories/logically_atomic.v
+++ b/theories/logically_atomic.v
@@ -15,7 +15,132 @@ Section Atomic.
   Definition atomic_shift (Φ : val → vProp Σ) : vProp Σ :=
     (∃ P, ▷ P ∗ □ (▷ P ={Eo,Ei}=∗ ∃ x : A,
         α x ∗ ((α x ={Ei,Eo}=∗ ▷ P) ∧ ∀ y, β x y ={Ei,Eo}=∗ Φ y)))%I.
-  
+
   Definition atomic_wp e : vProp Σ :=
       (∀ Φ, atomic_shift Φ -∗ WP e @ Eo {{ Φ }})%I.
-End Atomic.
\ No newline at end of file
+
+  Lemma LA_WP e {Ht : ∀x, Timeless (α x)} (Hlt : Ei ⊆ Eo) (Hwp : atomic_wp e) x:
+    α x -∗ WP e {{v, β x v}}%I.
+  Proof.
+    iIntros "Ha".
+    iApply wp_mask_mono; first auto.
+    iApply Hwp; unfold atomic_shift.
+    iExists (α x); iFrame.
+    iIntros "!# Ha"; iMod "Ha".
+    iPoseProof (fupd_intro_mask' Eo Ei) as "Hupd"; first done.
+    iMod "Hupd"; iModIntro.
+    iExists x; iFrame; iSplit.
+    - iIntros "Ha".
+      iMod "Hupd"; auto.
+    - iIntros (v) "Hv".
+      iMod "Hupd"; auto.
+  Qed.
+
+  Lemma LA_Hoare e {Ht : ∀x, Timeless (α x)} (Hlt : Ei ⊆ Eo):
+    to_val e = None → atomic_wp e → ∀x, {{{α x}}} e {{{v, RET v; β x v}}}.
+  Proof.
+    intros.
+    iIntros "Ha Hb".
+    iPoseProof (LA_WP with "Ha") as "Hwp"; [done | done |].
+    iPoseProof (wp_frame_step_l' with "[$Hb $Hwp]") as "Hwp"; first done.
+    iApply wp_mono; last done.
+    intros; iIntros "[Hpost Hb]".
+    iApply "Hpost"; done.
+  Qed.
+
+End Atomic.
+
+From igps Require Import invariants.
+
+Lemma LA_Frame {Σ} {fG : foundationG Σ} {A} (α : A → vProp Σ) β Ei Eo e P E':
+  atomic_wp α β Ei Eo e ->
+  atomic_wp (fun x => α x ∗ P)%I (fun x v => β x v ∗ P)%I (Ei ∪ E') Eo e.
+Proof.
+  unfold atomic_wp, atomic_shift; intros Ha.
+  iIntros (Φ) "H"; iDestruct "H" as (P') "[HP' #Hshift]".
+  iApply Ha; iExists P'; iFrame.
+  iIntros "!# HP".
+  iMod ("Hshift" with "HP") as (x) "[[Hx HP] Hclose]".
+  iPoseProof (fupd_intro_mask' (Ei ∪ E') Ei) as "Hupd".
+  { apply union_subseteq_l. }
+  iMod "Hupd"; iModIntro.
+  iExists x; iFrame; iSplit.
+  - iIntros "Hx".
+    iMod "Hupd".
+    iApply "Hclose"; iFrame.
+  - iIntros (v) "Hv".
+    iMod "Hupd".
+    iApply "Hclose"; iFrame.
+Qed.
+
+Lemma fupd_ext {Σ} {fG : foundationG Σ} (E1 E1' E2 E2' : coPset) (P P' : vProp Σ)
+  (Heq1 : E1 ≡ E1') (Heq2 : E2 ≡ E2') (HP : P -∗ P'):
+  ((|={E1,E2}=>P) -∗ |={E1',E2'}=>P')%I.
+Proof.
+  iIntros "H".
+  iPoseProof (fupd_intro_mask' E1' E1) as "Hupd1"; first by rewrite Heq1.
+  iPoseProof (fupd_intro_mask' E2 E2') as "Hupd2"; first by rewrite Heq2.
+  iMod "Hupd1"; iMod "H"; iMod "Hupd2".
+  iModIntro; iApply HP; done.
+Qed.
+
+Instance fupd_proper {Σ} {fG : foundationG Σ}:
+  Proper ((≡) ==> (≡) ==> bi_wand_iff ==> bi_wand_iff)
+         (fupd : _ -> _ -> vProp Σ -> vProp Σ).
+Proof.
+  repeat intro.
+  iSplit; iIntros "H"; iApply (fupd_ext with "H"); auto;
+    iIntros "H"; iApply H1; done.
+Qed.
+
+Lemma LA_Inv {Σ} {fG : foundationG Σ} {A} (α : A → vProp Σ) β Ei Eo e N R {_ : Objective R}:
+  ↑N ## Ei ->
+  atomic_wp (fun x => ▷R ∗ α x)%I (fun x v => ▷R ∗ β x v)%I Ei Eo e ->
+  atomic_wp (fun x => inv N R ∗ α x)%I β (↑N ∪ Ei) Eo e.
+Proof.
+  unfold atomic_wp, atomic_shift; intros ? Ha.
+  iIntros (Φ) "H"; iDestruct "H" as (P) "[HP #Hshift]".
+  iApply Ha; iExists P; iFrame.
+  iIntros "!# HP".
+  iMod ("Hshift" with "HP") as (x) "[[#Hinv Hx] Hclose]".
+  iPoseProof (inv_open (↑N ∪ Ei) with "Hinv") as "HR".
+  { apply union_subseteq_l. }
+  iMod "HR" as "[HR HcloseR]".
+  iPoseProof (all_now with "HR") as "HR".
+  assert (Ei ≡ (↑N ∪ Ei) ∖ ↑N).
+  { rewrite difference_union_distr_l difference_diag union_empty_l difference_disjoint; done. }
+  iApply fupd_ext; [done | reflexivity | done |].
+  iModIntro.
+  iExists x; iFrame; iSplit.
+  - iIntros "[HR Hx]".
+    iApply fupd_ext; [done | done | done |].
+    iMod ("HcloseR" with "[HR]").
+    { by iApply objective. }
+    iApply "Hclose"; auto.
+  - iIntros (v) "[HR Hv]".
+    iApply fupd_ext; [done | done | done |].
+    iMod ("HcloseR" with "[HR]").
+    { by iApply objective. }
+    iApply "Hclose"; auto.
+Qed.
+
+Lemma LA_Csq {Σ} {fG : foundationG Σ} {A} (α : A → vProp Σ) β Ei Eo e α' β' E'
+  (Hpre1: ∀x, α x ={Ei ∪ E', Ei}=∗ α' x) (Hpre2: ∀x, α' x ={Ei, Ei ∪ E'}=∗ α x)
+  (Ha : atomic_wp α' β' Ei Eo e)
+  (Hpost : ∀x v, β' x v ={Ei, Ei ∪ E'}=∗ β x v):
+  atomic_wp α β (Ei ∪ E') Eo e.
+Proof.
+  unfold atomic_wp, atomic_shift.
+  iIntros (Φ) "H"; iDestruct "H" as (P) "[HP #Hshift]".
+  iApply Ha; iExists P; iFrame.
+  iIntros "!# HP".
+  iMod ("Hshift" with "HP") as (x) "[Hx Hclose]".
+  iMod (Hpre1 with "Hx"); iModIntro.
+  iExists x; iFrame; iSplit.
+  - iIntros "Hx".
+    iMod (Hpre2 with "Hx").
+    iApply "Hclose"; done.
+  - iIntros (v) "Hv".
+    iMod (Hpost with "Hv").
+    iApply "Hclose"; done.
+Qed.
diff --git a/theories/machine.v b/theories/machine.v
index c272f5bf..7569321f 100644
--- a/theories/machine.v
+++ b/theories/machine.v
@@ -43,7 +43,7 @@ Section RAMachine.
     Lemma pairwise_disj_MS (M: MessageSet) m:
       pairwise_disj M -> MS_msg_disj M m -> pairwise_disj (M ∪ {[m]}).
     Proof.
-      move => PDisj MDisj m1 m2 
+      move => PDisj MDisj m1 m2
               /elem_of_union [In1|/elem_of_singleton Eq1]
               /elem_of_union [In2|/elem_of_singleton Eq2].
       - destruct (PDisj m1 m2 In1 In2); abstract naive_solver.
@@ -52,7 +52,6 @@ Section RAMachine.
       - subst. by left.
     Qed.
 
-
     Definition msg_ok (m : message) := m.(mview) !! m.(mloc) = Some m.(mtime).
     Global Arguments msg_ok _ /.
 
@@ -67,7 +66,7 @@ Section RAMachine.
 
     (* MaxTime represents a message in the memory M and a corresponding proof
       that it is maximal for its location.
-      - MaxTime M x (Some t) holds if such a maximal message exists 
+      - MaxTime M x (Some t) holds if such a maximal message exists
         for location x and its timestamp is t.
       - MaxTime M x None holds if no message exists for x.
     *)
@@ -88,7 +87,7 @@ Section RAMachine.
     Global Arguments MT_None [_ _] _.
     Global Arguments MT_Some [_ _ _] _ _ _ _.
 
-    Lemma MT_equiv_loc {M1 M2 x ot} : 
+    Lemma MT_equiv_loc {M1 M2 x ot} :
       MaxTime M1 x ot →
       of_loc x M1 ≡ of_loc x M2 →
       MaxTime M2 x ot.
@@ -166,9 +165,9 @@ Section RAMachine.
       - exfalso. apply: (MT_msg_In' Max1). exact: (MT_msg_In_of_loc).
     Qed.
 
-    Lemma MT_pred_mono 
-      `{∀ m, Decision (P1 m)} `{∀ m, Decision (P2 m)} {M x ot1 ot2} 
-      : MaxTime (filter P1 M) x ot1 -> MaxTime (filter P2 M) x ot2 
+    Lemma MT_pred_mono
+      `{∀ m, Decision (P1 m)} `{∀ m, Decision (P2 m)} {M x ot1 ot2}
+      : MaxTime (filter P1 M) x ot1 -> MaxTime (filter P2 M) x ot2
         -> (∀ m, P1 m -> P2 m) -> ot1 ⊑ ot2.
     Proof.
       case ot1, ot2 => MT1 MT2 HP //.
@@ -264,7 +263,7 @@ Section RAMachine.
       abstract set_solver+.
     Qed.
 
-    Inductive thread_red (V : View) (M : memory) 
+    Inductive thread_red (V : View) (M : memory)
     : option positive -> option positive -> event -> View -> memory -> Prop :=
     | Th_Read m acc
               (InM : m ∈ msgs M)
@@ -305,7 +304,7 @@ Section RAMachine.
 
     Global Instance state_Inhabited : Inhabited state := _.
     Proof.
-      do 2!econstructor; exact: inhabitant. 
+      do 2!econstructor; exact: inhabitant.
     Qed.
 
     Definition msgs_ok M := ∀ m, m ∈ msgs M → msg_ok m ∧ view_ok (msgs M) m.(mview).
@@ -327,7 +326,7 @@ Section RAMachine.
         exists m; repeat split; by auto.
       Qed.
 
-      Lemma lookup_join (X Y : View) x ot : 
+      Lemma lookup_join (X Y : View) x ot :
         (X ⊔ Y) !! x = ot → X !! x = ot ∨ Y !! x = ot.
       Proof.
         rewrite lookup_union_with.
@@ -350,7 +349,7 @@ Section RAMachine.
         repeat case : (_ !! x) => [?|]; cbn; try auto.
         f_equal. solve_jsl.
       Qed.
-      
+
       Lemma msg_ok_write V (x : loc) v t
         (Lt : V !! x ⊏ Some t)
         : msg_ok (<x→v@t,(V ⊔ {[x := t]})>).
@@ -391,7 +390,7 @@ Section RAMachine.
             * transitivity t...
             * reflexivity.
             * trivial ...
-          + rewrite (_ : _ ⊔ (t + 1) = t + 1); 
+          + rewrite (_ : _ ⊔ (t + 1) = t + 1);
               first rewrite (_ : t + 1 ⊔ t = t + 1).
             * reflexivity.
             * apply/Pos.max_l_iff...
@@ -418,7 +417,7 @@ Section RAMachine.
             first (apply/elem_of_union; right; by apply elem_of_singleton);
             by auto.
         - move => [[y vy ty Ry] /= [Hm1 [Hm2 [Hm3 Hm4]]]]; subst.
-          exists (<z → vy @ tz, Ry>) => /=. 
+          exists (<z → vy @ tz, Ry>) => /=.
           repeat (split; first try (reflexivity || eassumption)).
           + apply elem_of_union; by left.
           + rewrite Hm4. solve_jsl.
@@ -474,12 +473,12 @@ Section RAMachine.
           move => y ty Hyty.
           case/lookup_join: Hyty.
           + move/VOk => [[? vy ? Ry] /= [Hm1 [Hm2 [Hm3 Hm4]]]]; subst.
-            exists (<y → vy @ ty, Ry>) => /=. 
+            exists (<y → vy @ ty, Ry>) => /=.
             repeat (split; first (reflexivity || eassumption)).
             rewrite Hm4. solve_jsl.
           + move/MOk : InM => [_] ROk /ROk [[? vy ? Ry] /=
                         [Hm1 [Hm2 [Hm3 Hm4]]]]; subst.
-            exists (<y → vy @ ty, Ry>) => /=. 
+            exists (<y → vy @ ty, Ry>) => /=.
             repeat (split; first (reflexivity || eassumption)).
             rewrite Hm4. solve_jsl.
         - split.
@@ -500,7 +499,7 @@ Section RAMachine.
         move => ? ?. apply/elem_of_union. by left.
       Qed.
 
-      Lemma thread_red_msgs_mono V M otr otw evt V' M' : 
+      Lemma thread_red_msgs_mono V M otr otw evt V' M' :
         thread_red V M otr otw evt V' M' → msgs_mono M M'.
       Proof.
         inversion 1; subst.
@@ -572,7 +571,7 @@ Section RAMachine.
 
       Lemma nats_complete_add_ins (M : memory) m (D : MS_msg_disj M m) T :
         (∃ m', m' ∈ msgs M ∧ m'.(mloc) = m.(mloc))
-        → nats_complete M T 
+        → nats_complete M T
         → nats_complete (add_ins M _ D) T.
       Proof.
         move => [m' [In EqLoc]] NCom m'' /=.
@@ -589,7 +588,7 @@ Section RAMachine.
           try (rewrite Hσ' => /Inv [? [? [? ?]]];
                 eexists; repeat split; first exact: Mono; by assumption).
         rewrite Hσ' lookup_insert_Some. move => [[Eq Eq']|[_ /Inv]].
-        - subst. eexists; repeat split; first apply (MT_msg_In Max'); auto. 
+        - subst. eexists; repeat split; first apply (MT_msg_In Max'); auto.
         - move => [? [? [? ?]]].
           eexists; repeat split; first (exact: Mono); by assumption.
       Qed.
@@ -774,7 +773,7 @@ Section RAMachine.
           inversion Alloc as [t_d t_a TD TA ? ?].
           econstructor; eauto.
           - apply: (MT_impl (MT_equiv_loc TD (MT_msg_max_filter TD))).
-            apply set_unfold_2 => /=. split. 
+            apply set_unfold_2 => /=. split.
             + move => [-> [H]]. move: H Alloc0 Local => <- H1 H2. intuition.
               cbn in H1; by rewrite H1 H2.
             + intuition. apply (MT_max TD). abstract set_solver.
@@ -822,7 +821,7 @@ Section RAMachine.
 
       End alloc_inv_helpers.
 
-      Lemma alloc_thread_red_alloc σ x otr otw evt V V' σ' acc 
+      Lemma alloc_thread_red_alloc σ x otr otw evt V V' σ' acc
         (HAcc : acc_of evt = Some acc)
         (ThRed : thread_red V (mem σ) otr otw evt V' (mem σ'))
         (DRFRed : drf_red σ σ' V otr otw evt)
@@ -897,10 +896,10 @@ Section RAMachine.
         apply/elem_of_union; left; assumption.
       Qed.
 
-      Lemma dealloc_inv_helper2 M1 
+      Lemma dealloc_inv_helper2 M1
             `(HMax : MaxTime M1 (mloc m) (Some (mtime m)))
             `(NEqx : x ≠ mloc m) v t R:
-        inhabited (MaxTime (M1 ∪ {[<x→v@t,R>]}) 
+        inhabited (MaxTime (M1 ∪ {[<x→v@t,R>]})
                            (mloc m) (Some (mtime m))).
       Proof.
         rewrite (_ : Some (mtime m) = Some (mtime m) ⊔ None); last done.
@@ -919,7 +918,7 @@ Section RAMachine.
         rewrite <-(MT_max TD _ In') in Alloc.
         cbn in Alloc.
         rewrite <-(MT_Some_Lt HMax (MT_msg TA)), <-(MT_msg_time TA) in Alloc.
-        - exact: (irreflexive_fw _ _ Alloc). 
+        - exact: (irreflexive_fw _ _ Alloc).
         - move: (MT_msg_In TA). abstract set_solver+.
         - exact: MT_msg_loc.
       Qed.
@@ -946,9 +945,9 @@ Section RAMachine.
           case: (decide (x0 = mloc m)) => [Eqx0|NEqx0];
             last exact: dealloc_inv_helper2.
           subst. exfalso. eapply max_D_not_allocated => //.
-        - move => m /elem_of_union [|/elem_of_singleton ->] /=; 
+        - move => m /elem_of_union [|/elem_of_singleton ->] /=;
             last (by inversion 0);
-            move => In Eq; move : (DInv _ In Eq); intros [[HMax]|]; 
+            move => In Eq; move : (DInv _ In Eq); intros [[HMax]|];
                     [left|right; exact: dealloc_inv_helper1].
           case: (decide (x0 = mloc m)) => [Eqx0|NEqx0];
             last exact: dealloc_inv_helper2.
@@ -1010,7 +1009,7 @@ Section RAMachine.
         σ.mem
         + pairwise_disj <mem>
         + any message is well-formed
-            msgs_ok <mem> 
+            msgs_ok <mem>
           TODO: view_ok is useless!!!
         + any message is compatible with the history before it
             alloc_inv M
@@ -1118,4 +1117,4 @@ Section RAMachine.
 
 End RAMachine.
 
-Arguments machine_red_safe [_ _ _ _ _] _ _ _.
\ No newline at end of file
+Arguments machine_red_safe [_ _ _ _ _] _ _ _.
diff --git a/theories/malloc.v b/theories/malloc.v
index 0c6b9800..5d3c4078 100644
--- a/theories/malloc.v
+++ b/theories/malloc.v
@@ -42,7 +42,7 @@ Section proof.
 
   Lemma alloc (E : coPset) :
     ↑physN ⊆ E →
-    {{{ ⎡PSCtx⎤ ∗ True }}}
+    {{{ ⎡PSCtx⎤ }}}
       (Alloc) @ E 
     {{{ (l: loc), RET (#l); (l ↦ A) }}}%stdpp.
   Proof.
@@ -50,7 +50,7 @@ Section proof.
     rewrite /own_loc_na /own_loc_prim /= valloc_local_eq /valloc_local_def.
     rewrite WP_Def.wp_eq /WP_Def.wp_def.
     iStartProof (uPred _).
-    iIntros (?) "[#kI _]". iIntros (V2 HV2) "Post". iIntros (V3 HV3 π) "kS".
+    iIntros (?) "#kI". iIntros (V2 HV2) "Post". iIntros (V3 HV3 π) "kS".
     iApply program_logic.weakestpre.wp_mask_mono;
     last (iApply (f_alloc with "[$kI $kS]"));
     first auto.
@@ -107,7 +107,7 @@ Section proof.
     have ?: (1 ≤ n) by omega.
     rewrite (_ : Z.to_nat n = S (Z.to_nat (n-1))); last first.
     { rewrite -Z2Nat.inj_succ; last omega. f_equal; omega. }
-    rewrite [seq _ _]/=. 
+    rewrite [seq _ _]/=.
     iDestruct "Big" as "[B0 Big]".
     rewrite Z2Nat.id; last omega.
     iApply (dealloc with "[$kI $B0]").
@@ -189,7 +189,7 @@ Section proof.
         rewrite /= Z2Nat.id; last omega.
         iDestruct "oL" as (h) "oL".
         by iExists _.
-      + subst. wp_op. 
+      + subst. wp_op.
         rewrite bool_decide_false; last by auto. wp_if_false.
         wp_bind ((mdealloc _) _)%E.
         iApply (mdealloc_spec with "kI [$Big]").
@@ -201,4 +201,4 @@ Section proof.
         wp_lam. iApply ("IHmalloc" with "[%] [$Post]"); omega.
   Qed.
 
-End proof.
\ No newline at end of file
+End proof.
diff --git a/theories/na.v b/theories/na.v
index 2dddbb37..872c1910 100644
--- a/theories/na.v
+++ b/theories/na.v
@@ -159,7 +159,7 @@ Section ExclusiveRules.
   Lemma na_read (E : coPset) l v :
     ↑ physN ⊆ E →
     {{{ ⎡PSCtx⎤ ∗ l ↦ v }}}
-      ([ #l]_na) @ E 
+      ([ #l]_na) @ E
     {{{ z, RET (#z); ⌜z = v⌝ ∗ (l ↦ v)  }}}.
   Proof.
     intros.
@@ -178,7 +178,7 @@ Section ExclusiveRules.
     - iNext. iIntros (v_r V') "(% & $ & oH & % & H)".
       iDestruct "H" as (V_1) "(% & % & %)".
       move: H6 => /value_at_hd_singleton [[->] ->].
-      iApply ("Post" $! v_r V' with "[%]"); first done.
+      iApply ("Post" $! v_r).
       iSplit; first done.
       iExists _, _. iFrame "oH oI". iPureIntro.
       exists (VInj v_r), V_1. by rewrite elem_of_singleton.
@@ -201,7 +201,7 @@ Section ExclusiveRules.
     iApply program_logic.weakestpre.wp_mask_mono; last (iApply (f_write_na with "[$kI $S $oH]"));
       [auto|auto|].
     iNext. iIntros (V' h') "(% & $ & oH & % & % & % & _ & %)". subst.
-    iApply ("Post" $! V' with "[%]"); [by etrans|].
+    iApply "Post".
     iExists _, _. iFrame "oH oI". iPureIntro. by apply alloc_init_local.
   Qed.
 
@@ -218,18 +218,18 @@ Section FractionalRules.
     intros q1 q2. split => ?.
     iSplit.
     - iIntros "na".
-      iDestruct "na" as (V') "(#Local & na)". 
+      iDestruct "na" as (V') "(#Local & na)".
       iDestruct (fractor_splitjoin _ _ (λ _ _ X, ⌜X = encode (VInj v)⌝)%I
                                        (λ _ _ X, ⌜X = encode (VInj v)⌝)%I
         with "[na]") as "[na1 na2]".
       { iApply (fractor_mono with "[$na]"). iIntros (?) "?". by iSplit. }
       iSplitL "na1"; iExists V'; by iFrame.
     - iIntros "(na1 & na2)".
-      iDestruct "na1" as (V1) "(Local1 & na1)".
-      iDestruct "na2" as (V2) "(Local2 & na2)".
+      iDestruct "na1" as (V1) "(% & na1)".
+      iDestruct "na2" as (V2) "(% & na2)".
       iCombine "na1" "na2" as "na".
       rewrite fractor_join_l.
-      iExists _. iSplit; first by iExact "Local1".
+      iExists V1. iSplitR; first done.
       iApply (fractor_mono with "[$na]").
       by iIntros (?) "[? _]".
   Qed.
@@ -242,7 +242,7 @@ Section FractionalRules.
      l ↦{q} v ⊢ l ↦{q/2} v ∗ l ↦{q/2} v.
   Proof. rewrite -{1}(Qp_div_2 q). apply na_frac_split. Qed.
 
-  Lemma na_frac_agree l v1 v2 q1 q2: 
+  Lemma na_frac_agree l v1 v2 q1 q2:
     l ↦{q1} v1 ∗ l ↦{q2} v2 ⊢ ⌜v1 = v2⌝.
   Proof.
     rewrite /own_loc_na_frac /own_loc_na /own_loc_prim /= valloc_local_eq /valloc_local_def.
@@ -328,7 +328,7 @@ Section FractionalRules.
     ↑ physN ⊆ E →
     ↑fracN .@ l ⊆ E →
     {{{ ⎡PSCtx⎤ ∗ l ↦{q} v }}}
-      [ #l ]_na @ E 
+      [ #l ]_na @ E
     {{{ z, RET (#z); ⌜z = v⌝ ∗ l ↦{q} v }}}.
   Proof.
     intros.
@@ -347,7 +347,7 @@ Section FractionalRules.
     - iNext. iIntros (v_r V') "(% & $ & oH & % & H)".
       iDestruct "H" as (V_1) "(% & % & %)".
       move: H8 => /value_at_hd_singleton [[->] ->].
-      iApply ("Post" $! _ _ with "[%]"); [etrans;done|].
+      iApply ("Post").
       iMod ("HClose" $! (λ _ _ X, ⌜X = encode (VInj v_r)⌝)%I with "[$oH $HX]")
         as "oNA".
       iModIntro.
@@ -359,14 +359,13 @@ Section FractionalRules.
     ↑ physN ⊆ E →
     ↑fracN .@ l ⊆ E →
     {{{ ⎡PSCtx⎤ ∗ ⌞l ↦{q} v⌟ β ∗ aSeen β }}}
-      [ #l ]_na @ E 
+      [ #l ]_na @ E
     {{{ z, RET (#z); ⌜z = v⌝ ∗ ⌞l ↦{q} v⌟ β }}}.
   Proof.
     intros.
     rewrite /own_loc_na_frac /own_loc_na /own_loc_prim /=
             valloc_local_eq /valloc_local_def aSeen_eq /aSeen_def
-            abs_pred_eq /abs_pred_def
-    .
+            abs_pred_eq /abs_pred_def.
     rewrite WP_Def.wp_eq /WP_Def.wp_def.
     iStartProof (uPred _).
     iIntros (?) "[#kI (ol & Seen)]". iIntros (? H1) "Post". iIntros (V2 H2 π) "S".
@@ -385,7 +384,7 @@ Section FractionalRules.
     - iNext. iIntros (v_r V') "(% & $ & oH & % & H)".
       iDestruct "H" as (V_1) "(% & % & %)".
       move: H10 => /value_at_hd_singleton [[?] ?]. subst v_r V_1.
-      iApply ("Post" $! _ _ with "[%]"); [etrans; done|].
+      iApply ("Post" $! _ _ with "[%]"); [by etrans|].
       iMod ("HClose" $! (λ _ _ X, ⌜X = encode (VInj v)⌝)%I with "[$oH $HX]")
         as "ol".
       iModIntro. iSplitL ""; first done.
@@ -450,4 +449,4 @@ End FractionalRules.
 
 
 Arguments na_read [_ _] _ _ [_] _.
-Arguments na_write [_ _] _ _ _ _.
\ No newline at end of file
+Arguments na_write [_ _] _ _ _ _.
diff --git a/theories/notation.v b/theories/notation.v
index 8d9cb3c1..22e745e3 100644
--- a/theories/notation.v
+++ b/theories/notation.v
@@ -34,9 +34,9 @@ Notation "'[' e ']_na'" := (Load na e%E)
 Notation "'[' e ']_at'" := (Load at_hack e%E)
   (at level 9, format "[ e ]_at") : expr_scope.
 (* The unicode ← is already part of the notation "_ ← _; _" for bind. *)
-Notation "'[' e1 ']_na' <- e2 " := (Store na e1%E e2%E) 
+Notation "'[' e1 ']_na' <- e2 " := (Store na e1%E e2%E)
   (at level 9, right associativity, format "[ e1 ]_na  <-  e2") : expr_scope.
-Notation "'[' e1 ']_at' <- e2 " := (Store at_hack e1%E e2%E) 
+Notation "'[' e1 ']_at' <- e2 " := (Store at_hack e1%E e2%E)
   (at level 9, right associativity, format "[ e1 ]_at  <-  e2") : expr_scope.
 Notation "(int) e" := (UnOp CastIntOp e%E)
   (at level 35, right associativity) : expr_scope.
@@ -48,6 +48,8 @@ Notation "e1 + e2" := (BinOp PlusOp e1%E e2%E)
   (at level 50, left associativity) : expr_scope.
 Notation "e1 - e2" := (BinOp MinusOp e1%E e2%E)
   (at level 50, left associativity) : expr_scope.
+Notation "e1 * e2" := (BinOp TimesOp e1%E e2%E)
+  (at level 40, left associativity) : expr_scope.
 Notation "e1 '`mod`' e2" := (BinOp ModOp e1%E e2%E)
   (at level 35) : expr_scope.
 Notation "e1 ≤ e2" := (BinOp LeOp e1%E e2%E) (at level 70) : expr_scope.
diff --git a/theories/persistor.v b/theories/persistor.v
index 6de2a732..47947d08 100644
--- a/theories/persistor.v
+++ b/theories/persistor.v
@@ -77,7 +77,7 @@ Section Persistor.
     Qed.
 
     Lemma persistor_splitjoin l φ Ψ1 Ψ2:
-      persisted l φ Ψ1 ∗ persisted l φ Ψ2 
+      persisted l φ Ψ1 ∗ persisted l φ Ψ2
       ⊣⊢ persisted l φ (λ l X, Ψ1 l X ∗ Ψ2 l X).
     Proof.
       iSplit.
diff --git a/theories/proofmode.v b/theories/proofmode.v
index 71ab6447..a4791f30 100644
--- a/theories/proofmode.v
+++ b/theories/proofmode.v
@@ -1,23 +1,13 @@
 From iris.proofmode Require Export tactics.
 From igps Require Import viewpred.
 
-Class AutoExt (P : Type) := auto_ext : P.
-Global Instance Frame_vPred_ctx {Σ} p (P : vProp Σ) :
-  AutoExt (V ⊑ V') →
-  (* Frame (P V') (Q V') R → *)
-  Frame p (monPred_car P V) (monPred_car P V') (True) | 20.
-Proof.
-  intros. unfold Frame. iIntros "[? R]".
-  rewrite (@auto_ext _ X). destruct p => /=; by iFrame.
-Qed.
 Global Instance Frame_vPred Σ p (P : vProp Σ) :
   (* Frame (P V') (Q V') R → *)
   (* class_instances.MakeSep ⌜V ⊑ V'⌝ R T → *)
   Frame p (monPred_car P V) (monPred_car P V') ⌜V ⊑ V'⌝ | 30.
 Proof. intros. unfold Frame. iIntros "[? %]". rewrite H. by destruct p => /=. Qed.
-Hint Extern 10 (AutoExt (?V1 ⊑ ?V2)) =>
-(unify V1 V2; fail 2) || unfold AutoExt; solve_jsl : typeclass_instances.
-
+Hint Extern 10 (IsBiIndexRel _ _) => unfold IsBiIndexRel; solve_jsl
+            : typeclass_instances.
 
 Ltac iViewCnt fn :=
       let V := fresh "V" in
@@ -30,4 +20,4 @@ Ltac iViewUp :=
               let X := fresh "X" in
               match type of HV with
               | ?V_old ⊑ V => setoid_rewrite HV; try (clear V_old HV; rename V into V_old)
-              end).
\ No newline at end of file
+              end).
diff --git a/theories/rsl.v b/theories/rsl.v
index 379b41e5..89a49b8d 100644
--- a/theories/rsl.v
+++ b/theories/rsl.v
@@ -1,7 +1,7 @@
 From iris.program_logic Require Export weakestpre.
 From iris.program_logic Require Import ectx_lifting.
 From stdpp Require Import functions.
-From iris.base_logic Require Import big_op.
+From iris.bi Require Import big_op.
 From iris.base_logic.lib Require Import saved_prop sts cancelable_invariants.
 From iris.algebra Require Import auth frac sts.
 From iris.proofmode Require Import tactics.
@@ -18,11 +18,11 @@ Implicit Types (I : gset gname) (l: loc) (V: View) (s: state) (q: frac).
 
 Class rslG Σ := RSLG {
   rsl_stsG :> stsG Σ rsl_sts;
-  rsl_savedPropG :> savedPropG Σ (ofe_funCF Z (ofe_funCF View idCF));
+  rsl_savedPropG :> savedAnythingG Σ (▶ (Z -c> View -c> ∙));
 }.
 
 Definition rslΣ : gFunctors
- := #[stsΣ rsl_sts; savedPropΣ (ofe_funCF Z (ofe_funCF View idCF))].
+ := #[stsΣ rsl_sts; savedAnythingΣ (▶ (Z -c> View -c> ∙))].
 
 Instance subG_rslΣ {Σ} : subG rslΣ Σ → rslG Σ.
 Proof. solve_inG. Qed.
@@ -34,22 +34,25 @@ Context `{!foundationG Σ, !rslG Σ}.
 Set Default Proof Using "Type".
 
 Local Notation iProp := (iProp Σ).
+Local Notation vProp := (vProp Σ).
+
+Local Notation to_saved Φ := (Next (λ x V, Φ x V) : (▶ (Z -c> View -c> ∙))%CF iProp).
 
 Implicit Types (Ψ: gname → Z → View → iProp) (φ: Z → View → iProp).
 
 
 Local Open Scope I.
 
-Definition sconv φ Ψ I : iProp 
+Definition sconv φ Ψ I : iProp
   := □ ∀ v V, φ v V → ([∗ set] i ∈ I, Ψ i v V).
 
 Definition conv φ φ' : iProp := □ ∀ v V, φ v V → φ' v V.
 
 Definition ress φ Ψ I : iProp :=
-  ▷ (sconv φ Ψ I) ∗ [∗ set] i ∈ I, saved_prop_own i (Ψ i).
+  ▷ (sconv φ Ψ I) ∗ [∗ set] i ∈ I, saved_anything_own i (to_saved (Ψ i)).
 
 Definition ress2 φ Ψ I : iProp :=
-  ([∗ set] i ∈ I, saved_prop_own i (Ψ i))
+  ([∗ set] i ∈ I, saved_anything_own i (to_saved (Ψ i)))
   ∗ ▷ □ ∀ i v V, ⌜i ∈ I⌝ → (Ψ i v V ↔ True).
 
 Definition rsl_inv l φ φd s: iProp :=
@@ -82,34 +85,34 @@ Definition rsl_inv l φ φd s: iProp :=
       ∗ ∃ Ψ, ress2 φ Ψ I0%I
   end.
 
-Definition RSLInv φ φd l jγ : iProp 
+Definition RSLInv φ φd l jγ : iProp
   := ∃ (j γ: gname), ⌜jγ = encode (j,γ)⌝ ∗ sts_inv γ (rsl_inv l φ φd).
 
-Definition RelRaw V (φ: Z → vPred) φR jγ : iProp
+Definition RelRaw V (φ: Z → vProp) φR jγ : iProp
   := ∃ s (j γ: gname), ⌜jγ = encode (j,γ)⌝
         ∗ sts_own γ s ∅ ∗ ⌜alloc_local (rhist s) V⌝
-        ∗ ▷ conv φ φR ∗ saved_prop_own j φR.
+        ∗ ▷ conv φ φR ∗ saved_anything_own j (to_saved φR).
 
-Definition AcqRaw V (φ: Z → vPred) jγ : iProp
+Definition AcqRaw V (φ: Z → vProp) jγ : iProp
   := ∃ s (j γ: gname) i Ψi, ⌜jγ = encode (j,γ)⌝
         ∗ sts_own γ s {[ Change i ]}
         ∗ ⌜s ∈ i_states i⌝
         ∗ ⌜alloc_local (rhist s) V⌝ ∗ ▷ conv Ψi φ
-        ∗ saved_prop_own i Ψi.
+        ∗ saved_anything_own i (to_saved Ψi).
 
 Definition RMWAcqRaw V jγ : iProp
   := ∃ s (j γ: gname), ⌜jγ = encode (j,γ)⌝
-        ∗ sts_own γ s ∅ 
+        ∗ sts_own γ s ∅
         ∗ ∃ I0 h, ⌜s = cUninit I0 h ∨ s = cInit I0 h⌝
         ∗ ⌜alloc_local h V⌝.
 
 Definition InitRaw V jγ : iProp
   := ∃ s (j γ: gname), ⌜jγ = encode (j,γ)⌝
-        ∗ sts_own γ s ∅ 
+        ∗ sts_own γ s ∅
         ∗ ⌜∃ h, ((∃ I1 I2 Sub, s = aInit I1 I2 h Sub) ∨ (∃ I0, s = cInit I0 h))
                   ∧ init_local h V⌝.
 
-Instance RelRaw_persistent V (φ: Z → vPred) φR jγ
+Instance RelRaw_persistent V (φ: Z → vProp) φR jγ
   : Persistent (RelRaw V φ φR jγ) := _.
 
 Instance RMWAcqRaw_persistent V jγ : Persistent (RMWAcqRaw V jγ) := _.
@@ -119,8 +122,8 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Section Helpers.
   Set Default Proof Using "Type".
 
-  Lemma big_sepS_gblock_ends_ins_update 
-    l (h: History) p0 (Ψ: Val * View → iProp) 
+  Lemma big_sepS_gblock_ends_ins_update
+    l (h: History) p0 (Ψ: Val * View → iProp)
     (Disj: (h ## {[p0]})%stdpp)
     (GB: gblock_ends_ins l h p0 (gblock_ends l h) (gblock_ends l ({[p0]} ∪ h)))
     : Ψ p0 ∗ ([∗ set] p ∈ gblock_ends l h, Ψ p)
@@ -139,37 +142,37 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Qed.
 
   Lemma rsl_saved_pred_unfold i (Ψ1 Ψ2: (Z -c> View -c> ∙)%CF iProp):
-    saved_prop_own i Ψ1 ∗ saved_prop_own i Ψ2 ⊢ □ ▷ (∀ v V, Ψ1 v V ≡ Ψ2 v V).
+    saved_anything_own i (to_saved Ψ1) ∗ saved_anything_own i (to_saved Ψ2)
+    ⊢ □ ▷ (∀ v V, Ψ1 v V ≡ Ψ2 v V).
   Proof.
-    iIntros "(Saved1 & Saved2)". 
-    iDestruct (saved_prop_agree i with "Saved1 Saved2") as "#Eq".
-    by do 2 setoid_rewrite ofe_funC_equivI.
+    iIntros "(Saved1 & Saved2)".
+    iDestruct (saved_anything_agree i with "Saved1 Saved2") as "#Eq".
+    rewrite later_equivI. by do 2 setoid_rewrite ofe_fun_equivI.
   Qed.
 
   Lemma rsl_saved_pred_unfold_big_sepS
     I0 i Ψi Ψ (In: i ∈ I0):
-    (saved_prop_own i Ψi ∗ [∗ set] i0 ∈ I0, saved_prop_own i0 (Ψ i0))
+    saved_anything_own i (to_saved Ψi) ∗
+    ([∗ set] i0 ∈ I0, saved_anything_own i0 (to_saved (Ψ i0)))
     ⊢ □ ▷ (∀ v V, Ψi v V ≡ Ψ i v V).
   Proof.
     iIntros "(#Saved & #SSaved)".
-    iDestruct (big_sepS_dup _ _
-                (λ i, saved_prop_own i (Ψ i : (Z -c> View -c> ∙)%CF iProp)) In 
-                with "[$SSaved]") as "(#SΨi1 & _)".
+    iDestruct (big_sepS_dup _ _ (λ i, saved_anything_own i (to_saved (Ψ i))) In
+            with "[$SSaved]") as "(#SΨi1 & _)".
       { iIntros "#?". by iSplit. }
     iApply (rsl_saved_pred_unfold i with "[]"). iFrame "Saved SΨi1".
   Qed.
 
   Lemma rsl_ress2_insert_c φ I0 i (NI: i ∉ I0):
-    saved_prop_own i ((λ _ _, True): Z → View → iProp) ∗ (∃ Ψ, ress2 φ Ψ I0)
+    saved_anything_own i (to_saved (λ _ _, True)) ∗ (∃ Ψ, ress2 φ Ψ I0)
     ⊢ ∃ Ψ, ress2 φ Ψ ({[i]} ∪ I0).
   Proof.
     iIntros "(#Si & ress2)".
     iDestruct "ress2" as (Ψ) "(Saved & #True)".
     iExists (<[i:= (λ _ _, True)]> Ψ). iSplitL "Saved".
-    - rewrite (big_sepS_fn_insert
-                (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp)));
+    - rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f)));
        last exact NI. by iFrame.
-    - iNext. rewrite persistently_elim. iIntros "!#".
+    - iNext. rewrite affinely_persistently_elim. iIntros "!#".
       iIntros (i' v' V') "In".
       iDestruct "In" as %[->%elem_of_singleton_1|In]%elem_of_union.
       + rewrite fn_lookup_insert. iSplit; by iIntros "_".
@@ -180,7 +183,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
 
   Lemma rsl_inv_insert_c I0 h i s l φ φd
     (Eqc: s = cUninit I0 h ∨ s = cInit I0 h) (NI: i ∉ rISet s)
-    : saved_prop_own i ((λ _ _, True): Z → View → iProp)
+    : saved_anything_own i (to_saved (λ _ _, True))
       ∗ rsl_inv l φ φd s ⊢ rsl_inv l φ φd (state_ISet_insert i s).
   Proof.
     iIntros "(#Saved & Inv)".
@@ -190,14 +193,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       iFrame; by iApply (rsl_ress2_insert_c with "[$Saved $ress2]").
   Qed.
 
-  Lemma conv_id φ: True ⊢ conv φ φ.
-  Proof.
-    iIntros "!#". by iIntros (? ?) "?".
-  Qed.
+  Lemma conv_id φ: conv φ φ.
+  Proof. iIntros (??) "!# $". Qed.
 
-  Lemma conv_switch_a i Ψi Ψ (φ: Z → vPred) v:
+  Lemma conv_switch_a i Ψi Ψ (φ: Z → vProp) v:
     conv Ψi φ ∗ □ (∀ (v0 : Z) (V0 : View), Ψi v0 V0 ≡ Ψ i v0 V0)
-    ⊢ conv (<[v:=λ _ : View, True]> (Ψ i)) (<[v:=True%VP]> φ).
+    ⊢ conv (<[v:=λ _ : View, True]> (Ψ i)) (<[v:=True]> φ).
   Proof.
     iIntros "(#Conv & #Eq)".
     iIntros "!#".
@@ -229,14 +230,13 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
 
 
   Lemma saved_pred_switch_a i i' Ψ v I1 (NI: i' ∉ I1) (In: i ∈ I1):
-    ([∗ set] x ∈ I1, saved_prop_own x (Ψ x))
-    ∗ saved_prop_own i' (<[v:=λ _ : View, True]> (Ψ i))
-    ⊢ [∗ set] i0 ∈ ISet_switch i i' I1, 
-                saved_prop_own i0 (<[i':=<[v:=λ _ : View, True]> (Ψ i)]> Ψ i0).
+    ([∗ set] x ∈ I1, saved_anything_own x (to_saved (Ψ x)))
+    ∗ saved_anything_own i' (to_saved (<[v:=λ _ : View, True]> (Ψ i)))
+    ⊢ [∗ set] i0 ∈ ISet_switch i i' I1,
+          saved_anything_own i0 (to_saved (<[i':=<[v:=λ _ : View, True]> (Ψ i)]> Ψ i0)).
   Proof.
     iIntros "(#Saved & #?)".
-    rewrite (big_sepS_fn_insert
-                (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp)));
+    rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f)));
       last first.
       { by move => /elem_of_difference [? _]. }
     rewrite (big_sepS_delete _ _ _ In).
@@ -245,13 +245,13 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
 
   End Helpers.
 
-  Lemma RelRaw_write l (φ φd φR: Z → vPred) v V n E π:
+  Lemma RelRaw_write l (φ φd φR: Z → vProp) v V n E π:
     nclose physN ⊆ E →
-      {{{ PSCtx ∗ ▷ Seen π V ∗ ▷ φ v V ∗ ▷ φd v V 
+      {{{ PSCtx ∗ ▷ Seen π V ∗ ▷ φ v V ∗ ▷ φd v V
             ∗ ▷ RSLInv φR φd l n ∗ RelRaw V φ φR n }}}
         ([(#l)]_at <- #v, V) @ E
       {{{ V', RET (#(), V') ;
-            ⌜V ⊑ V'⌝ ∗ Seen π V' 
+            ⌜V ⊑ V'⌝ ∗ Seen π V'
             ∗ ▷ RSLInv φR φd l n ∗ RelRaw V' φ φR n ∗ InitRaw V' n}}}%stdpp.
   Proof.
     intros.
@@ -269,7 +269,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
 
     iDestruct "Inv" as "(Hist & Inv)".
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first auto.
+    iApply (wp_mask_mono _ (↑physN)); first auto.
 
     iApply (f_write_at with "[$kI $Hist $Seen]"); first auto.
 
@@ -292,65 +292,54 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
             iExists Ψ. iDestruct "ress" as "(#EConv & #Saved)".
             destruct H11 as [V'' EqV]. rewrite -EqV in H6.
             iFrame "EConv Saved". iNext. iSplitL "Itpr".
-            { iApply (big_sepS_mono _ _ ({[(VInj v, V');(A, V'')]}: History)).
+            { iApply (big_sepS_subseteq _ ({[(VInj v, V');(A, V'')]}: History)).
               - by rewrite -H6.
-              - reflexivity.
               - rewrite big_sepS_insert; last by rewrite elem_of_singleton.
                 rewrite big_sepS_singleton.
-                iSplitR ""; last auto. iApply "EConv". iApply "Conv".
-                by iApply (vPred_mono V). }
-            { iApply (big_sepS_mono _ _ ({[(VInj v, V');(A, V'')]}: History)).
+                iSplitR ""; last auto. iApply "EConv". by iApply "Conv". }
+            { iApply (big_sepS_subseteq _ ({[(VInj v, V');(A, V'')]}: History)).
               - by rewrite -H6.
-              - reflexivity.
               - rewrite big_sepS_insert; last by rewrite elem_of_singleton.
-                rewrite big_sepS_singleton.
-                iSplitR ""; last auto. by iApply (vPred_mono V). }
+                rewrite big_sepS_singleton. auto. }
           + iDestruct "ress" as (Ψ) "ress". (* TODO: duplicated proofs *)
             iExists Ψ. iDestruct "ress" as "[[#EConv #Saved] [ress ressD]]".
             iFrame "EConv Saved". iNext.
             pose p : Val * View := (VInj v, V').
             iSplitL "Itpr ress".
-            * iApply (big_sepS_mono _ _  ({[VInj v, V']} ∪ h));
-                [by rewrite H6|reflexivity|].
+            * iApply (big_sepS_subseteq _  ({[VInj v, V']} ∪ h)); [by rewrite H6|].
               rewrite big_sepS_insert;
                 last by apply disjoint_singleton_r in H7.
               iFrame "ress".
-              iApply (big_sepS_mono _ _ I1); [auto|auto|].
-              iApply "EConv". iApply "Conv". by iApply (vPred_mono V).
-            * iApply (big_sepS_mono _ _  ({[VInj v, V']} ∪ h));
-                [by rewrite H6|reflexivity|].
+              iApply (big_sepS_subseteq _ I1); [done|].
+              iApply "EConv". by iApply "Conv".
+            * iApply (big_sepS_subseteq _  ({[VInj v, V']} ∪ h)); [by rewrite H6|].
               rewrite big_sepS_insert;
                 last by apply disjoint_singleton_r in H7.
-              iFrame "ressD". by iApply (vPred_mono V).
+              by iFrame "ressD".
           + iDestruct "ress" as (Ψ) "#ress2".
-            destruct H11 as [V'' EqV]. rewrite -EqV in H6.
+            destruct H11 as [V'' EqV]. rewrite -EqV in H6 .
             iSplitL "Itpr"; last iSplitL "ItprD".
             { iNext.
-              iApply (big_sepS_mono _ _ ({[(VInj v, V');(A, V'')]}: History)).
+              iApply (big_sepS_subseteq _ ({[(VInj v, V');(A, V'')]}: History)).
               * rewrite -H6. by apply subseteq_gset_filter.
-              * reflexivity.
               * rewrite big_sepS_insert; last by rewrite elem_of_singleton.
-                rewrite big_sepS_singleton. iSplitR ""; last done.
-                iApply "Conv". by iApply (vPred_mono V). }
+                rewrite big_sepS_singleton. iSplitR ""; last done. by iApply "Conv". }
             { iNext.
-              iApply (big_sepS_mono _ _ ({[(VInj v, V');(A, V'')]}: History)).
+              iApply (big_sepS_subseteq _ ({[(VInj v, V');(A, V'')]}: History)).
               * by rewrite -H6.
-              * reflexivity.
               * rewrite big_sepS_insert; last by rewrite elem_of_singleton.
-                rewrite big_sepS_singleton.
-                iSplitR ""; last done. by iApply (vPred_mono V). }
-            { iExists Ψ. iDestruct "ress2" as "(? & ?)". by iFrame "ress2". }
+                rewrite big_sepS_singleton. auto. }
+            { iExists Ψ. iDestruct "ress2" as "(? & ?)". by iSplit. }
           + iDestruct "ress" as "(ress & ressD & ress2)". iFrame "ress2".
             pose p : Val * View := (VInj v, V'). iNext.
             iSplitL "Itpr ress".
-            { iApply (big_sepS_mono _ _  (gblock_ends l ({[VInj v, V']} ∪ h)));
-                [by rewrite H6|reflexivity|].
+            { iApply (big_sepS_subseteq _  (gblock_ends l ({[VInj v, V']} ∪ h)));
+                [by rewrite H6|].
               iApply (big_sepS_gblock_ends_ins_update); [auto|..|iFrame].
               * apply blocks.gblock_ends_ins_update; last auto.
                 apply (hTime_ok_mono _ _ h'); last auto. by rewrite H6.
-              * iApply "Conv". by iApply (vPred_mono V). }
-            { iApply (big_sepS_mono _ _  ({[VInj v, V']} ∪ h));
-                [by rewrite H6|reflexivity|].
+              * by iApply "Conv". }
+            { iApply (big_sepS_subseteq  _  ({[VInj v, V']} ∪ h)); [by rewrite H6|].
               rewrite big_sepS_insert;
                 last by apply disjoint_singleton_r in H7.
               by iFrame. } }
@@ -366,7 +355,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Qed.
 
 
-  Lemma AcqRaw_read l (φ φd φR: Z → vPred) V n E π:
+  Lemma AcqRaw_read l (φ φd φR: Z → vProp) V n E π:
     nclose physN ⊆ E →
     {{{ inv physN PSInv ∗ ▷ Seen π V ∗ (∀ V v, φd v V -∗ φd v V ∗ φd v V)
         ∗ ▷ RSLInv φR φd l n ∗ AcqRaw V φ n ∗ InitRaw V n }}}
@@ -374,7 +363,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     {{{ (v:Z) V', RET (#v, V');
         ⌜V ⊑ V'⌝ ∗ Seen π V'
         ∗ ▷ RSLInv φR φd l n
-        ∗ AcqRaw V' (<[v:=True%VP]>φ) n ∗ InitRaw V' n
+        ∗ AcqRaw V' (<[v:=True]>φ) n ∗ InitRaw V' n
         ∗ ▷ φ v V' ∗ ▷ φd v V'}}}%stdpp.
   Proof.
     iIntros (? Φ) "(#kI & kS & Dup & oI & AR & IR) Post".
@@ -402,7 +391,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     destruct (frame_steps_Init_branch s2 s ∅)
       as [h Hh]; [auto|by eexists|].
 
-    iApply (wp_mask_mono (↑physN)); first auto.
+    iApply (wp_mask_mono _ (↑physN)); first auto.
     iApply wp_fupd.
 
     iApply (f_read_at with "[$Hist $kS $kI]"); first auto.
@@ -418,10 +407,8 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       iDestruct "ress" as (Ψ) "(#Saved & ress & ressD)".
       iDestruct "Saved" as "(SConv & Saved)".
 
-      iMod (saved_prop_alloc_strong 
-                  (<[v:=(λ _, True)]> (Ψ i): (Z -c> View -c> ∙)%CF iProp)
-                  (rISet s))
-      as (i') "[% #SΨi']".
+      iMod (saved_anything_alloc_strong (to_saved (<[v:=(λ _, True)]> (Ψ i))) (rISet s))
+        as (i') "[% #SΨi']".
       rewrite Hs /= in H8.
 
       (* State update *)
@@ -449,9 +436,6 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         as "(Hφd & ressD)".
         { iNext. iIntros "?". by iApply "Dup". }
 
-      iAssert (▷ (φd v) V') with "[Hφd]" as "Hφd".
-        { iNext. by iApply (vPred_mono V1). }
-
       (* Predicate equality *)
       iDestruct (rsl_saved_pred_unfold_big_sepS with "[$SΨi $Saved]") as "#Eq".
         { by apply Sub. }
@@ -465,7 +449,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       iDestruct "Hφ" as "(Hφi & Hφ)".
 
       iAssert (▷ (φ v) V') with "[Hφi]" as "Hφi".
-        { iNext. iApply (vPred_mono V1); first auto.
+        { iNext. iApply (monPred_mono _ V1); first auto.
           iApply "Conv". by iRewrite ("Eq" $! v V1). }
 
       iMod (Hist_hTime_ok with "[$Hist $kI]") as "(Hist & %)"; first auto.
@@ -483,8 +467,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         iExists ((<[i':= <[v:=λ _ : View, True]> (Ψ i)]>)Ψ).
         iSplitL ""; first iSplitL "".
         * iNext. iApply (sconv_switch_a with "[$SConv]"); [auto|by apply Sub].
-        * iApply (saved_pred_switch_a with "[$Saved $SΨi']");
-            [auto|by apply Sub].
+        * iApply saved_pred_switch_a; auto.
         * iNext.
           rewrite (big_sepS_delete _ _ _ H6). iSplitL "Hφ".
           { rewrite (big_sepS_fn_insert
@@ -504,13 +487,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         rewrite /sts_own. iFrame. rewrite /s' /=. iFrame "SΨi'".
         repeat iSplitL ""; first auto.
         * iPureIntro. by apply state_ISet_switch_included.
-        * iPureIntro. apply state_ISet_switch_alloc_local.
-          rewrite Hs. by apply (alloc_local_Mono _ _ V).
+        * iPureIntro. apply state_ISet_switch_alloc_local. rewrite Hs /= -(_:V⊑V') //.
         * iModIntro. iNext. by iApply (conv_switch_a with "[$Conv]").
       + iExists s',j,γ. rewrite /sts_own.  iFrame "Own1".
         iSplitL ""; first auto. iPureIntro. exists h. split.
         * left. rewrite /s' Hs. by do 3 eexists.
-        * by apply (init_local_Mono _ _ V).
+        * rewrite -(_:V⊑V') //.
 
     - rewrite Hs /= in HInit, HAlloc.
       rewrite {2}Hs.
@@ -522,8 +504,6 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         as "(Hφd & ressD)".
         { iNext. iIntros "?". by iApply "Dup". }
 
-      iAssert (▷ (φd v) V') with "[Hφd]" as "Hφd".
-        { iNext. by iApply (vPred_mono V1). }
 
       (* Predicate equality *)
       assert (His: i ∈ I1).
@@ -550,7 +530,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       + iExists s,j,γ,i, Ψi. rewrite /sts_own. iFrame "SΨi Own'".
         repeat iSplitL ""; [auto|..].
         * iPureIntro. by apply HSub.
-        * iPureIntro. rewrite Hs /=. by apply (alloc_local_Mono _ _ V).
+        * iPureIntro. rewrite Hs /= -(_:V⊑V') //.
         * iModIntro. iNext. iIntros "!#". iIntros (v2 V2) "?".
           case (decide (v = v2)) => [->|NEqv].
           { by rewrite fn_lookup_insert. }
@@ -558,15 +538,15 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       + iExists s,j,γ. rewrite /sts_own. iFrame "Own1".
         iSplitL ""; first auto. iPureIntro. exists h. split.
         * right. by eexists.
-        * by apply (init_local_Mono _ _ V).
+        * rewrite -(_:V⊑V') //.
       + iModIntro. iNext. iApply "Conv". iRewrite ("Eq" $! v V').
         iDestruct ("True" $! i v V' with "[#]") as "EC"; [done|].
         by iApply "EC".
   Qed.
 
-  Lemma RMWAcqRaw_split l (φ φd: Z → vPred) V n E:
+  Lemma RMWAcqRaw_split l (φ φd: Z → vProp) V n E:
     ▷ RSLInv φ φd l n ∗ RMWAcqRaw V n
-    ⊢ |={E}=> ▷ RSLInv φ φd l n ∗ RMWAcqRaw V n ∗ AcqRaw V (λ _, True)%VP n.
+    ⊢ |={E}=> ▷ RSLInv φ φd l n ∗ RMWAcqRaw V n ∗ AcqRaw V (λ _, True) n.
   Proof.
     iIntros "(Inv & #RR)".
     iDestruct "RR" as (s j γ) "(% & #Own & RR)".
@@ -575,14 +555,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     apply encode_inj in H2. inversion H2. subst j1 γ1.
 
     iMod (sts_acc with "[$Inv $Own]") as (s') "(% & Inv & SClose)".
-    iMod (saved_prop_alloc_strong 
-                  ((λ _ _, True): (Z -c> View -c> ∙)%CF iProp)
-                  (rISet s'))
+    iMod (saved_anything_alloc_strong (to_saved (λ _ _, True)) (rISet s'))
       as (i) "[% #Hi]".
     destruct (frame_steps_RMW_branch _ _ _ H _ _ H0) as [I' [h' Eq]].
 
     iMod ("SClose" $! (state_ISet_insert i s') {[Change i]}
-              with "[Inv]") as "(Inv & Own')". 
+              with "[Inv]") as "(Inv & Own')".
         { iSplitL "".
           - iPureIntro. by apply state_ISet_insert_steps.
           - iNext. by iApply (rsl_inv_insert_c with "[$Hi $Inv]"). }
@@ -592,7 +570,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     - iNext. iExists j,γ. by iFrame.
     - iExists s,j,γ. iFrame "Own". iSplitL ""; first auto.
       iExists I0, h. by iSplitL "".
-    - iExists _, j,γ, i, (λ _ : Z, True%VP). iFrame "Own' Hi".
+    - iExists _, j,γ, i, (λ _ _, True). iFrame "Own' Hi".
       repeat iSplitL ""; first auto.
       + iPureIntro. by apply state_ISet_insert_included.
       + iPureIntro.
@@ -600,16 +578,15 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         apply (alloc_local_mono_rsl_state s).
         * destruct H0 as [Hs|Hs]; by rewrite Hs.
         * apply (frame_steps_steps _ _ _ H).
-      + iNext. by rewrite -conv_id.
+      + by iIntros "!> !# *".
   Qed.
 
-  Lemma AcqRaw_alloc_vs l h (φ φd: Z → vPred) V:
+  Lemma AcqRaw_alloc_vs l h (φ φd: Z → vProp) V:
     Hist l h ∗ ⌜alloc_local h V⌝ ∗ ⌜uninit h⌝
     ⊢ |==> ∃ n, ▷ RSLInv φ φd l n ∗ RelRaw V φ φ n ∗ AcqRaw V φ n.
   Proof.
     iIntros "(Hist & % & %)".
-    iMod (saved_prop_alloc (φ: (Z -c> View -c> ∙)%CF iProp))
-      as (i) "#SΨi".
+    iMod (saved_anything_alloc (to_saved φ)) as (i) "#SΨi".
 
     iAssert (rsl_inv l φ φd (aUninit {[i]} h)) with "[Hist]" as "Inv".
       { iFrame (H0) "Hist".
@@ -619,7 +596,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
           by rewrite big_sepS_singleton fn_lookup_insert.
         - by rewrite big_sepS_singleton fn_lookup_insert. }
 
-    iMod (own_alloc (sts_auth (aUninit {[i]} h) {[Change i]})) 
+    iMod (own_alloc (sts_auth (aUninit {[i]} h) {[Change i]}))
       as (γ) "Own".
       { apply sts_auth_valid, tok_disj_state_singleton => /=.
         by apply elem_of_singleton. }
@@ -641,16 +618,16 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       + iNext. by rewrite -conv_id.
   Qed.
 
-  Lemma AcqRaw_alloc (φ φd: Z → vPred) E π V:
+  Lemma AcqRaw_alloc (φ φd: Z → vProp) E π V:
     nclose physN ⊆ E →
     {{{ inv physN PSInv ∗ ▷ Seen π V }}}
-      (Alloc, V) @ E
-    {{{ (l: loc) V' n X, RET (LitV $ LitLoc l, V');
+      (base.Alloc, V) @ E
+    {{{ (l: loc) V' n X, RET (#l, V');
           ⌜V ⊑ V'⌝ ∗ Seen π V' ∗ Info l 1 X
         ∗ ▷ RSLInv φ φd l n ∗ RelRaw V' φ φ n ∗ AcqRaw V' φ n }}}%stdpp.
   Proof.
     iIntros (? Φ) "(kI & kS) Post".
-    iApply (wp_mask_mono (↑physN)); first auto.
+    iApply (wp_mask_mono _ (↑physN)); first auto.
     iApply wp_fupd.
 
     iApply (f_alloc with "[$kI $kS]"). iNext.
@@ -662,13 +639,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Qed.
 
 
-  Lemma RMWAcqRaw_alloc_vs l h (φ φd: Z → vPred) V:
+  Lemma RMWAcqRaw_alloc_vs l h (φ φd: Z → vProp) V:
     Hist l h ∗ ⌜alloc_local h V⌝ ∗ ⌜uninit h⌝
     ⊢ |==> ∃ n, ▷ RSLInv φ φd l n ∗ RelRaw V φ φ n ∗ RMWAcqRaw V n.
   Proof.
     iIntros "(Hist & % & %)".
-    iMod (saved_prop_alloc (φ: (Z -c> View -c> ∙)%CF iProp))
-      as (i) "#SΨi".
+    iMod (saved_anything_alloc (to_saved φ)) as (i) "#SΨi".
 
     iAssert (rsl_inv l φ φd (cUninit ∅ h)) with "[$Hist]" as "Inv".
       { iFrame (H0). iExists (λ _ _ _,True).
@@ -692,16 +668,16 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     - iFrame "Own'". iExists ∅, h. iFrame (H). by iLeft.
   Qed.
 
-  Lemma RMWAcqRaw_alloc (φ φd: Z → vPred) E π V:
+  Lemma RMWAcqRaw_alloc (φ φd: Z → vProp) E π V:
     nclose physN ⊆ E →
     {{{ inv physN PSInv ∗ ▷ Seen π V }}}
-      (Alloc, V) @ E
-    {{{ (l: loc) V' n X, RET (LitV $ LitLoc l, V');
+      (base.Alloc, V) @ E
+    {{{ (l: loc) V' n X, RET (#l, V');
           ⌜V ⊑ V'⌝ ∗ Seen π V' ∗ Info l 1 X
         ∗ ▷ RSLInv φ φd l n ∗ RelRaw V' φ φ n ∗ RMWAcqRaw V' n }}}%stdpp.
   Proof.
     iIntros (? Φ) "(kI & kS) Post".
-    iApply (wp_mask_mono (↑physN)); [done|].
+    iApply (wp_mask_mono _ (↑physN)); [done|].
     iApply wp_fupd.
 
     iApply (f_alloc with "[$kI $kS]").
@@ -712,14 +688,13 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     iApply ("Post" $! l V' n X with "[$VV' $kS' $Info $IRA]").
   Qed.
 
-  Lemma AcqRaw_init_vs l h (φ φd: Z → vPred) v V:
+  Lemma AcqRaw_init_vs l h (φ φd: Z → vProp) v V:
     Hist l h ∗ ⌜alloc_local h V⌝ ∗ ⌜init l h⌝ ∗ ⌜h = {[VInj v, V]}⌝
     ∗ ▷ φ v V ∗ ▷ φd v V
     ⊢ |==> ∃ n, ▷ RSLInv φ φd l n ∗ RelRaw V φ φ n ∗ AcqRaw V φ n ∗ InitRaw V n.
   Proof.
     iIntros "(Hist & % & % & % & Hφ & Hφd)".
-    iMod (saved_prop_alloc (φ: (Z -c> View -c> ∙)%CF iProp))
-      as (i) "#SΨi".
+    iMod (saved_anything_alloc (to_saved φ)) as (i) "#SΨi".
 
     set s:= aInit {[i]} {[i]} h (reflexivity {[i]}).
     iAssert (rsl_inv l φ φd s)
@@ -733,7 +708,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         - by rewrite H1 !big_sepS_singleton fn_lookup_insert.
         - by rewrite H1 big_sepS_singleton. }
 
-    iMod (own_alloc (sts_auth s {[Change i]})) 
+    iMod (own_alloc (sts_auth s {[Change i]}))
       as (γ) "Own".
       { apply sts_auth_valid, tok_disj_state_singleton => /=.
         by apply elem_of_singleton. }
@@ -757,14 +732,13 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       + exists (VInj v), V. by rewrite H1 elem_of_singleton.
   Qed.
 
-  Lemma RMWAcqRaw_init_vs l h (φ φd: Z → vPred) v V:
+  Lemma RMWAcqRaw_init_vs l h (φ φd: Z → vProp) v V:
     Hist l h ∗ ⌜alloc_local h V⌝ ∗ ⌜init l h⌝ ∗ ⌜h = {[VInj v, V]}⌝
     ∗ ▷ φ v V ∗ ▷ φd v V
     ⊢ |==> ∃ n, ▷ RSLInv φ φd l n ∗ RelRaw V φ φ n ∗ RMWAcqRaw V n ∗ InitRaw V n.
   Proof.
     iIntros "(Hist & % & % & % & Hφ & Hφd)".
-    iMod (saved_prop_alloc (φ: (Z -c> View -c> ∙)%CF iProp))
-      as (i) "#SΨi".
+    iMod (saved_anything_alloc (to_saved φ)) as (i) "#SΨi".
 
     set s:= cInit ∅ h.
     iAssert (rsl_inv l φ φd s)
@@ -772,8 +746,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       { iFrame (H0) "Hist".
         iSplitL "Hφ"; last iSplitL "Hφd".
         - iNext.
-          iApply (big_sepS_mono _ _ h with "[Hφ]");
-            [by apply subseteq_gset_filter |reflexivity|].
+          iApply (big_sepS_subseteq _ h with "[Hφ]"); [by apply subseteq_gset_filter|].
           by rewrite H1 big_sepS_singleton.
         - by rewrite H1 big_sepS_singleton.
         - iExists (λ _ _ _,True). iSplitL "".
@@ -801,7 +774,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Qed.
 
 
-  Lemma AcqRaw_init l v h (φ φd: Z → vPred) E π V:
+  Lemma AcqRaw_init l v h (φ φd: Z → vProp) E π V:
     nclose physN ⊆ E →
     {{{ inv physN PSInv ∗ ▷ Seen π V
         ∗ ▷ Hist l h ∗ ⌜alloc_local h V⌝
@@ -813,14 +786,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Proof.
     iIntros (? Φ) "(#kI & kS & Hist & #Alloc & Hφ & Hφd) Post".
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
 
     iApply (f_write_na with "[$kI $kS $Hist $Alloc]").
     iNext.
     iIntros (V' h') "(% & kS' & Hist' & % & % & #HInit & % & #Hh')".
 
-    iDestruct (vPred_mono V V' with "Hφ") as "Hφ"; first auto.
-    iDestruct (vPred_mono V V' with "Hφd") as "Hφd"; first auto.
     iMod (AcqRaw_init_vs l h' φ φd v V' with "[$Hist' $Hh' $HInit $Hφ $Hφd]")
       as (n) "(Inv & Rel & Acq & Init)".
       { iPureIntro. by apply alloc_init_local. }
@@ -828,26 +799,24 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     by iApply ("Post" $! V' n with "[$kS' $Inv $Rel $Acq $Init]").
   Qed.
 
-  Lemma RMWAcqRaw_init l v h (φ φd: Z → vPred) E π V:
+  Lemma RMWAcqRaw_init l v h (φ φd: Z → vProp) E π V:
     nclose physN ⊆ E →
     {{{ inv physN PSInv ∗ ▷ Seen π V
         ∗ ▷ Hist l h ∗ ⌜alloc_local h V⌝
         ∗ ▷ φ v V ∗ ▷ φd v V}}}
       ([ #l ]_na <- #v, V) @ E
     {{{ V' n, RET (#(), V');
-          ⌜V ⊑ V'⌝ ∗ Seen π V' ∗ ▷ RSLInv φ φd l n 
+          ⌜V ⊑ V'⌝ ∗ Seen π V' ∗ ▷ RSLInv φ φd l n
           ∗ RelRaw V' φ φ n ∗ RMWAcqRaw V' n ∗ InitRaw V' n }}}%stdpp.
   Proof.
     iIntros (? Φ) "(#kI & kS & Hist & #Alloc & Hφ & Hφd) Post".
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first done.
+    iApply (wp_mask_mono _ (↑physN)); first done.
 
     iApply (f_write_na with "[$kI $kS $Hist $Alloc]").
     iNext.
     iIntros (V' h') "(% & kS' & Hist' & % & % & #HInit & % & #Hh')".
 
-    iDestruct (vPred_mono V V' with "Hφ") as "Hφ"; first auto.
-    iDestruct (vPred_mono V V' with "Hφd") as "Hφd"; first auto.
     iMod (RMWAcqRaw_init_vs l h' φ φd v V' with "[$Hist' $Hh' $HInit $Hφ $Hφd]")
       as (n) "(Inv & Rel & RMW & Init)".
       { iPureIntro. by apply alloc_init_local. }
@@ -856,17 +825,17 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Qed.
 
   Lemma RMWAcqRaw_CAS
-    l (φ φd: Z → vPred) (P: vPred) (R: bool → Z → vPred) (v_r v_w: Z) V n E π:
+    l (φ φd: Z → vProp) (P: vProp) (R: bool → Z → vProp) (v_r v_w: Z) V n E π:
     nclose physN ⊆ E →
     □ (∀ V v, φd v V -∗ φd v V ∗ φd v V)
     ⊢
-    {{{ (P ∗ ▷ φ v_r ∗ ▷ φd v_r ={ E }=∗ ▷ φ v_w∗ ▷ φd v_w ∗ R true v_r)%VP V
-        ∗ (∀ v, ■ (v ≠ v_r) → P ∗ ▷ φd v ={ E }=∗ R false v)%VP V
-        ∗ inv physN PSInv ∗ ▷ Seen π V 
+    {{{ (P ∗ ▷ φ v_r ∗ ▷ φd v_r ={ E }=∗ ▷ φ v_w∗ ▷ φd v_w ∗ R true v_r) V
+        ∗ (∀ v, ⌜v ≠ v_r⌝ → P ∗ ▷ φd v ={ E }=∗ R false v) V
+        ∗ inv physN PSInv ∗ ▷ Seen π V
         ∗ ▷ RSLInv φ φd l n ∗ RMWAcqRaw V n ∗ InitRaw V n ∗ P V}}}
-        (CAS #l #v_r #v_w, V) @ E
-    {{{ (b: bool) (v': Z) V', RET (LitV $ LitInt b, V');
-          ⌜V ⊑ V'⌝ ∗ Seen π V' 
+        (base.CAS #l #v_r #v_w, V) @ E
+    {{{ (b: bool) (v': Z) V', RET (#(base.Z_of_bool b), V');
+          ⌜V ⊑ V'⌝ ∗ Seen π V'
           ∗ ▷ RSLInv φ φd l n ∗ RMWAcqRaw V' n ∗ InitRaw V' n ∗ R b v' V'}}}.
   Proof.
     iIntros (?) "#Dup !#".
@@ -902,7 +871,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
     iDestruct "Inv" as "(Hist & >% & ress & ressD & ress2)".
 
     iApply wp_fupd.
-    iApply (wp_mask_mono (↑physN)); first auto.
+    iApply (wp_mask_mono _ (↑physN)); first auto.
     iApply (f_CAS with "[$kI $kS $Hist]").
       { iPureIntro.
         rewrite (_: h = (rhist s)); last by rewrite Hs. exact HInit. }
@@ -916,15 +885,10 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       as "(Hφd & ressD)".
       { iNext. iIntros "?". by iApply "Dup". }
 
-    iAssert (▷ (φd v) V') with "[Hφd]" as "Hφd".
-      { iNext. by iApply (vPred_mono V1). }
-
-    iDestruct (vPred_mono V V' with "P") as "P"; first auto.
-
     iDestruct "CCase" as "[CS|CF]".
     - iDestruct "CS" as "(% & % & % & % & % & % & % & %)".
       subst v.
-      iMod (Hist_hTime_ok with "[$kI $Hist']") as "(Hist' & %)"; first auto. 
+      iMod (Hist_hTime_ok with "[$kI $Hist']") as "(Hist' & %)"; first auto.
 
       assert (HO: hTime_ok l ({[VInj v_w, V']} ∪ h)).
         { apply (hTime_ok_mono _ _ h'); [by rewrite H11|auto]. }
@@ -948,13 +912,11 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
           - apply disjoint_empty_r.
           - apply disjoint_empty_r. }
       rewrite -sts_op_auth_frag_up.
-      iDestruct "Own" as "(Own & #Own')". 
+      iDestruct "Own" as "(Own & #Own')".
 
       (* get full interpretation *)
       rewrite (big_sepS_delete _ _ _ Hp').
       iDestruct "ress" as "(Hφ & ress)".
-      iAssert (▷ (φ v_r) V') with "[Hφ]" as "Hφ".
-        { iNext. by iApply (vPred_mono V1). }
 
       iMod ("CSucc" $! V' with "[] [$P $Hφ $Hφd]") as "(Hφ & Hφd & R)";
         first auto.
@@ -974,11 +936,9 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         iExists j, γ. iModIntro. iSplitL ""; first by iNext.
         iExists s'. iFrame "Own Hist' ress2".
         iNext. iSplit; first auto. iNext. iSplitR "ressD"; last first.
-        * iApply (big_sepS_mono _ _ _ h' with "[$ressD]");
-            last reflexivity.
+        * iApply (big_sepS_subseteq _ _ h' with "[$ressD]").
           by rewrite H11 union_comm_L.
-        * iApply (big_sepS_mono _ _ _ (gblock_ends l h') with "[$ress]");
-            last reflexivity.
+        * iApply (big_sepS_subseteq _ _ (gblock_ends l h') with "[$ress]").
           rewrite H11.
           apply gblock_ends_ins_sL_update => //; last by apply HO.
           by apply gblock_ends_ins_update.
@@ -992,8 +952,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         * by exists (VInj v_w), V'.
 
     - iDestruct "CF" as "(% & #NE & % & %)".
-      iMod ("CFail" $! v V' with "[] NE [] [$P $Hφd]") as "R";
-        [auto|auto|].
+      iMod ("CFail" $! v V' with "[//] [NE//] [$P $Hφd]") as "R".
 
       rewrite -sts_op_auth_frag_up.
       iDestruct "Own" as "(Own & #Own')".
@@ -1002,14 +961,15 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       iSplitR ""; last iSplitL "".
       + rewrite /RSLInv /sts_inv. iModIntro. iNext. iExists j,γ.
         iSplitL ""; first auto. iExists s. iFrame.
-        rewrite Hs H9 /rsl_inv. by iFrame.
+        rewrite Hs H9 /rsl_inv /ress2. iFrame. iSplit=>//.
+        iDestruct "ress2" as (Ψ) "[??]". iExists Ψ. iFrame.
       + iExists s,j,γ. rewrite /sts_own. iFrame "Own'". iSplitL ""; first auto.
         iExists I0, h. iSplitL ""; iPureIntro; first by right.
-        rewrite Hs in HAlloc. by apply (alloc_local_Mono _ _ V).
+        rewrite Hs in HAlloc. rewrite -(_:V⊑V') //.
       + iExists s,j,γ. rewrite /sts_own. iFrame "Own'". iSplitL ""; first auto.
         iPureIntro. exists h. split.
         * right. by eexists.
-        * rewrite Hs in HInit. by apply (init_local_Mono _ _ V).
+        * rewrite Hs in HInit. rewrite -(_:V⊑V') //.
   Qed.
 
   Section RSL_SplitJoin.
@@ -1017,8 +977,8 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
   Lemma ress_inv_join i1 i2 i I1 I2 Ψi Ψi1 Ψi2 φR Ψ:
       I2 ⊆ I1 →
       i ∉ I1 → i1 ∈ I2 → i2 ∈ I2 → i1 ≠ i2 →
-      saved_prop_own i1 Ψi1 ∗ saved_prop_own i2 Ψi2
-        ∗ saved_prop_own i Ψi
+      saved_anything_own i1 (to_saved Ψi1) ∗ saved_anything_own i2 (to_saved Ψi2)
+        ∗ saved_anything_own i (to_saved Ψi)
         ∗ conv (λ v V, Ψi1 v V ∗ Ψi2 v V) Ψi
         ∗ ress φR Ψ I2
         ⊢ ress φR (<[i:=(Ψi: Z → View → iProp)]>Ψ) ({[i]} ∪ I2 ∖ {[i1; i2]}).
@@ -1053,17 +1013,16 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
           - by iRewrite ("Heq2" $! v' V'). }
       rewrite (big_sepS_fn_insert (λ _ (f: Z → View → iProp), f v' V')); [|done].
       iFrame.
-      by iApply (big_sepS_mono _ _ _ _ HSub with "[$ress]").
-    - rewrite (big_sepS_fn_insert
-                (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp))); [|done].
+      by iApply (big_sepS_subseteq _ _ _ HSub with "[$ress]").
+    - rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f))); [|done].
       iSplit; [done|].
-      by iApply (big_sepS_mono _ _ _ _ HSub).
+      by iApply (big_sepS_subseteq _ _ _ HSub).
   Qed.
 
   Lemma ress2_inv_join i1 i2 i I1 Ψi Ψi1 Ψi2 φR:
     i ∉ I1 → i1 ∈ I1 → i2 ∈ I1 → i1 ≠ i2 →
-    saved_prop_own i1 Ψi1 ∗ saved_prop_own i2 Ψi2
-      ∗ saved_prop_own i Ψi
+    saved_anything_own i1 (to_saved Ψi1) ∗ saved_anything_own i2 (to_saved Ψi2)
+      ∗ saved_anything_own i (to_saved Ψi)
       ∗ conv (λ v V, Ψi1 v V ∗ Ψi2 v V) Ψi
       ∗ (∃ Ψ , ress2 φR Ψ I1)
       ⊢ ∃ Ψ , ress2 φR Ψ ({[i]} ∪ I1 ∖ {[i1; i2]}).
@@ -1087,13 +1046,11 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
 
     iSplit.
     - rewrite big_sepS_insert ?fn_lookup_insert; [|done]. iSplit; [done|].
-      iApply (big_sepS_mono
-                (λ (i: gname), saved_prop_own i (Ψ i: (Z -c> View -c> ∙)%CF iProp))).
-      + apply HSub.
+      iApply (big_sepS_mono (λ (i: gname), saved_anything_own i (to_saved (Ψ i)))).
       + move => i' In' /=. apply gset_difference_subseteq in In'.
         rewrite fn_lookup_insert_ne => //.
-        move => ?; subst. by apply Ini. 
-      + done.
+        move => ?; subst. by apply Ini.
+      + iApply (@big_sepS_subseteq with "Saved2"). set_solver.
     - iDestruct (rsl_saved_pred_unfold i1 Ψi1 (Ψ i1) with "[$Hi1 $SΨi1]") as "#Heq1".
       iDestruct (rsl_saved_pred_unfold i2 Ψi2 (Ψ i2) with "[$Hi2 $SΨi2]") as "#Heq2".
       iNext. iIntros "!#". iIntros (i' v' V') "Ini'".
@@ -1113,10 +1070,10 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         by apply elem_of_difference in Ini' as [].
   Qed.
 
-  Lemma rsl_inv_join l i1 i2 i s Ψi Ψi1 Ψi2 (φd: Z → vPred) φR:
+  Lemma rsl_inv_join l i1 i2 i s Ψi Ψi1 Ψi2 (φd: Z → vProp) φR:
       i ∉ rISet s → i1 ∈ rISet2 s → i2 ∈ rISet2 s → i1 ≠ i2 →
-        saved_prop_own i1 Ψi1 ∗ saved_prop_own i2 Ψi2
-        ∗ saved_prop_own i Ψi
+        saved_anything_own i1 (to_saved Ψi1) ∗ saved_anything_own i2 (to_saved Ψi2)
+        ∗ saved_anything_own i (to_saved Ψi)
         ∗ conv (λ v V, Ψi1 v V ∗ Ψi2 v V) Ψi
         ∗ rsl_inv l φR φd s ⊢ rsl_inv l φR φd (state_ISet_join i1 i2 i s).
   Proof.
@@ -1146,7 +1103,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         iSplitL "".
           * iApply ress_inv_join; try done. by apply Sub. by apply Sub.
             iSplitL; first by iApply "Hi1".
-            iSplitL; first by iApply "Hi2". 
+            iSplitL; first by iApply "Hi2".
             by iFrame "#".
           * iNext. iApply (big_sepS_impl _ _ h with "[$RS]").
             iIntros "!#". iIntros ([[| |v'] V']) "% ress"; [done|done|..].
@@ -1162,20 +1119,20 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
               { iApply "Conv". iSplitL "rΨi1".
                 - by iRewrite ("Heq1" $! v' V').
                 - by iRewrite ("Heq2" $! v' V'). }
-            iApply (big_sepS_mono (λ (i: gname), Ψ i v' V')); last by iFrame.
+            iApply (big_sepS_subseteq _ (I2 ∖ {[i1 ]} ∖ {[i2]})).
               { by rewrite difference_twice_union. }
+            iApply (@big_sepS_mono with "ress").
             move => i' In' /=. apply gset_difference_subseteq in In'.
-            rewrite fn_lookup_insert_ne => //.
-            move => ?; subst. by apply Ini, Sub.
+            rewrite fn_lookup_insert_ne => //. set_solver.
       - iDestruct "Inv" as "(? & ? & ress2)". iFrame.
         iApply ress2_inv_join; try done; []. by iFrame "∗ #".
       - iDestruct "Inv" as "(? & ? & ? & ? & ress2)". iFrame.
         iApply ress2_inv_join; try done; []. by iFrame "∗ #".
     Qed.
 
-    Lemma AcqRaw_join l (φ1 φ2 φR φd: Z → vPred) n V E:
+    Lemma AcqRaw_join l (φ1 φ2 φR φd: Z → vProp) n V E:
       ▷ RSLInv φR φd l n ∗ AcqRaw V φ1 n ∗ AcqRaw V φ2 n
-      ⊢ |={E}=> ▷ RSLInv φR φd l n ∗ AcqRaw V (λ v, φ1 v ∗ φ2 v)%VP n.
+      ⊢ |={E}=> ▷ RSLInv φR φd l n ∗ AcqRaw V (λ v, φ1 v ∗ φ2 v) n.
     Proof.
       iIntros "(Inv & RR1 & RR2)".
       rewrite {1}/RSLInv.
@@ -1194,7 +1151,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       rewrite -sts_op_frag; [|auto|auto|auto].
 
       iCombine "OwnA" "Own" as "OwnA".
-      iDestruct (own_valid with "OwnA") 
+      iDestruct (own_valid with "OwnA")
         as %[Hs1 Hs2]%sts_auth_frag_valid_inv%elem_of_intersection.
       assert (Hi1: s ∈ i_states i1).
         { eapply (sts.up_subseteq s1); eauto. by apply i_states_closed. }
@@ -1203,8 +1160,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       rewrite sts_op_auth_frag; [|done|by apply sts.closed_op].
 
       pose Ψi := (λ v V, Ψi1 v V ∗ Ψi2 v V).
-      iMod (saved_prop_alloc_strong (Ψi: (Z -c> View -c> ∙)%CF iProp) (rISet s))
-        as (i) "[% #Hi]".
+      iMod (saved_anything_alloc_strong (to_saved Ψi)) as (i) "[% #Hi]".
       pose s' := state_ISet_join i1 i2 i s.
 
       assert (alloc_local (rhist s') V).
@@ -1232,8 +1188,8 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
           * by iApply "Conv2".
     Qed.
 
-    Lemma RelRaw_split (φ1 φ2 φR: Z → vPred) n V:
-      RelRaw V (λ v, φ1 v ∨ φ2 v)%VP φR n
+    Lemma RelRaw_split (φ1 φ2 φR: Z → vProp) n V:
+      RelRaw V (λ v, φ1 v ∨ φ2 v) φR n
       ⊢  RelRaw V φ1 φR n ∗ RelRaw V φ2 φR n.
     Proof.
       iIntros "#RR". iDestruct "RR" as (s j γ) "(% & #oF & % & #Conv & #Saved)".
@@ -1242,9 +1198,9 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         iApply "Conv"; [by iLeft|by iRight].
     Qed.
 
-    Lemma RelRaw_join (φ1 φ2 φR1 φR2: Z → vPred) n V:
+    Lemma RelRaw_join (φ1 φ2 φR1 φR2: Z → vProp) n V:
       RelRaw V φ1 φR1 n ∗ RelRaw V φ2 φR2 n
-      ⊢  RelRaw V (λ v, φ1 v ∨ φ2 v)%VP φR1 n.
+      ⊢  RelRaw V (λ v, φ1 v ∨ φ2 v) φR1 n.
     Proof.
       iIntros "(RR1 & RR2)".
       iDestruct "RR1" as (s1 j1 γ1) "(% & #Rel1 & % & #Conv1 & #SP1)".
@@ -1257,12 +1213,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       - iRewrite ("LEq" $! v' V'). by iApply "Conv2".
     Qed.
 
-    Lemma ress_inv_split_1 i1 i2 i I0 I2 Ψi (φ1 φ2 : Z → vPred) φR Ψ:
+    Lemma ress_inv_split_1 i1 i2 i I0 I2 Ψi (φ1 φ2 : Z → vProp) φR Ψ:
       I2 ⊆ I0 →
       i ∈ I2 → i1 ∉ I0 → i2 ∉ I0 → i1 ≠ i2 →
-      saved_prop_own i1 φ1 ∗ saved_prop_own i2 φ2
-        ∗ saved_prop_own i Ψi
-        ∗ conv Ψi (λ v, φ1 v ∗ φ2 v)%VP
+      saved_anything_own i1 (to_saved φ1) ∗ saved_anything_own i2 (to_saved φ2)
+        ∗ saved_anything_own i (to_saved Ψi)
+        ∗ conv Ψi (λ v, φ1 v ∗ φ2 v)
         ∗ ress φR Ψ I2
         ⊢ ress φR (<[i2:=(φ2: Z → View → iProp)]> (<[i1:=(φ1: Z → View → iProp)]> Ψ))
                   ({[i1; i2]} ∪ I2 ∖ {[i]}).
@@ -1294,23 +1250,21 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
           { by iRewrite ("Heq" $! v' V'). }
         iDestruct ("Conv" $! v' V' with "rΨi") as "(R1 & R2)".
 
-        iApply (big_sepS_mono _ _  ({[i2]} ∪ ({[i1]} ∪ I2 ∖ {[i]}))); [done|done|].
+        iApply (big_sepS_subseteq _  ({[i2]} ∪ ({[i1]} ∪ I2 ∖ {[i]}))); [done|].
         rewrite (big_sepS_fn_insert (λ _ (f: Z → View → iProp), f v' V')); [|done].
         rewrite (big_sepS_fn_insert (λ _ (f: Z → View → iProp), f v' V')); [|done].
         by iFrame.
-      - iApply (big_sepS_mono _ _  ({[i2]} ∪ ({[i1]} ∪ I2 ∖ {[i]}))); [done|done|].
-        rewrite (big_sepS_fn_insert
-                  (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp))); [|done].
-        rewrite (big_sepS_fn_insert
-                  (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp))); [|done].
+      - iApply (big_sepS_subseteq _  ({[i2]} ∪ ({[i1]} ∪ I2 ∖ {[i]}))); [done|].
+        rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f))); [|done].
+        rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f))); [|done].
         by repeat iSplit.
     Qed.
 
-    Lemma ress2_inv_split i1 i2 i I0 Ψi (φ1 φ2 : Z → vPred) φR:
+    Lemma ress2_inv_split i1 i2 i I0 Ψi (φ1 φ2 : Z → vProp) φR:
       i ∈ I0 → i1 ∉ I0 → i2 ∉ I0 → i1 ≠ i2 →
-      saved_prop_own i1 φ1 ∗ saved_prop_own i2 φ2
-        ∗ saved_prop_own i Ψi
-        ∗ conv Ψi (λ v, φ1 v ∗ φ2 v)%VP
+      saved_anything_own i1 (to_saved φ1) ∗ saved_anything_own i2 (to_saved φ2)
+        ∗ saved_anything_own i (to_saved Ψi)
+        ∗ conv Ψi (λ v, φ1 v ∗ φ2 v)
         ∗ (∃ Ψ , ress2 φR Ψ I0)
         ⊢ ∃ Ψ , ress2 φR Ψ ({[i1; i2]} ∪ I0 ∖ {[i]}).
     Proof.
@@ -1333,14 +1287,12 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
       iExists (<[i2:=(φ2: Z → View → iProp)]> (<[i1:=(φ1: Z → View → iProp)]> Ψ)).
 
       iSplit.
-      - iApply (big_sepS_mono _ _  ({[i2]} ∪ ({[i1]} ∪ I0 ∖ {[i]}))); [done|done|..].
-        rewrite (big_sepS_fn_insert
-                  (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp))); [|done].
-        rewrite (big_sepS_fn_insert
-                   (λ i0 f, saved_prop_own i0 (f: (Z -c> View -c> ∙)%CF iProp))); [|done].
+      - iApply (big_sepS_subseteq _  ({[i2]} ∪ ({[i1]} ∪ I0 ∖ {[i]}))); [done|..].
+        rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f))); [|done].
+        rewrite (big_sepS_fn_insert (λ i0 f, saved_anything_own i0 (to_saved f))); [|done].
         by iFrame "#".
       - iDestruct (rsl_saved_pred_unfold i Ψi (Ψ i) with "[$Hi $SΨi]") as "#Heq".
-        iNext. iIntros "!#". iIntros (i' v' V') "%".
+        iNext. iIntros "!#". iIntros (i' v' V' Hi').
         iDestruct ("EConv" $! i v' V' with "[#]") as "EC"; [done|].
         case (decide (i' = i1)) => [->|?].
           { rewrite fn_lookup_insert_ne ?fn_lookup_insert; [|done].
@@ -1356,16 +1308,16 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
             by iDestruct ("Conv" $! v' V' with "[$rΨi]") as "(_ & ?)". }
         rewrite !fn_lookup_insert_ne; [|auto|auto].
         case (decide (i' = i)) => [->|?]; first auto.
-        iApply "EConv". iPureIntro. move : H2.
+        iApply "EConv". iPureIntro. move:Hi'.
         rewrite -assoc_L !elem_of_union elem_of_difference !elem_of_singleton.
         move => [?|[?|[? _]]] => //.
     Qed.
 
-    Lemma rsl_inv_split l i1 i2 i s Ψi (φ1 φ2 φd: Z → vPred) φR:
+    Lemma rsl_inv_split l i1 i2 i s Ψi (φ1 φ2 φd: Z → vProp) φR:
       i ∈ rISet2 s → i1 ∉ rISet s → i2 ∉ rISet s → i1 ≠ i2 →
-        saved_prop_own i1 φ1 ∗ saved_prop_own i2 φ2
-        ∗ saved_prop_own i Ψi
-        ∗ conv Ψi (λ v, φ1 v ∗ φ2 v)%VP
+        saved_anything_own i1 (to_saved φ1) ∗ saved_anything_own i2 (to_saved φ2)
+        ∗ saved_anything_own i (to_saved Ψi)
+        ∗ conv Ψi (λ v, φ1 v ∗ φ2 v)
         ∗ rsl_inv l φR φd s ⊢ rsl_inv l φR φd (state_ISet_split i1 i2 i s).
     Proof.
       move => Ini Ini1 Ini2 iEq.
@@ -1393,7 +1345,7 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
             iAssert (Ψi v' V') with "[rΨi]" as "rΨi".
               { by iRewrite ("Heq" $! v' V'). }
             iDestruct ("Conv" $! v' V' with "rΨi") as "(R1 & R2)".
-            iApply (big_sepS_mono _ _  ({[i2]} ∪ ({[i1]} ∪ I2 ∖ {[i]}))); try done; [|].
+            iApply (big_sepS_subseteq _  ({[i2]} ∪ ({[i1]} ∪ I2 ∖ {[i]}))).
               { rewrite union_assoc. apply union_mono_r.
                 by rewrite union_comm. }
             rewrite (big_sepS_fn_insert (λ _ (f: Z → View → iProp), f v' V'));
@@ -1406,13 +1358,13 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
                 by apply Ini1, Sub. }
             by iFrame.
       - iDestruct "Inv" as "(? & ? & ress2)". iFrame.
-        iApply ress2_inv_split; try done; []. by iFrame "∗ #".
+        iApply ress2_inv_split=>//. by iFrame "∗ #".
       - iDestruct "Inv" as "(? & ? & ? & ? & ress2)". iFrame.
-        iApply ress2_inv_split; try done; []. by iFrame "∗ #".
+        iApply ress2_inv_split=>//. by iFrame "∗ #".
     Qed.
 
-    Lemma AcqRaw_split l (φ1 φ2 φR φd: Z → vPred) V n E:
-      ▷ RSLInv φR φd l n ∗ AcqRaw V (λ v : Z, (φ1 v ∗ φ2 v)%VP) n 
+    Lemma AcqRaw_split l (φ1 φ2 φR φd: Z → vProp) V n E:
+      ▷ RSLInv φR φd l n ∗ AcqRaw V (λ v : Z, (φ1 v ∗ φ2 v)) n
       ⊢ |={E}=> ▷ RSLInv φR φd l n ∗ AcqRaw V φ1 n ∗ AcqRaw V φ2 n.
     Proof.
       iIntros "[Inv AC]".
@@ -1427,9 +1379,8 @@ Instance InitRaw_persistent V jγ : Persistent (InitRaw V jγ) := _.
         { move: Step. apply sts.closed_steps; last auto.
           by apply i_states_closed. }
 
-      iMod (saved_prop_alloc_strong (φ1: (Z -c> View -c> ∙)%CF iProp) (rISet s'))
-        as (i1) "[% #Hi1]".
-      iMod (saved_prop_alloc_strong (φ2: (Z -c> View -c> ∙)%CF iProp) (rISet s' ∪ {[i1]}))
+      iMod (saved_anything_alloc_strong (to_saved φ1) (rISet s')) as (i1) "[% #Hi1]".
+      iMod (saved_anything_alloc_strong (to_saved φ2) (rISet s' ∪ {[i1]}))
         as (i2) "[Hi2' #Hi2]"; iDestruct "Hi2'" as %Hi2.
       rewrite ->not_elem_of_union, elem_of_singleton in Hi2; destruct Hi2.
 
diff --git a/theories/rsl_instances.v b/theories/rsl_instances.v
index 59bea7cf..0bee0fd6 100644
--- a/theories/rsl_instances.v
+++ b/theories/rsl_instances.v
@@ -16,86 +16,48 @@ Context `{!foundationG Σ, !rslG Σ}.
 Set Default Proof Using "Type".
 
 Local Notation iProp := (iProp Σ).
-Local Notation vPred := (@vPred Σ).
+Local Notation vProp := (vProp Σ).
 Implicit Types  (l: loc) (V: View)  (q: frac)
-                (Ψ: gname → Z → View → iProp) (φ: vPred).
+                (Ψ: gname → Z → View → iProp) (φ: vProp).
 
 Local Open Scope I.
 
   Section Plain.
   Context `{!persistorG Σ}.
 
-  Definition Rel' l (φ φd: Z → vPred) V
-    := ∃ (φR: Z → vPred), persisted l (RSLInv φR φd) (λ _ n, RelRaw V φ φR n).
-
-  Definition Acq' l (φ φd: Z → vPred) V
-    := ∃ (φR: Z → vPred), persisted l (RSLInv φR φd) (λ _ n, AcqRaw V φ n).
-
-  Definition RMWAcq' l (φR φd: Z → vPred) V
-    := persisted l (RSLInv φR φd) (λ _ n, RMWAcqRaw V n).
-
-  Definition Init' l V
-    := ∃ (φR φd: Z → vPred), persisted l (RSLInv φR φd) (λ _ n, InitRaw V n).
-
-  Lemma vPred_Rel'_mono l (φ φd: Z → vPred): vPred_Mono (Rel' l φ φd).
-  Proof.
-    repeat intro. iIntros "Rel".
-    iDestruct "Rel" as (φR) "FR". iDestruct "FR" as (jγ) "(FV & ?)".
-    iExists φR, jγ. iFrame.
-    iDestruct "FV" as (? ? ?) "(? & ? & % & ?)".
-    iExists _, _, _. iFrame.
-    iPureIntro. by apply (alloc_local_Mono _ _ V).
-  Qed.
-
-  Canonical Structure Rel l (φ φd: Z → vPred)
-    := Build_vPred _ (vPred_Rel'_mono l φ φd).
-
-  Lemma vPred_Acq'_mono l (φ φd: Z → vPred): vPred_Mono (Acq' l φ φd).
-  Proof.
-    repeat intro. iIntros "Acq".
-    iDestruct "Acq" as (φR) "FA". iDestruct "FA" as (jγ) "(FV & ?)".
-    iExists φR, jγ. iFrame.
-    iDestruct "FV" as (? ? ? ? ?) "(? & ? & ? & % & ?)".
-    iExists _,_,_,_,_. iFrame.
-    iPureIntro. by apply (alloc_local_Mono _ _ V).
-  Qed.
-
-  Canonical Structure Acq l (φ φd: Z → vPred)
-    := Build_vPred _ (vPred_Acq'_mono l φ φd).
-
-  Lemma vPred_RMWAcq'_mono l (φ φd: Z → vPred): vPred_Mono (RMWAcq' l φ φd).
-  Proof.
-    repeat intro. iIntros "FA".
-    iDestruct "FA" as (jγ) "(FV & ?)". iExists jγ. iFrame.
-    iDestruct "FV" as (s j γ) "(? & ? & Eq)".
-    iDestruct "Eq" as (? ?) "(? & %)".
-    iExists _,_,_. iFrame. iExists _, _. iFrame "Eq".
-    iPureIntro. by apply (alloc_local_Mono _ _ V).
-  Qed.
-
-  Canonical Structure RMWAcq l (φ φd: Z → vPred)
-    := Build_vPred _ (vPred_RMWAcq'_mono l φ φd).
-
-  Lemma vPred_Init'_mono l: vPred_Mono (Init' l).
-  Proof.
-    repeat intro. iIntros "Init".
-    iDestruct "Init" as (φR φd) "FA". iDestruct "FA" as (jγ) "(FV & ?)".
-    iExists φR, φd, jγ. iFrame.
-    iDestruct "FV" as (s j γ) "(? & ? & Eq)".
-    iDestruct "Eq" as %[? [? ?]].
-    iExists _,_,_. iFrame. iPureIntro. eexists.
-    split => //. by apply (init_local_Mono _ _ V).
-  Qed.
-
-  Canonical Structure Init l
-    := Build_vPred _ (vPred_Init'_mono l).
-
-  Lemma Rel_write l (φ φd: Z → vPred) v (E : coPset)
+  Program Definition Rel_def l (φ φd: Z → vProp) : vProp := MonPred (λ V,
+      ∃ (φR: Z → vProp), persisted l (RSLInv φR φd) (λ _ n, RelRaw V φ φR n)) _.
+  Next Obligation. unfold RelRaw, persisted. solve_proper. Qed.
+  Definition Rel_aux : seal Rel_def. by eexists. Qed.
+  Definition Rel := unseal Rel_aux.
+  Definition Rel_eq : Rel = _ := seal_eq _.
+
+  Program Definition Acq_def l (φ φd: Z → vProp) : vProp := MonPred (λ V,
+      ∃ (φR: Z → vProp), persisted l (RSLInv φR φd) (λ _ n, AcqRaw V φ n)) _.
+  Next Obligation. unfold AcqRaw, persisted. solve_proper. Qed.
+  Definition Acq_aux : seal Acq_def. by eexists. Qed.
+  Definition Acq := unseal Rel_aux.
+  Definition Acq_eq : Acq = _ := seal_eq _.
+
+  Program Definition RMWAcq_def l (φR φd: Z → vProp) : vProp := MonPred (λ V,
+      persisted l (RSLInv φR φd) (λ _ n, RMWAcqRaw V n)) _.
+  Next Obligation. unfold RMWAcqRaw, persisted. solve_proper. Qed.
+  Definition RMWAcq_aux : seal RMWAcq_def. by eexists. Qed.
+  Definition RMWAcq := unseal RMWAcq_aux.
+  Definition RMWAcq_eq : Rel = _ := seal_eq _.
+
+  Program Definition Init_def l : vProp := MonPred (λ V,
+      ∃ (φR φd: Z → vProp), persisted l (RSLInv φR φd) (λ _ n, InitRaw V n)) _.
+  Next Obligation. unfold InitRaw, persisted. solve_proper. Qed.
+  Definition Init_aux : seal Init_def. by eexists. Qed.
+  Definition Init := unseal Init_aux.
+  Definition Init_eq : Init = _ := seal_eq _.
+
+  Lemma Rel_write l (φ φd: Z → vProp) v (E : coPset)
     (HN: ↑physN ⊆ E) (Hl: ↑persist_locN .@ l ⊆ E):
-    {{{{ ▷ φ v ∗ ▷ φd v ∗ Rel l φ φd }}}}
+    {{{ ▷ φ v ∗ ▷ φd v ∗ Rel l φ φd }}}
       ([ #l ]_at <- #v) @ E
-    {{{{ RET #() ; Rel l φ φd ∗ Init l }}}}.
-
+    {{{ RET #() ; Rel l φ φd ∗ Init l }}}.
   Proof.
     intros; iViewUp. iIntros "kI kS (Itpr & ItprD & Rel) Post".
 
@@ -112,7 +74,7 @@ Local Open Scope I.
     iSplitL "Rel"; [by iExists _|by iExists _,_].
   Qed.
 
-  Lemma Acq_read l (φ φd: Z → vPred) (E: coPset)
+  Lemma Acq_read l (φ φd: Z → vProp) (E: coPset)
     (HN: ↑physN ⊆ E) (Hl: ↑persist_locN .@ l ⊆ E):
     (□ ∀ v, φd v -∗ φd v ∗ φd v)%VP
     ⊩ {{{{ Acq l φ φd ∗ Init l }}}}
@@ -141,7 +103,7 @@ Local Open Scope I.
   Qed.
 
   Lemma RMWAcq_CAS
-    l (φ φd: Z → vPred) (P: vPred) (R: bool → Z → vPred) (v_r v_w: Z) (E: coPset)
+    l (φ φd: Z → vProp) (P: vProp) (R: bool → Z → vProp) (v_r v_w: Z) (E: coPset)
     (HN: ↑physN ⊆ E) (Hl: ↑persist_locN .@ l ⊆ E):
     (□ (∀ v, φd v -∗ φd v ∗ φd v))%VP
     ⊩ {{{{ (P ∗ ▷ φ v_r ∗ ▷ φd v_r
@@ -159,7 +121,7 @@ Local Open Scope I.
     iDestruct (persistor_join_l with "[$RMW $Init]") as "AI".
     iMod (persistor_open with "[$AI]") as (jγ) "(Inv & (AR & IR) & PClose)";
       [done|].
-    iPoseProof (RMWAcqRaw_CAS l φ φd P R v_r v_w _ jγ (E ∖ ↑persist_locN .@ l) 
+    iPoseProof (RMWAcqRaw_CAS l φ φd P R v_r v_w _ jγ (E ∖ ↑persist_locN .@ l)
       with "[]")
       as "RMW"; [solve_ndisj|..].
     { iIntros "!#". iIntros (V v) "?". by iApply ("Dup" $! V v with "[]"). }
@@ -174,7 +136,7 @@ Local Open Scope I.
      by iExists _,_.
   Qed.
 
-  Lemma Acq_alloc (φ φd: Z → vPred) (E: coPset)
+  Lemma Acq_alloc (φ φd: Z → vProp) (E: coPset)
     (HN: ↑physN ⊆ E) (Hp: ↑persistN ⊆ E):
     {{{{ ☐ PersistorCtx }}}}
       Alloc @ E
@@ -192,7 +154,7 @@ Local Open Scope I.
     iSplitL "Rel"; by iExists _.
   Qed.
 
-  Lemma RMWAcq_alloc (φ φd: Z → vPred)(E: coPset)
+  Lemma RMWAcq_alloc (φ φd: Z → vProp)(E: coPset)
     (HN: ↑physN ⊆ E) (Hp: ↑persistN ⊆ E):
     {{{{ ☐ PersistorCtx  }}}}
       Alloc @ E
@@ -210,7 +172,7 @@ Local Open Scope I.
     by iExists _.
   Qed.
 
-  Lemma RMWAcq_init l v (φ φd: Z → vPred) (E : coPset)
+  Lemma RMWAcq_init l v (φ φd: Z → vProp) (E : coPset)
     (HN: ↑physN ⊆ E) (Hp: ↑persistN ⊆ E):
     {{{{ ☐ PersistorCtx ∗ own_loc l ∗ ▷ φ v ∗ ▷ φd v }}}}
       ([ #l ]_na <- #v) @ E
@@ -232,7 +194,7 @@ Local Open Scope I.
     iSplitL "Rel"; [by iExists _|by iExists _,_].
   Qed.
 
-  Lemma Acq_init l v (φ φd: Z → vPred) (E : coPset)
+  Lemma Acq_init l v (φ φd: Z → vProp) (E : coPset)
     (HN: ↑physN ⊆ E) (Hp: ↑persistN ⊆ E):
     {{{{ ☐ PersistorCtx ∗ own_loc l ∗ ▷ φ v ∗ ▷ φd v }}}}
       ([ #l ]_na <- #v) @ E
@@ -254,7 +216,7 @@ Local Open Scope I.
     iSplitL "Rel"; last iSplitL "Acq"; by [iExists _|iExists _|iExists _,_].
   Qed.
 
-  Lemma Acq_join l (φ1 φ2 φd: Z → vPred) (E : coPset)
+  Lemma Acq_join l (φ1 φ2 φd: Z → vProp) (E : coPset)
       (HNl : ↑persist_locN .@ l ⊆ E):
       (Acq l φ1 φd ∗ Acq l φ2 φd ={E}=> Acq l (λ v, φ1 v ∗ φ2 v) φd)%VP.
   Proof.
@@ -271,7 +233,7 @@ Local Open Scope I.
   Qed.
 
 
-  Lemma Acq_split l (φ1 φ2 φd: Z → vPred) (E : coPset)
+  Lemma Acq_split l (φ1 φ2 φd: Z → vProp) (E : coPset)
       (HNl : ↑persist_locN .@ l ⊆ E):
     (Acq l (λ v, φ1 v ∗ φ2 v) φd ={E}=> Acq l φ1 φd ∗ Acq l φ2 φd)%VP.
   Proof.
@@ -286,7 +248,7 @@ Local Open Scope I.
     iSplitL "AR1"; by iExists _.
   Qed.
 
-  Lemma Rel_split l (φ1 φ2 φd: Z → vPred):
+  Lemma Rel_split l (φ1 φ2 φd: Z → vProp):
       Rel l (λ v, φ1 v ∨ φ2 v)%VP φd ⊧  Rel l φ1 φd ∗ Rel l φ2 φd.
   Proof.
     constructor => ? V _.
@@ -296,7 +258,7 @@ Local Open Scope I.
     iSplitL "RR1"; iExists φR, n; by iFrame.
   Qed.
 
-  Lemma Rel_join l (φ1 φ2 φd: Z → vPred):
+  Lemma Rel_join l (φ1 φ2 φd: Z → vProp):
     Rel l φ1 φd ∗ Rel l φ2 φd ⊧ Rel l (λ v, φ1 v ∨ φ2 v)%VP φd.
   Proof.
     constructor => ? V _.
@@ -308,7 +270,7 @@ Local Open Scope I.
     iExists φR, n. by iFrame.
   Qed.
 
-  Lemma RMWAcq_split l (φ φd: Z → vPred) (E : coPset)
+  Lemma RMWAcq_split l (φ φd: Z → vProp) (E : coPset)
     (HNl: ↑persist_locN .@ l ⊆ E):
     (RMWAcq l φ φd ={E}=>  RMWAcq l φ φd ∗ Acq l (λ _, True) φd)%VP.
   Proof.
@@ -327,23 +289,23 @@ Local Open Scope I.
   Section Frac.
 
   Context `{!cinvG Σ}.
-  Definition Relp' l q (φ φd: Z → vPred) V
-    := ∃ (φR: Z → vPred), fractored l (RSLInv φR φd) (λ _ _ n, RelRaw V φ φR n) q.
+  Definition Relp' l q (φ φd: Z → vProp) V
+    := ∃ (φR: Z → vProp), fractored l (RSLInv φR φd) (λ _ _ n, RelRaw V φ φR n) q.
 
-  Definition Acqp' l q (φ φd: Z → vPred) V
-    := ∃ (φR: Z → vPred), fractored l (RSLInv φR φd) (λ _ _ n, AcqRaw V φ n) q.
+  Definition Acqp' l q (φ φd: Z → vProp) V
+    := ∃ (φR: Z → vProp), fractored l (RSLInv φR φd) (λ _ _ n, AcqRaw V φ n) q.
 
   Definition Initp' l q V
-    := ∃ (φR φd: Z → vPred), fractored l (RSLInv φR φd) (λ _ _ n, InitRaw V n) q.
+    := ∃ (φR φd: Z → vProp), fractored l (RSLInv φR φd) (λ _ _ n, InitRaw V n) q.
 
   Definition Emptyp l q V
-    := ∃ (φR φd: Z → vPred), fractored l (RSLInv φR φd) (λ _ _ _, True) q.
+    := ∃ (φR φd: Z → vProp), fractored l (RSLInv φR φd) (λ _ _ _, True) q.
 
-  Definition RMWAcqp' l q (φR φd: Z → vPred) V
+  Definition RMWAcqp' l q (φR φd: Z → vProp) V
     := fractored l (RSLInv φR φd) (λ _ _ n, RMWAcqRaw V n) q.
 
-  Definition vPred_Relp'_mono l q (φ φd: Z → vPred):
-    vPred_Mono (Relp' l q φ φd).
+  Definition vProp_Relp'_mono l q (φ φd: Z → vProp):
+    vProp_Mono (Relp' l q φ φd).
   Proof.
     repeat intro. iIntros "Rel".
     iDestruct "Rel" as (φR) "FR". iDestruct "FR" as (jγ) "(FV & ?)".
@@ -353,12 +315,12 @@ Local Open Scope I.
     iPureIntro. by apply (alloc_local_Mono _ _ V).
   Qed.
 
-  Canonical Structure Relp l q (φ φd: Z → vPred)
-    := Build_vPred _ (vPred_Relp'_mono l q φ φd).
+  Canonical Structure Relp l q (φ φd: Z → vProp)
+    := Build_vProp _ (vProp_Relp'_mono l q φ φd).
 
 
-  Definition vPred_Acqp'_mono l q (φ φd: Z → vPred):
-    vPred_Mono (Acqp' l q φ φd).
+  Definition vProp_Acqp'_mono l q (φ φd: Z → vProp):
+    vProp_Mono (Acqp' l q φ φd).
   Proof.
     repeat intro. iIntros "Acq".
     iDestruct "Acq" as (φR) "FA". iDestruct "FA" as (jγ) "(FV & ?)".
@@ -368,11 +330,11 @@ Local Open Scope I.
     iPureIntro. by apply (alloc_local_Mono _ _ V).
   Qed.
 
-  Canonical Structure Acqp l q (φ φd: Z → vPred)
-    := Build_vPred _ (vPred_Acqp'_mono l q φ φd).
+  Canonical Structure Acqp l q (φ φd: Z → vProp)
+    := Build_vProp _ (vProp_Acqp'_mono l q φ φd).
 
-  Definition vPred_Initp'_mono l q:
-    vPred_Mono (Initp' l q).
+  Definition vProp_Initp'_mono l q:
+    vProp_Mono (Initp' l q).
   Proof.
     repeat intro. iIntros "Init".
     iDestruct "Init" as (φR φd) "FA". iDestruct "FA" as (jγ) "(FV & ?)".
@@ -384,10 +346,10 @@ Local Open Scope I.
   Qed.
 
   Canonical Structure Initp l q
-    := Build_vPred _ (vPred_Initp'_mono l q).
+    := Build_vProp _ (vProp_Initp'_mono l q).
 
-  Definition vPred_RMWAcqp'_mono l q (φ φd: Z → vPred):
-    vPred_Mono (RMWAcqp' l q φ φd).
+  Definition vProp_RMWAcqp'_mono l q (φ φd: Z → vProp):
+    vProp_Mono (RMWAcqp' l q φ φd).
   Proof.
     repeat intro. iIntros "FA".
     iDestruct "FA" as (jγ) "(FV & ?)". iExists jγ. iFrame.
@@ -397,11 +359,11 @@ Local Open Scope I.
     iPureIntro. by apply (alloc_local_Mono _ _ V).
   Qed.
 
-  Canonical Structure RMWAcqp l q (φ φd: Z → vPred)
-    := Build_vPred _ (vPred_RMWAcqp'_mono l q φ φd).
+  Canonical Structure RMWAcqp l q (φ φd: Z → vProp)
+    := Build_vProp _ (vProp_RMWAcqp'_mono l q φ φd).
 
 
-  Lemma Relp_write l q1 q2 (φ φd: Z → vPred) v (E: coPset) 
+  Lemma Relp_write l q1 q2 (φ φd: Z → vProp) v (E: coPset)
     (HN: ↑physN ⊆ E) (Hl: ↑fracN .@ l ⊆ E):
     {{{{ ▷ φ v ∗ ▷ φd v ∗ Relp l (q1 + q2) φ φd }}}}
       ([(#l)]_at <- #v) @ E
@@ -416,14 +378,14 @@ Local Open Scope I.
     iNext. iIntros (V') "(VV' & kS' & Inv & Rel' & Init)".
     iMod ("PClose" $! (λ l _ X, RelRaw V' φ φR X ∗ InitRaw V' X)
               with "[$Inv $Rel' $Init]") as "RI".
-    iDestruct (fractor_splitjoin l _ 
+    iDestruct (fractor_splitjoin l _
                 (λ l _ X, RelRaw V' φ φR X)
                 (λ l _ X, InitRaw V' X) with "RI") as "[Rel' Init]".
     iApply ("Post" $! V' with "VV' kS'").
     iSplitL "Rel'"; [by iExists _|by iExists _,_].
   Qed.
 
-  Lemma Acqp_read l (φ φd: Z → vPred) q1 q2 (E: coPset) 
+  Lemma Acqp_read l (φ φd: Z → vProp) q1 q2 (E: coPset)
     (HN: ↑physN ⊆ E) (Hl: ↑fracN .@ l ⊆ E):
     (□ (∀ v, φd v -∗ φd v ∗ φd v))%VP
     ⊩ {{{{ Acqp l q1 φ φd ∗ Initp l q2 }}}}
@@ -453,7 +415,7 @@ Local Open Scope I.
   Qed.
 
   Lemma RMWAcqp_CAS
-    l (φ φd: Z → vPred) (P: vPred) (R: bool → Z → vPred) (v_r v_w: Z) q1 q2
+    l (φ φd: Z → vProp) (P: vProp) (R: bool → Z → vProp) (v_r v_w: Z) q1 q2
       (E: coPset) (HN: ↑physN ⊆ E) (Hl: ↑fracN .@ l ⊆ E):
     (□ (∀ v, φd v -∗ φd v ∗ φd v))%VP
     ⊩ {{{{ (P ∗ ▷ φ v_r ∗ ▷ φd v_r
@@ -488,7 +450,7 @@ Local Open Scope I.
   Qed.
 
 
-  Lemma Acqp_alloc (φ φd: Z → vPred) (E: coPset) (HN: ↑physN ⊆ E):
+  Lemma Acqp_alloc (φ φd: Z → vProp) (E: coPset) (HN: ↑physN ⊆ E):
     {{{{ True }}}}
       Alloc @ E
     {{{{ (l: loc), RET (LitV $ LitLoc l);
@@ -509,7 +471,7 @@ Local Open Scope I.
     iSplitL "Rel"; by iExists _.
   Qed.
 
-  Lemma RMWAcqp_alloc (φ φd: Z → vPred) (E: coPset) (HN: ↑physN ⊆ E):
+  Lemma RMWAcqp_alloc (φ φd: Z → vProp) (E: coPset) (HN: ↑physN ⊆ E):
     {{{{ True }}}}
       Alloc @ E
     {{{{ (l: loc), RET (LitV $ LitLoc l);
@@ -529,7 +491,7 @@ Local Open Scope I.
     by iExists _.
   Qed.
 
-  Lemma RMWAcqp_init l v (φ φd: Z → vPred) (E: coPset) (HN: ↑physN ⊆ E):
+  Lemma RMWAcqp_init l v (φ φd: Z → vProp) (E: coPset) (HN: ↑physN ⊆ E):
     {{{{ own_loc l ∗ ▷ φ v ∗ ▷ φd v }}}}
       ([ #l ]_na <- #v) @ E
     {{{{ RET #();
@@ -555,7 +517,7 @@ Local Open Scope I.
     iSplitL "Rel"; [by iExists _|by iExists _,_].
   Qed.
 
-  Lemma Acqp_init l v (φ φd: Z → vPred) (E: coPset) (HN: ↑physN ⊆ E):
+  Lemma Acqp_init l v (φ φd: Z → vProp) (E: coPset) (HN: ↑physN ⊆ E):
     {{{{ own_loc l ∗ ▷ φ v ∗ ▷ φd v }}}}
       ([ #l ]_na <- #v) @ E
     {{{{ RET #();
@@ -581,9 +543,9 @@ Local Open Scope I.
     iSplitL "Rel"; last iSplitL "Acq"; by [iExists _|iExists _|iExists _,_].
   Qed.
 
-  Lemma Acqp_join l (φ1 φ2 φd: Z → vPred) q1 q2 (E : coPset)
+  Lemma Acqp_join l (φ1 φ2 φd: Z → vProp) q1 q2 (E : coPset)
       (HNl : ↑fracN .@ l ⊆ E):
-      (Acqp l q1 φ1 φd ∗ Acqp l q2 φ2 φd 
+      (Acqp l q1 φ1 φd ∗ Acqp l q2 φ2 φd
           ={E}=> Acqp l (q1+q2) (λ v, φ1 v ∗ φ2 v) φd)%VP.
   Proof.
     constructor => ? V _.
@@ -598,7 +560,7 @@ Local Open Scope I.
     by iExists _.
   Qed.
 
-  Lemma Acqp_split l (φ1 φ2 φd: Z → vPred) q1 q2 (E : coPset)
+  Lemma Acqp_split l (φ1 φ2 φd: Z → vProp) q1 q2 (E : coPset)
       (HNl : ↑fracN .@ l ⊆ E):
       (Acqp l (q1+q2) (λ v, φ1 v ∗ φ2 v)%VP φd
         ={E}=> Acqp l q1 φ1 φd ∗ Acqp l q2 φ2 φd)%VP.
@@ -616,7 +578,7 @@ Local Open Scope I.
     iSplitL "AR1"; by iExists _.
   Qed.
 
-  Lemma Relp_split l (φ1 φ2 φd: Z → vPred) q1 q2:
+  Lemma Relp_split l (φ1 φ2 φd: Z → vProp) q1 q2:
     Relp l (q1+q2) (λ v, φ1 v ∨ φ2 v)%VP φd ⊧  Relp l q1 φ1 φd ∗ Relp l q2 φ2 φd.
   Proof.
     constructor => ? V _.
@@ -631,7 +593,7 @@ Local Open Scope I.
     iSplitL "RR1"; by iExists _.
   Qed.
 
-  Lemma Relp_join l (φ1 φ2 φd: Z → vPred) q1 q2:
+  Lemma Relp_join l (φ1 φ2 φd: Z → vProp) q1 q2:
     Relp l q1 φ1 φd ∗ Relp l q2 φ2 φd ⊧ Relp l (q1+q2) (λ v, φ1 v ∨ φ2 v)%VP φd.
   Proof.
     constructor => ? V _. iIntros "(RR1 & RR2)".
@@ -644,7 +606,7 @@ Local Open Scope I.
     by iExists _.
   Qed.
 
-  Lemma RMWAcqp_split l (φ φd: Z → vPred) q1 q2 (E : coPset)
+  Lemma RMWAcqp_split l (φ φd: Z → vProp) q1 q2 (E : coPset)
     (HNl: ↑fracN .@ l ⊆ E):
     (RMWAcqp l (q1+q2) φ φd
       ={E}=> RMWAcqp l q1 φ φd ∗ Acqp l q2 (λ _, True) φd)%VP.
diff --git a/theories/rsl_sts.v b/theories/rsl_sts.v
index 68f0ba4a..1ea51775 100644
--- a/theories/rsl_sts.v
+++ b/theories/rsl_sts.v
@@ -7,7 +7,7 @@ From igps.base Require Export ghosts.
 
 Implicit Types (I : gset gname) (h: History) (i: gname).
 
-Inductive state := 
+Inductive state :=
     aUninit I h
   | aInit I1 I2 h (Sub: I2 ⊆ I1)
   | cUninit I h
@@ -65,12 +65,12 @@ Lemma rISet2_subseteq s:
   rISet2 s ⊆ rISet s.
 Proof. by destruct s. Qed.
 
-Definition tok s: set token 
+Definition tok s: set token
   := {[ t | ∃ i, t = Change i ∧ i ∉ rISet2 s ]}.
 
 Global Arguments tok !_ /.
 
-Canonical Structure rsl_sts := sts.STS prim_step tok.
+Canonical Structure rsl_sts := sts.Sts prim_step tok.
 
 
 Definition i_states i : set state := {[ s | i ∈ rISet2 s ]}.
@@ -99,7 +99,7 @@ Proof.
   by apply: frame_step_step.
 Qed.
 
-Lemma frame_steps_RMW_branch s s' T: 
+Lemma frame_steps_RMW_branch s s' T:
   sts.frame_steps T s s' →
   ∀ I0 h, (s = cUninit I0 h ∨ s = cInit I0 h)
     → ∃ I' h', s' = cUninit I' h' ∨ s' = cInit I' h'.
@@ -114,10 +114,10 @@ Proof.
     + apply (IHrtc I0 h'). by right.
 Qed.
 
-Lemma frame_steps_Init_branch s s' T: 
+Lemma frame_steps_Init_branch s s' T:
   sts.frame_steps T s s' →
    (∃ h, ((∃ I1 I2 (Sub: I2 ⊆ I1), s = aInit I1 I2 h Sub) ∨ (∃ I1, s = cInit I1 h)))
-    → ∃ h', 
+    → ∃ h',
       (∃ I1' I2' (Sub' : I2' ⊆ I1'), (s' = aInit I1' I2' h' Sub'))
          ∨ (∃ I1', s' = cInit I1' h').
 Proof.
@@ -139,7 +139,7 @@ Proof.
   move => Step1 Step2 Eq1 Eq2.
   destruct Eq1 as [I1 [h1 Eq]].
   destruct (frame_steps_RMW_branch _ _ _ Step1 _ _ Eq) as [I'1 [h'1 H1]].
-  destruct (frame_steps_Init_branch _ _ _ Step2 Eq2) 
+  destruct (frame_steps_Init_branch _ _ _ Step2 Eq2)
     as [h'2 [[I1' [I2' [Sub' H2]]]|[I1' H2]]].
   - rewrite H2 in H1. destruct H1 as [H1|H1]; by inversion H1.
   - by do 2 eexists.
@@ -223,7 +223,7 @@ Definition state_ISet_join i1 i2 i s :=
   match s with
   | aUninit I0 h => aUninit (ISet_join i1 i2 i I0) h
   | aInit I1 I2 h Sub => aInit (ISet_join i1 i2 i I1)
-                                (ISet_join i1 i2 i I2) h 
+                                (ISet_join i1 i2 i I2) h
                                 (ISet_join_subseteq i1 i2 i I1 I2 Sub)
   | cUninit I0 h => cUninit (ISet_join i1 i2 i I0) h
   | cInit I0 h => cInit (ISet_join i1 i2 i I0) h
@@ -234,7 +234,7 @@ Definition state_ISet_split i1 i2 i s :=
   match s with
   | aUninit I0 h => aUninit (ISet_split i1 i2 i I0) h
   | aInit I1 I2 h Sub => aInit (ISet_split i1 i2 i I1)
-                                (ISet_split i1 i2 i I2) h 
+                                (ISet_split i1 i2 i I2) h
                                 (ISet_split_subseteq i1 i2 i I1 I2 Sub)
   | cUninit I0 h => cUninit (ISet_split i1 i2 i I0) h
   | cInit I0 h => cInit (ISet_split i1 i2 i I0) h
@@ -245,7 +245,7 @@ Definition state_ISet_switch i i' s :=
   match s with
   | aUninit I0 h => aUninit (ISet_switch i i' I0) h
   | aInit I1 I2 h Sub => aInit (ISet_switch i i' I1)
-                                (ISet_switch i i' I2) h 
+                                (ISet_switch i i' I2) h
                                 (ISet_switch_subseteq i i' I1 I2 Sub)
   | cUninit I0 h => cUninit (ISet_switch i i' I0) h
   | cInit I0 h => cInit (ISet_switch i i' I0) h
@@ -255,7 +255,7 @@ Definition state_ISet_insert i s :=
   match s with
   | aUninit I0 h => aUninit ({[i]} ∪ I0) h
   | aInit I1 I2 h Sub => aInit ({[i]} ∪ I1)
-                               ({[i]} ∪ I2) h 
+                               ({[i]} ∪ I2) h
                                (union_mono_l {[i]} I2 I1 Sub)
   | cUninit I0 h => cUninit ({[i]} ∪ I0) h
   | cInit I0 h => cInit ({[i]} ∪ I0) h
diff --git a/theories/tactics.v b/theories/tactics.v
index 39c85323..08ee554a 100644
--- a/theories/tactics.v
+++ b/theories/tactics.v
@@ -220,7 +220,9 @@ Definition atomic0 (e : expr0) :=
   end.
 Definition atomic (e : expr) := atomic0 (e.1).
 
-Lemma atomic_correct e π : atomic0 e  → Atomic ((to_expr' e, π) : language.expr ra_lang).
+Lemma atomic_correct e π :
+  atomic0 e  →
+  Atomic WeaklyAtomic ((to_expr' e, π) : language.expr ra_lang).
 Proof.
   (* intros He. *)
   (* - inversion H; clear H. *)
@@ -315,17 +317,18 @@ Hint Extern 10 (AsVal _) => solve_as_val : typeclass_instances.
 
 Ltac solve_atomic :=
   match goal with
-  | |- language.Atomic (base.Fork ?e0, ?Ï€) =>
+  | |- language.Atomic _ (base.Fork ?e0, ?Ï€) =>
     rewrite -(W.to_expr_eq e0);
-    change (@language.Atomic ra_lang (W.to_expr' (W.Fork (W.to_expr0 e0)), π));
+    change (@language.Atomic ra_lang WeaklyAtomic (W.to_expr' (W.Fork (W.to_expr0 e0)), π));
+    apply W.atomic_correct; vm_compute; exact I
+  | |- language.Atomic _ (?e, ?Ï€) =>
+    simpl e;
+    let e' := W.of_expr0 e in
+    change (@language.Atomic ra_lang WeaklyAtomic (W.to_expr' e', π));
     apply W.atomic_correct; vm_compute; exact I
-  | |- language.Atomic (?e, ?Ï€) =>
-    simpl e; 
-     let e' := W.of_expr0 e in change (@language.Atomic ra_lang (W.to_expr' e', π));
-     apply W.atomic_correct; vm_compute; exact I
   end.
 (* For the side-condition of elim_vs_pvs_wp_atomic *)
-Hint Extern 10 (language.Atomic _) => solve_atomic : typeclass_instances.
+Hint Extern 10 (language.Atomic _ _) => solve_atomic : typeclass_instances.
 
 (** Substitution *)
 Ltac simpl_subst :=
diff --git a/theories/tactics_vProp.v b/theories/tactics_vProp.v
index 9e5742a6..00f35889 100644
--- a/theories/tactics_vProp.v
+++ b/theories/tactics_vProp.v
@@ -6,9 +6,9 @@ From igps Require tactics.
 Import base.
 
 
-Lemma atomic_correct e V : tactics.W.atomic0 e → Atomic (Λ:=ra_lang) (tactics.W.to_expr' e,V).
+Lemma atomic_correct e V : tactics.W.atomic0 e →
+                           Atomic (Λ:=ra_lang) WeaklyAtomic (tactics.W.to_expr' e,V).
 Proof.
-  
 (*   intros He. apply ectx_language_atomic; last first. *)
 (*   - apply ectxi_language_sub_redexes_are_values => /= Ki [e' V'] Hfill. *)
 (*     (destruct e=> //); destruct Ki; repeat (simplify_eq; case_match=>//); *)
@@ -30,8 +30,8 @@ Admitted.
 Ltac solve_closed :=
   match goal with
   | |- Closed ?X ?e =>
-     let e' := tactics.W.of_expr0 e in change (Closed X (tactics.W.to_expr0 e'));
-     apply tactics.W.is_closed_correct; vm_compute; exact I
+     let e' := tactics.W.of_expr0 e in change (Closed X (tactics.W.to_expr' e'));
+     apply tactics.W.is_closed_correct0; vm_compute; exact I
   end.
 Hint Extern 0 (Closed _ _) => solve_closed : typeclass_instances.
 
diff --git a/theories/tests/message_passing.v b/theories/tests/message_passing.v
index 86225c2a..4192dc9f 100644
--- a/theories/tests/message_passing.v
+++ b/theories/tests/message_passing.v
@@ -37,4 +37,4 @@ Qed.
 
 Print Assumptions message_passing.
 
-End MP.
\ No newline at end of file
+End MP.
diff --git a/theories/types.v b/theories/types.v
index 70812a1f..1545c506 100644
--- a/theories/types.v
+++ b/theories/types.v
@@ -15,7 +15,7 @@ Section Definitions.
   Inductive isval : ∀ (v : Val), Prop := val_is_val v : isval (VInj v).
 
   Lemma NEqAD : A ≠ D. Proof. by inversion 1. Qed.
-  
+
   Definition View := gmap positive positive.
   Record message :=
     mkMessage
@@ -83,9 +83,9 @@ Section Definitions.
     Lemma loc_acc_of evt acc : acc_of evt = Some acc -> ∃ x, loc_of evt = Some x.
     Proof. by destruct evt; eauto. Qed.
 
-    Definition acc_to_sum : access -> _ 
+    Definition acc_to_sum : access -> _
       := λ acc, match acc with | na => inl () | at_hack => inr () end.
-    Definition sum_to_acc : _ -> access 
+    Definition sum_to_acc : _ -> access
       := λ s, match s with | inl () => na | inr () => at_hack end.
     Instance acc_to_sum_inj : Inj (=) (=) acc_to_sum.
     Proof. intros [] []; naive_solver. Qed.
@@ -99,7 +99,7 @@ Section Definitions.
 
 End Definitions.
 
-Canonical Structure View_bi_index := BIIndex View (gmap_SqSubsetEq) _.
+Canonical Structure View_bi_index := BiIndex View _ (gmap_SqSubsetEq) _.
 
 Notation "'<' x → v @ t , R >" :=
   (mkMessage x v t R)
diff --git a/theories/weakestpre.v b/theories/weakestpre.v
index c84f5c8c..8ec7e531 100644
--- a/theories/weakestpre.v
+++ b/theories/weakestpre.v
@@ -11,7 +11,7 @@ Module WP_Def.
     MonPred (
         monPred_upclosed
           (λ V, ∀ π, Seen π V -∗
-                          weakestpre.wp E (e, V) (λ '(v,V), Seen π V ∗ monPred_car (Φ v) V))
+              weakestpre.wp NotStuck E (e, V) (λ '(v,V), Seen π V ∗ monPred_car (Φ v) V))
       )%I _.
   Definition wp_aux : seal (@wp_def). by eexists. Qed.
   Definition wp := unseal (wp_aux).
@@ -115,7 +115,7 @@ Proof.
   apply bi.forall_intro => ?. apply bi.wand_intro_l.
   (* apply bi.forall_intro => ?. apply bi.forall_intro => Eq. *)
   rewrite monPred_mono //.
-  by refine (weakestpre.wp_value' (Λ:=ra_lang) _ _ (v,V)).
+  by refine (weakestpre.wp_value' (Λ:=ra_lang) _ _ _ (v,V)).
   (* by iIntros "($&?&$)". (* TODO: don't use proofmode. also figure out why not *) *)
 Qed.
 (* TODO: we don't seem to need this but it should be fixed and proven *)
@@ -130,7 +130,9 @@ Qed.
 Import language.
 
 Lemma wp_strong_mono_iris E1 E2 e φ ψ V0:
-  E1 ⊆ E2 → (∀ v V (_ : V0 ⊑ V), φ (v,V) ={E2}=∗ ψ (v,V)) ∗ weakestpre.wp E1 (e,V0) (φ) ⊢ weakestpre.wp E2 (e,V0) ψ.
+  E1 ⊆ E2 → (∀ v V (_ : V0 ⊑ V), φ (v,V) ={E2}=∗ ψ (v,V)) ∗
+            weakestpre.wp weakestpre.NotStuck E1 (e,V0) (φ)
+          ⊢ weakestpre.wp weakestpre.NotStuck E2 (e,V0) ψ.
 Proof.
   iIntros (?) "[HΦ H]". iLöb as "IH" forall (e V0).
   setoid_rewrite weakestpre.wp_unfold at 3 4.
@@ -160,7 +162,9 @@ Proof.
   iIntros (Vo) "[VS WP]". iIntros (V HV). iSpecialize ("WP" $! V with "[//]").
   iIntros (π) "S". iSpecialize ("WP" $! π with "S").
   iApply (wp_strong_mono_iris with "[-$WP]"); first done.
-  iIntros (???). iIntros "[$ HΦ]". iApply ("VS" $! _ _ with "[%] HΦ"). by etrans.
+  iIntros (? V' HV'). iIntros "[$ HΦ]".
+  setoid_rewrite HV. setoid_rewrite HV'.
+  iApply ("VS" with "HΦ").
 Qed.
 
 Lemma fupd_wp E e Φ : (|={E}=> WP e @ E {{ Φ }}) ⊢ WP e @ E {{ Φ }}.
@@ -180,13 +184,13 @@ Lemma wp_fupd E e Φ : WP e @ E {{ v, |={E}=> Φ v }} ⊢ WP e @ E {{ Φ }}.
 Proof. iIntros "H". iApply (wp_strong_mono E); try iFrame; auto. Qed.
 
 Lemma wp_atomic E1 E2 e Φ :
-  (∀ V, Atomic ((e ,V) : ra_lang.expr)) →
+  (∀ V, Atomic WeaklyAtomic ((e ,V) : ra_lang.expr)) →
   (|={E1,E2}=> WP e @ E2 {{ v, |={E2,E1}=> Φ v }}) ⊢ WP e @ E1 {{ Φ }}.
 Proof.
   intros Hatomic. iStartProof (uPred _). iIntros (?) "H".
   rewrite wp_eq /wp_def /=. iIntros (?? π) "S".
-  iApply (weakestpre.wp_atomic E1 E2) => //. iMod "H". iModIntro.
-  iApply (weakestpre.wp_strong_mono E2); first done.
+  iApply (weakestpre.wp_atomic _ E1 E2) => //. iMod "H". iModIntro.
+  iApply (weakestpre.wp_strong_mono _ E2); first done.
   iDestruct ("H" with "[%] S") as "$"; first done.
   iIntros (?) "H". case_match. iModIntro.
   by iDestruct "H" as "[$ H]".
@@ -210,17 +214,16 @@ Proof.
   iApply (wp_strong_mono_iris E2 with "[-$H]"); first done.
   move/prim_step_view_mono : Hstep => ?.
   iIntros (???) "[$ HP]".
-  iApply ("HP" $! _); first done. iApply (monPred_mono with "HR").
-  repeat etrans=>//.
+  iApply ("HP" with "[HR]"). iApply ("HR" $! _ with "[%]"). repeat etrans=>//.
 Qed.
 
 
 (* TODO: file issue with iris bc of LanguageCtx stuff *)
 (* TODO: figure out was is meant by above comment *)
 Lemma wp_bind_iris K E e φ V0:
-  weakestpre.wp E (e,V0) (λ '(v,Vp),
-                          weakestpre.wp E (base.fill K (base.of_val v),Vp) φ)
-                ⊢ weakestpre.wp E (base.fill K e,V0) φ.
+  weakestpre.wp weakestpre.NotStuck E (e,V0) (λ '(v,Vp),
+      weakestpre.wp weakestpre.NotStuck E (base.fill K (base.of_val v),Vp) φ)
+                ⊢ weakestpre.wp weakestpre.NotStuck E (base.fill K e,V0) φ.
 Proof.
   iIntros "H".
   rewrite (_ : (base.fill _ _, _) = ectx_language.fill K (e,V0)); last first.
@@ -313,6 +316,53 @@ Section proofmode_classes.
   Implicit Types P Q : vProp Σ.
   Implicit Types Φ : base.val → vProp Σ.
 
+  Import iprop.
+
+  Global Instance into_forall_wp E e Φ V :
+    V ⊑ V' →
+    IntoForall
+      (monPred_car (WP e @ E {{ Φ }})%I V)
+      (λ π, Seen π V' -∗
+                 weakestpre.wp weakestpre.NotStuck E (e, V') (λ '(v,V), Seen π V ∗ monPred_car (Φ v) V))%I.
+  Proof.
+    intros.
+    rewrite wp_eq /wp_def /= /IntoForall
+    bi.forall_elim
+    bi.pure_impl_forall
+    bi.forall_elim; try done. 
+  Qed.
+
+  Global Instance from_forall_wp E e Φ V :
+    FromForall
+      (monPred_car (WP e @ E {{ Φ }})%I V)
+      (λ V', ⌜V ⊑ V'⌝ → ∀ π, Seen π V' -∗
+                 weakestpre.wp weakestpre.NotStuck E (e, V') (λ '(v,V), Seen π V ∗ monPred_car (Φ v) V))%I.
+  Proof.
+    rewrite wp_eq /wp_def /= /FromForall.
+    apply bi.forall_intro => V0.
+    rewrite bi.pure_impl_forall.
+    apply bi.forall_intro => HV'.
+    apply bi.forall_intro => π.
+    by rewrite (bi.forall_elim V0)
+    bi.pure_impl_forall
+    (bi.forall_elim HV')
+    (bi.forall_elim π).
+  Qed.
+
+  Global Instance into_wand_wp p q E e Φ V π:
+    IntoWand p q
+      (monPred_car (WP e @ E {{ Φ }})%I V)
+      (Seen π V)
+            (weakestpre.wp weakestpre.NotStuck E (e, V) (λ '(v,V), Seen π V ∗ monPred_car (Φ v) V))%I.
+  Proof.
+    rewrite wp_eq /wp_def /= /IntoWand
+    bi.forall_elim
+    bi.pure_impl_forall
+    bi.forall_elim; [|done].
+    rewrite bi.forall_elim.
+    apply (_ : IntoWand _ _ _ _ _).
+  Qed.
+
   Global Instance frame_wp p E e R Φ Ψ :
     (∀ v, Frame p R (Φ v) (Ψ v)) → Frame p R (WP e @ E {{ Φ }}) (WP e @ E {{ Ψ }}).
   Proof. rewrite /Frame=> HR. rewrite wp_frame_l. apply wp_mono, HR. Qed.
@@ -330,13 +380,60 @@ Section proofmode_classes.
 
   (* lower precedence, if possible, it should always pick elim_upd_fupd_wp *)
   Global Instance elim_modal_fupd_wp_atomic E1 E2 e P Φ :
-    (∀ V, Atomic ((e,V) : ra_lang.expr)) →
+    (∀ V, Atomic WeaklyAtomic ((e,V) : ra_lang.expr)) →
     ElimModal (|={E1,E2}=> P) P
             (WP e @ E1 {{ Φ }}) (WP e @ E2 {{ v, |={E2,E1}=> Φ v }})%I | 100.
   Proof. intros. by rewrite /ElimModal fupd_frame_r bi.wand_elim_r wp_atomic. Qed.
 
 End proofmode_classes.
 
+Section proofmode_classes_embedding.
+  Context `{foundationG Σ}.
+  Implicit Types P Q : vProp Σ.
+  Implicit Types Φ : base.val → vProp Σ.
+
+  Global Instance elim_modal_bupd_wp_emb E e A Φ :
+    ElimModal (⎡|==> A⎤) ⎡A⎤ (WP e @ E {{ Φ }}) (WP e @ E {{ Φ }}).
+  Proof.
+    rewrite /ElimModal.
+    iStartProof (upred.uPredSI _); iIntros (?) "[U WP]".
+    iIntros (V -> π) "kS".
+    iMod "U".
+    iSpecialize ("WP" with "[U //]").
+    iApply ("WP" with "kS").
+  Qed.
+
+  Global Instance elim_modal_fupd_wp_emb E e A Φ :
+    ElimModal (⎡|={E}=> A⎤) ⎡A⎤ (WP e @ E {{ Φ }}) (WP e @ E {{ Φ }}).
+  Proof.
+    rewrite /ElimModal.
+    iStartProof (upred.uPredSI _); iIntros (?) "[U WP]".
+    iIntros (V -> π) "kS".
+    iMod "U".
+    iSpecialize ("WP" with "[U //]").
+    iApply ("WP" with "kS").
+  Qed.
+
+  (* lower precedence, if possible, it should always pick elim_upd_fupd_wp *)
+  Global Instance elim_modal_fupd_wp_atomic_emb E1 E2 e A Φ :
+    (∀ V, Atomic WeaklyAtomic ((e,V) : ra_lang.expr)) →
+    ElimModal (⎡|={E1,E2}=> A⎤) ⎡A⎤
+              (WP e @ E1 {{ Φ }}) (WP e @ E2 {{ v, |={E2,E1}=> Φ v }})%I | 100.
+  Proof.
+    intros.
+    rewrite /ElimModal.
+    iStartProof (upred.uPredSI _). iIntros (?) "[U WP]".
+    iIntros (V -> π) "kS".
+    iMod "U".
+    iSpecialize ("WP" with "[U //]").
+    iApply iris.program_logic.weakestpre.wp_mono; cycle 1.
+    iApply ("WP" with "kS").
+    rewrite monPred_fupd_eq /=.
+    by iIntros ([v V']) "[$ >$]".
+  Qed.
+
+End proofmode_classes_embedding.
+
 Hint Extern 0 (atomic _) => assumption : typeclass_instances.
 
-Delimit Scope val_scope with Val. (* TODO: what is this? *)
\ No newline at end of file
+Delimit Scope val_scope with Val. (* TODO: what is this? *)
diff --git a/theories/wp_tactics_vProp.v b/theories/wp_tactics_vProp.v
index 322291e3..e5aca333 100644
--- a/theories/wp_tactics_vProp.v
+++ b/theories/wp_tactics_vProp.v
@@ -95,4 +95,4 @@ Tactic Notation "wp_apply" open_constr(lem) :=
       | Some (_,?P) => fail "wp_apply: cannot apply" P
       end
     | _ => fail "wp_apply: not a 'wp'"
-    end).
\ No newline at end of file
+    end).
-- 
GitLab