From 9ba9ef1b761b690a1eab30e69e435497014c68b7 Mon Sep 17 00:00:00 2001
From: David Swasey <swasey@mpi-sws.org>
Date: Sat, 21 Feb 2015 15:05:06 +0100
Subject: [PATCH] Notation and metavariables.
MIME-Version: 1.0
Content-Type: text/plain; charset=UTF-8
Content-Transfer-Encoding: 8bit

Moved connective notation to BI.

Added ⁺T for ra_pos T and eliminated BI.pres since I'd rather see ⁺res than BI.pres.

Bound mask_scope to type mask.
---
 iris_core.v      | 137 ++++++++++++++++++++++-------------------------
 iris_meta.v      |  16 ++----
 iris_unsafe.v    |   4 +-
 iris_vs.v        |  66 +++++++++++------------
 iris_wp.v        |  22 +++-----
 lib/ModuRes/BI.v |  14 ++++-
 lib/ModuRes/RA.v |  65 +++++++++++-----------
 masks.v          |   1 +
 8 files changed, 156 insertions(+), 169 deletions(-)

diff --git a/iris_core.v b/iris_core.v
index 5c7260882..d2b7a6c17 100644
--- a/iris_core.v
+++ b/iris_core.v
@@ -24,6 +24,8 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
   Module Export WP := WorldProp R.
 
   Delimit Scope iris_scope with iris.
+  Local Open Scope ra_scope.
+  Local Open Scope bi_scope.
   Local Open Scope iris_scope.
 
   (** Instances for a bunch of types (some don't even have Setoids) *)
@@ -61,24 +63,26 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
   
   (** And now we're ready to build the IRIS-specific connectives! *)
 
+  Implicit Types (P Q : Props) (w : Wld) (n i k : nat) (m : mask) (r : pres) (u v : res) (σ : state).
+
   Section Necessitation.
     (** Note: this could be moved to BI, since it's possible to define
-        for any UPred over a monoid. **)
+        for any UPred over a RA. **)
 
     Local Obligation Tactic := intros; resp_set || eauto with typeclass_instances.
 
     Program Definition box : Props -n> Props :=
-      n[(fun p => m[(fun w => mkUPred (fun n r => p w n ra_pos_unit) _)])].
+      n[(fun P => m[(fun w => mkUPred (fun n r => P w n ra_pos_unit) _)])].
     Next Obligation.
       intros n m r s HLe _ Hp; rewrite-> HLe; assumption.
     Qed.
     Next Obligation.
       intros w1 w2 EQw m r HLt; simpl.
-      eapply (met_morph_nonexp _ _ p); eassumption.
+      eapply (met_morph_nonexp _ _ P); eassumption.
     Qed.
     Next Obligation.
       intros w1 w2 Subw n r; simpl.
-      apply p; assumption.
+      apply P; assumption.
     Qed.
     Next Obligation.
       intros p1 p2 EQp w m r HLt; simpl.
@@ -87,6 +91,34 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
 
   End Necessitation.
 
+  Notation "â–¡ P" := (box P) (at level 30, right associativity) : iris_scope.
+
+  (** Lemmas about box **)
+  Lemma box_intro P Q (Hpr : □P ⊑ Q) :
+    □P ⊑ □Q.
+  Proof.
+    intros w n r Hp; simpl; apply Hpr, Hp.
+  Qed.
+
+  Lemma box_elim P :
+    □P ⊑ P.
+  Proof.
+    intros w n r Hp; simpl in Hp.
+    eapply uni_pred, Hp; [reflexivity |].
+    now eapply unit_min.
+  Qed.
+
+  Lemma box_top : ⊤ == □⊤.
+  Proof.
+    intros w n r; simpl; unfold const; reflexivity.
+  Qed.
+
+  Lemma box_disj P Q :
+    □(P ∨ Q) == □P ∨ □Q.
+  Proof.
+    intros w n r; reflexivity.
+  Qed.
+
   (** "Internal" equality **)
   Section IntEq.
     Context {T} `{mT : metric T}.
@@ -124,10 +156,10 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
   Section Invariants.
 
     (** Invariants **)
-    Definition invP (i : nat) (p : Props) (w : Wld) : UPred pres :=
-      intEqP (w i) (Some (ı' p)).
+    Definition invP i P w : UPred pres :=
+      intEqP (w i) (Some (ı' P)).
     Program Definition inv i : Props -n> Props :=
-      n[(fun p => m[(invP i p)])].
+      n[(fun P => m[(invP i P)])].
     Next Obligation.
       intros w1 w2 EQw; unfold invP; simpl morph.
       destruct n; [apply dist_bound |].
@@ -148,34 +180,12 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
 
   End Invariants.
 
-  Notation "â–¡ p" := (box p) (at level 30, right associativity) : iris_scope.
-  Notation "⊤" := (top : Props) : iris_scope.
-  Notation "⊥" := (bot : Props) : iris_scope.
-  Notation "p ∧ q" := (and p q : Props) (at level 40, left associativity) : iris_scope.
-  Notation "p ∨ q" := (or p q : Props) (at level 50, left associativity) : iris_scope.
-  Notation "p * q" := (sc p q : Props) (at level 40, left associativity) : iris_scope.
-  Notation "p → q" := (BI.impl p q : Props) (at level 55, right associativity) : iris_scope.
-  Notation "p '-*' q" := (si p q : Props) (at level 55, right associativity) : iris_scope.
-  Notation "∀ x , p" := (all n[(fun x => p)] : Props) (at level 60, x ident, no associativity) : iris_scope.
-  Notation "∃ x , p" := (xist n[(fun x => p)] : Props) (at level 60, x ident, no associativity) : iris_scope.
-  Notation "∀ x : T , p" := (all n[(fun x : T => p)] : Props) (at level 60, x ident, no associativity) : iris_scope.
-  Notation "∃ x : T , p" := (xist n[(fun x : T => p)] : Props) (at level 60, x ident, no associativity) : iris_scope.
-
-  Lemma valid_iff p :
-    valid p <-> (⊤ ⊑ p).
-  Proof.
-    split; intros Hp.
-    - intros w n r _; apply Hp.
-    - intros w n r; apply Hp; exact I.
-  Qed.
-
   Section Ownership.
-    Local Open Scope ra.
 
     (** Ownership **)
     (* We define this on *any* resource, not just the positive (valid) ones.
        Note that this makes ownR trivially *False* for invalid u: There is no
-       elment v such that u · v = r (where r is valid) *)
+       element v such that u · v = r (where r is valid) *)
     Program Definition ownR: res -=> Props :=
       s[(fun u => pcmconst (mkUPred(fun n r => u ⊑ ra_proj r) _) )].
     Next Obligation.
@@ -212,7 +222,7 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
 
     (** Proper physical state: ownership of the machine state **)
     Program Definition ownS : state -n> Props :=
-      n[(fun s => ownR (ex_own _ s, 1%ra))].
+      n[(fun s => ownR (ex_own _ s, 1))].
     Next Obligation.
       intros r1 r2 EQr; destruct n as [| n]; [apply dist_bound |].
       rewrite EQr. reflexivity.
@@ -220,58 +230,32 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
 
     (** Proper ghost state: ownership of logical **)
     Program Definition ownL : RL.res -n> Props :=
-      n[(fun r => ownR (1%ra, r))].
+      n[(fun r : RL.res => ownR (1, r))].
     Next Obligation.
       intros r1 r2 EQr. destruct n as [| n]; [apply dist_bound |eapply dist_refl].
       simpl in EQr. intros w m t. simpl. change ( (ex_unit state, r1) ⊑ (ra_proj t) <->  (ex_unit state, r2) ⊑ (ra_proj t)). rewrite EQr. reflexivity.
     Qed.
 
     (** Ghost state ownership **)
-    Lemma ownL_sc (u t : RL.res) :
-      ownL (u · t) == ownL u * ownL t.
+    Lemma ownL_sc (r s : RL.res) :
+      ownL (r · s) == ownL r * ownL s.
     Proof.
-      assert (Heq: (ex_unit state, u · t) == ((ex_unit state, u) · (ex_unit state, t)) ) by reflexivity.
+      assert (Heq: (1, r · s) == ((1, r) · (1, s)) ) by reflexivity.
       (* I cannot believe I have to write this... *)
-      change (ownR (ex_unit state, u · t) == ownR (ex_unit state, u) * ownR (ex_unit state, t)).
+      change (ownR (1, r · s) == ownR (1, r) * ownR (1, s)).
       rewrite Heq.
       now eapply ownR_sc.
     Qed.
 
   End Ownership.
 
-  (** Lemmas about box **)
-  Lemma box_intro p q (Hpq : □ p ⊑ q) :
-    □ p ⊑ □ q.
-  Proof.
-    intros w n r Hp; simpl; apply Hpq, Hp.
-  Qed.
-
-  Lemma box_elim p :
-    □ p ⊑ p.
-  Proof.
-    intros w n r Hp; simpl in Hp.
-    eapply uni_pred, Hp; [reflexivity |].
-    now eapply unit_min.
-  Qed.
-
-  Lemma box_top : ⊤ == □ ⊤.
-  Proof.
-    intros w n r; simpl; unfold const; reflexivity.
-  Qed.
-
-  Lemma box_disj p q :
-    □ (p ∨ q) == □ p ∨ □ q.
-  Proof.
-    intros w n r; reflexivity.
-  Qed.
-
   (** Timeless *)
 
-  Definition timelessP (p : Props) w n :=
-    forall w' k r (HSw : w ⊑ w') (HLt : k < n) (Hp : p w' k r), p w' (S k) r.
+  Definition timelessP P w n :=
+    forall w' k r (HSw : w ⊑ w') (HLt : k < n) (Hp : P w' k r), P w' (S k) r.
 
-  Program Definition timeless (p : Props) : Props :=
-    m[(fun w => mkUPred (fun n r => timelessP p w n) _)].
+  Program Definition timeless P : Props :=
+    m[(fun w => mkUPred (fun n r => timelessP P w n) _)].
   Next Obligation.
     intros n1 n2 _ _ HLe _ HT w' k r HSw HLt Hp; eapply HT, Hp; [eassumption |].
     omega.
@@ -280,9 +264,9 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
     intros w1 w2 EQw k; simpl; intros _ HLt; destruct n as [| n]; [now inversion HLt |].
     split; intros HT w' m r HSw HLt' Hp.
     - symmetry in EQw; assert (HD := extend_dist _ _ _ _ EQw HSw); assert (HS := extend_sub _ _ _ _ EQw HSw).
-      apply (met_morph_nonexp _ _ p) in HD; apply HD, HT, HD, Hp; now (assumption || eauto with arith).
+      apply (met_morph_nonexp _ _ P) in HD; apply HD, HT, HD, Hp; now (assumption || eauto with arith).
     - assert (HD := extend_dist _ _ _ _ EQw HSw); assert (HS := extend_sub _ _ _ _ EQw HSw).
-      apply (met_morph_nonexp _ _ p) in HD; apply HD, HT, HD, Hp; now (assumption || eauto with arith).
+      apply (met_morph_nonexp _ _ P) in HD; apply HD, HT, HD, Hp; now (assumption || eauto with arith).
   Qed.
   Next Obligation.
     intros w1 w2 HSw n; simpl; intros _ HT w' m r HSw' HLt Hp.
@@ -290,8 +274,6 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
   Qed.
 
   Section WorldSatisfaction.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
 
     (* First, we need to compose the resources of a finite map. This won't be pretty, for
        now, since the library does not provide enough
@@ -359,7 +341,7 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
 
     Global Instance preo_unit : preoType () := disc_preo ().
 
-    Program Definition wsat (σ : state) (m : mask) (r : res) (w : Wld) : UPred () :=
+    Program Definition wsat σ m (r : res) w : UPred () :=
       â–¹ (mkUPred (fun n _ => exists rs : nat -f> pres,
                     state_sat (r · (comp_map rs)) σ
                       /\ forall i (Hm : m i),
@@ -383,7 +365,7 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
         rewrite fdLookup_in; setoid_rewrite <- EQw; rewrite <- fdLookup_in; reflexivity.
     Qed.
 
-    Global Instance wsat_dist n σ m r : Proper (dist n ==> dist n) (wsat σ m r).
+    Global Instance wsat_dist n σ m u : Proper (dist n ==> dist n) (wsat σ m u).
     Proof.
       intros w1 w2 EQw [| n'] [] HLt; [reflexivity |]; destruct n as [| n]; [now inversion HLt |].
       split; intros [rs [HE HM] ]; exists rs.
@@ -412,6 +394,15 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
 
   End WorldSatisfaction.
 
-  Notation " p @ k " := ((p : UPred ()) k tt) (at level 60, no associativity).
+  Notation " P @ k " := ((P : UPred ()) k tt) (at level 60, no associativity).
+
+
+  Lemma valid_iff P :
+    valid P <-> (⊤ ⊑ P).
+  Proof.
+    split; intros Hp.
+    - intros w n r _; apply Hp.
+    - intros w n r; apply Hp; exact I.
+  Qed.
 
 End IrisCore.
diff --git a/iris_meta.v b/iris_meta.v
index b60ae7bf6..db948ce63 100644
--- a/iris_meta.v
+++ b/iris_meta.v
@@ -7,14 +7,13 @@ Set Bullet Behavior "Strict Subproofs".
 Module IrisMeta (RL : RA_T) (C : CORE_LANG).
   Module Export WP := IrisWP RL C.
 
-  Delimit Scope iris_scope with iris.
+  Local Open Scope lang_scope.
+  Local Open Scope ra_scope.
+  Local Open Scope bi_scope.
   Local Open Scope iris_scope.
 
   Section Adequacy.
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
-    Local Open Scope lang_scope.
+
     Local Open Scope list_scope.
 
     (* weakest-pre for a threadpool *)
@@ -201,11 +200,6 @@ Module IrisMeta (RL : RA_T) (C : CORE_LANG).
 
   Section Lifting.
 
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
-    Local Open Scope lang_scope.
-
     Implicit Types (P : Props) (i : nat) (safe : bool) (m : mask) (e : expr) (Q R : vPred) (r : pres).
 
 
@@ -242,6 +236,6 @@ Module IrisMeta (RL : RA_T) (C : CORE_LANG).
     Proof.
     Admitted.
 
-  End Lifting. 
+  End Lifting.
 
 End IrisMeta.
diff --git a/iris_unsafe.v b/iris_unsafe.v
index d8a425165..1ebfce9ed 100644
--- a/iris_unsafe.v
+++ b/iris_unsafe.v
@@ -6,10 +6,9 @@ Require Import ModuRes.RA ModuRes.UPred ModuRes.BI ModuRes.PreoMet ModuRes.Finma
 Module Unsafety (RL : RA_T) (C : CORE_LANG).
 
   Module Export Iris := IrisWP RL C.
-  Local Open Scope mask_scope.
+  Local Open Scope lang_scope.
   Local Open Scope ra_scope.
   Local Open Scope bi_scope.
-  Local Open Scope lang_scope.
   Local Open Scope iris_scope.
 
   (* PDS: Move to iris_core.v *)
@@ -52,7 +51,6 @@ Module Unsafety (RL : RA_T) (C : CORE_LANG).
 
 	PDS: Should be moved or discarded.
   *)
-  Notation "p * q" := (BI.sc p q) (at level 40, left associativity) : iris_scope. (* RJ: there's already notation for this in iris_core? *) (* PDS: The notation in Iris core uses sc : UPred (ra_pos res) -> UPred (ra_pos res) -> UPred (ra_pos res) rather than BI.sc. This variant is generic, so it survives more simplification. *)
 
   Lemma wpO {safe m e Q w r} : wp safe m e Q w O r.
   Proof.
diff --git a/iris_vs.v b/iris_vs.v
index 8f1003bf0..9a2ba7f9c 100644
--- a/iris_vs.v
+++ b/iris_vs.v
@@ -7,20 +7,21 @@ Set Bullet Behavior "Strict Subproofs".
 Module IrisVS (RL : RA_T) (C : CORE_LANG).
   Module Export CORE := IrisCore RL C.
 
-  Delimit Scope iris_scope with iris.
+  Local Open Scope ra_scope.
+  Local Open Scope bi_scope.
   Local Open Scope iris_scope.
 
+  Implicit Types (P Q R : Props) (w : Wld) (n i k : nat) (m : mask) (r : pres) (σ : state).
+
   Section ViewShifts.
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
     Local Obligation Tactic := intros.
 
-    Program Definition preVS (m1 m2 : mask) (p : Props) (w : Wld) : UPred pres :=
+    Program Definition preVS m1 m2 P w : UPred pres :=
       mkUPred (fun n r => forall w1 (rf: res) mf σ k (HSub : w ⊑ w1) (HLe : k < n)
                                  (HD : mf # m1 ∪ m2)
                                  (HE : wsat σ (m1 ∪ mf) (ra_proj r · rf) w1 @ S k),
                           exists w2 r',
-                            w1 ⊑ w2 /\ p w2 (S k) r'
+                            w1 ⊑ w2 /\ P w2 (S k) r'
                             /\ wsat σ (m2 ∪ mf) (r' · rf) w2 @ S k) _.
     Next Obligation.
       intros n1 n2 r1 r2 HLe [rd HR] HP; intros.
@@ -35,8 +36,8 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
         eapply uni_pred, HP'; [reflexivity|]. exists rd. reflexivity.
     Qed.
 
-    Program Definition pvs (m1 m2 : mask) : Props -n> Props :=
-      n[(fun p => m[(preVS m1 m2 p)])].
+    Program Definition pvs m1 m2 : Props -n> Props :=
+      n[(fun P => m[(preVS m1 m2 P)])].
     Next Obligation.
       intros w1 w2 EQw n' r HLt; destruct n as [| n]; [now inversion HLt |]; split; intros HP w2'; intros.
       - symmetry in EQw; assert (HDE := extend_dist _ _ _ _ EQw HSub).
@@ -46,7 +47,7 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
         + symmetry in HDE; assert (HDE' := extend_dist _ _ _ _ HDE HW).
           assert (HSE' := extend_sub _ _ _ _ HDE HW); destruct HH as [HP HE'];
           exists (extend w1'' w2') r'; split; [assumption | split].
-          * eapply (met_morph_nonexp _ _ p), HP ; [symmetry; eassumption | omega].
+          * eapply (met_morph_nonexp _ _ P), HP ; [symmetry; eassumption | omega].
           * eapply wsat_dist, HE'; [symmetry; eassumption | omega].
       - assert (HDE := extend_dist _ _ _ _ EQw HSub); assert (HSE := extend_sub _ _ _ _ EQw HSub); specialize (HP (extend w2' w2)).
         edestruct HP as [w1'' [r' [HW HH] ] ]; try eassumption; clear HP; [ | ].
@@ -54,7 +55,7 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
         + symmetry in HDE; assert (HDE' := extend_dist _ _ _ _ HDE HW).
           assert (HSE' := extend_sub _ _ _ _ HDE HW); destruct HH as [HP HE'];
           exists (extend w1'' w2') r'; split; [assumption | split].
-          * eapply (met_morph_nonexp _ _ p), HP ; [symmetry; eassumption | omega].
+          * eapply (met_morph_nonexp _ _ P), HP ; [symmetry; eassumption | omega].
           * eapply wsat_dist, HE'; [symmetry; eassumption | omega].
     Qed.
     Next Obligation.
@@ -71,22 +72,17 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
         apply EQp; [now eauto with arith | assumption].
     Qed.
 
-    Definition vs (m1 m2 : mask) (p q : Props) : Props :=
-      □ (p → pvs m1 m2 q).
+    Definition vs m1 m2 P Q : Props :=
+      □(P → pvs m1 m2 Q).
 
   End ViewShifts.
 
   Section ViewShiftProps.
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
-
-    Implicit Types (p q r : Props) (i : nat) (m : mask).
 
     Definition mask_sing i := mask_set mask_emp i True.
 
-    Lemma vsTimeless m p :
-      timeless p ⊑ vs m m (▹ p) p.
+    Lemma vsTimeless m P :
+      timeless P ⊑ vs m m (▹P) P.
     Proof.
       intros w' n r1 HTL w HSub; rewrite ->HSub in HTL; clear w' HSub.
       intros np rp HLe HS Hp w1; intros.
@@ -96,13 +92,13 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
       rewrite <- HSub; apply HTL, Hp; [reflexivity | assumption].
     Qed.
 
-    Lemma vsOpen i p :
-      valid (vs (mask_sing i) mask_emp (inv i p) (â–¹ p)).
+    Lemma vsOpen i P :
+      valid (vs (mask_sing i) mask_emp (inv i P) (â–¹P)).
     Proof.
       intros pw nn r w _; clear r pw.
       intros n r _ _ HInv w'; clear nn; intros.
       do 14 red in HInv; destruct (w i) as [μ |] eqn: HLu; [| contradiction].
-      apply ı in HInv; rewrite ->(isoR p) in HInv.
+      apply ı in HInv; rewrite ->(isoR P) in HInv.
       (* get rid of the invisible 1/2 *)
       do 8 red in HInv.
       destruct HE as [rs [HE HM] ].
@@ -132,13 +128,13 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
         destruct (Peano_dec.eq_nat_dec i i); tauto.
     Qed.
 
-    Lemma vsClose i p :
-      valid (vs mask_emp (mask_sing i) (inv i p * ▹ p) ⊤).
+    Lemma vsClose i P :
+      valid (vs mask_emp (mask_sing i) (inv i P * ▹P) ⊤).
     Proof.
       intros pw nn r w _; clear r pw.
       intros n r _ _ [r1 [r2 [HR [HInv HP] ] ] ] w'; clear nn; intros.
       do 14 red in HInv; destruct (w i) as [μ |] eqn: HLu; [| contradiction].
-      apply ı in HInv; rewrite ->(isoR p) in HInv.
+      apply ı in HInv; rewrite ->(isoR P) in HInv.
       (* get rid of the invisible 1/2 *)
       do 8 red in HInv.
       destruct HE as [rs [HE HM] ].
@@ -179,8 +175,8 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
         rewrite <-HR, assoc. reflexivity.
     Qed.
 
-    Lemma vsTrans p q r m1 m2 m3 (HMS : m2 ⊆ m1 ∪ m3) :
-      vs m1 m2 p q ∧ vs m2 m3 q r ⊑ vs m1 m3 p r.
+    Lemma vsTrans P Q R m1 m2 m3 (HMS : m2 ⊆ m1 ∪ m3) :
+      vs m1 m2 P Q ∧ vs m2 m3 Q R ⊑ vs m1 m3 P R.
     Proof.
       intros w' n r1 [Hpq Hqr] w HSub; specialize (Hpq _ HSub); rewrite ->HSub in Hqr; clear w' HSub.
       intros np rp HLe HS Hp w1; intros; specialize (Hpq _ _ HLe HS Hp).
@@ -198,8 +194,8 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
       setoid_rewrite HSw12; eauto 8.
     Qed.
 
-    Lemma vsEnt p q m :
-      □ (p → q) ⊑ vs m m p q.
+    Lemma vsEnt P Q m :
+      □(P → Q) ⊑ vs m m P Q.
     Proof.
       intros w' n r1 Himp w HSub; rewrite ->HSub in Himp; clear w' HSub.
       intros np rp HLe HS Hp w1; intros.
@@ -208,8 +204,8 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
       unfold lt in HLe0; rewrite ->HLe0, <- HSub; assumption.
     Qed.
 
-    Lemma vsFrame p q r m1 m2 mf (HDisj : mf # m1 ∪ m2) :
-      vs m1 m2 p q ⊑ vs (m1 ∪ mf) (m2 ∪ mf) (p * r) (q * r).
+    Lemma vsFrame P Q R m1 m2 mf (HDisj : mf # m1 ∪ m2) :
+      vs m1 m2 P Q ⊑ vs (m1 ∪ mf) (m2 ∪ mf) (P * R) (Q * R).
     Proof.
       intros w' n r1 HVS w HSub; specialize (HVS _ HSub); clear w' r1 HSub.
       intros np rpr HLe _ [rp [rr [HR [Hp Hr] ] ] ] w'; intros.
@@ -270,7 +266,7 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
     Qed.
 
     Program Definition inv' m : Props -n> {n : nat | m n} -n> Props :=
-      n[(fun p => n[(fun n => inv n p)])].
+      n[(fun P => n[(fun n : {n | m n} => inv n P)])].
     Next Obligation.
       intros i i' EQi; destruct n as [| n]; [apply dist_bound |].
       simpl in EQi; rewrite EQi; reflexivity.
@@ -297,19 +293,19 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
       intros σ σc Hp; apply Hp.
     Qed.
 
-    Lemma vsNewInv p m (HInf : mask_infinite m) :
-      valid (vs m m (â–¹ p) (xist (inv' m p))).
+    Lemma vsNewInv P m (HInf : mask_infinite m) :
+      valid (vs m m (â–¹P) (xist (inv' m P))).
     Proof.
       intros pw nn r w _; clear r pw.
       intros n r _ _ HP w'; clear nn; intros.
       destruct n as [| n]; [now inversion HLe | simpl in HP].
       rewrite ->HSub in HP; clear w HSub; rename w' into w.
       destruct (fresh_region w m HInf) as [i [Hm HLi] ].
-      assert (HSub : w ⊑ fdUpdate i (ı' p) w).
+      assert (HSub : w ⊑ fdUpdate i (ı' P) w).
       { intros j; destruct (Peano_dec.eq_nat_dec i j); [subst j; rewrite HLi; exact I|].
         now rewrite ->fdUpdate_neq by assumption.
       }
-      exists (fdUpdate i (ı' p) w) (ra_pos_unit); split; [assumption | split].
+      exists (fdUpdate i (ı' P) w) (ra_pos_unit); split; [assumption | split].
       - exists (exist _ i Hm). do 22 red.
         unfold proj1_sig. rewrite fdUpdate_eq; reflexivity.
       - unfold ra_pos_unit, proj1_sig. erewrite ra_op_unit by apply _.
diff --git a/iris_wp.v b/iris_wp.v
index 248e7167b..69559bbbb 100644
--- a/iris_wp.v
+++ b/iris_wp.v
@@ -8,15 +8,13 @@ Module IrisWP (RL : RA_T) (C : CORE_LANG).
   Module Export L  := Lang C.
   Module Export VS := IrisVS RL C.
 
-  Delimit Scope iris_scope with iris.
+  Local Open Scope lang_scope.
+  Local Open Scope ra_scope.
+  Local Open Scope bi_scope.
   Local Open Scope iris_scope.
 
   Section HoareTriples.
   (* Quadruples, really *)
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
-    Local Open Scope lang_scope.
 
     Instance LP_isval : LimitPreserving is_value.
     Proof.
@@ -199,15 +197,11 @@ Module IrisWP (RL : RA_T) (C : CORE_LANG).
 
     Opaque wp.
 
-    Definition ht safe m P e Q := □ (P → wp safe m e Q).
+    Definition ht safe m P e Q := □(P → wp safe m e Q).
 
   End HoareTriples.
   
   Section HoareTripleProperties.
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
-    Local Open Scope lang_scope.
 
     Existing Instance LP_isval.
 
@@ -474,7 +468,7 @@ Module IrisWP (RL : RA_T) (C : CORE_LANG).
     Lemma htAFrame safe m m' P R e Q
           (HD  : m # m')
           (HAt : atomic e) :
-      ht safe m P e Q ⊑ ht safe (m ∪ m') (P * ▹ R) e (lift_bin sc Q (umconst R)).
+      ht safe m P e Q ⊑ ht safe (m ∪ m') (P * ▹R) e (lift_bin sc Q (umconst R)).
     Proof.
       intros w n rz He w' HSw n' r HLe _ [r1 [r2 [EQr [HP HLR] ] ] ].
       specialize (He _ HSw _ _ HLe (unit_min _ _) HP).
@@ -521,7 +515,7 @@ Module IrisWP (RL : RA_T) (C : CORE_LANG).
 
     (** Fork **)
     Lemma htFork safe m P R e :
-      ht safe m P e (umconst ⊤) ⊑ ht safe m (▹ P * ▹ R) (fork e) (lift_bin sc (eqV (exist _ fork_ret fork_ret_is_value)) (umconst R)).
+      ht safe m P e (umconst ⊤) ⊑ ht safe m (▹P * ▹R) (fork e) (lift_bin sc (eqV (exist _ fork_ret fork_ret_is_value)) (umconst R)).
     Proof.
       intros w n rz He w' HSw n' r HLe _ [r1 [r2 [EQr [HP HLR] ] ] ].
       destruct n' as [| n']; [apply wpO |].
@@ -553,10 +547,6 @@ Module IrisWP (RL : RA_T) (C : CORE_LANG).
   End HoareTripleProperties.
 
   Section DerivedRules.
-    Local Open Scope mask_scope.
-    Local Open Scope ra_scope.
-    Local Open Scope bi_scope.
-    Local Open Scope lang_scope.
 
     Existing Instance LP_isval.
 
diff --git a/lib/ModuRes/BI.v b/lib/ModuRes/BI.v
index f05946fc5..2864bc6d3 100644
--- a/lib/ModuRes/BI.v
+++ b/lib/ModuRes/BI.v
@@ -102,7 +102,19 @@ Class Later (T : Type) `{pcmT : pcmType T} {vT : Valid T} :=
     loeb (t : T) (HL : later t ⊑ t) : valid t
   }.
 
+Delimit Scope bi_scope with bi.
 Notation " â–¹ p " := (later p) (at level 20) : bi_scope.
+Notation "⊤" := (top) : bi_scope.
+Notation "⊥" := (bot) : bi_scope.
+Notation "p ∧ q" := (and p q) (at level 40, left associativity) : bi_scope.
+Notation "p ∨ q" := (or p q) (at level 50, left associativity) : bi_scope.
+Notation "p * q" := (sc p q) (at level 40, left associativity) : bi_scope.
+Notation "p → q" := (impl p q) (at level 55, right associativity) : bi_scope.
+Notation "p '-*' q" := (si p q) (at level 55, right associativity) : bi_scope.
+Notation "∀ x , p" := (all n[(fun x => p)]) (at level 60, x ident, no associativity) : bi_scope.
+Notation "∃ x , p" := (xist n[(fun x => p)]) (at level 60, x ident, no associativity) : bi_scope.
+Notation "∀ x : T , p" := (all n[(fun x : T => p)]) (at level 60, x ident, no associativity) : bi_scope.
+Notation "∃ x : T , p" := (xist n[(fun x : T => p)]) (at level 60, x ident, no associativity) : bi_scope.
 
 Require Import UPred.
 
@@ -144,7 +156,7 @@ Section UPredBI.
   Local Open Scope ra_scope.
   Local Obligation Tactic := intros; eauto with typeclass_instances.
 
-  Definition pres := ra_pos res.
+  Let pres := ⁺res.
 
   (* Standard interpretations of propositional connectives. *)
   Global Program Instance top_up : topBI (UPred pres) := up_cr (const True).
diff --git a/lib/ModuRes/RA.v b/lib/ModuRes/RA.v
index 827920079..0874c92dc 100644
--- a/lib/ModuRes/RA.v
+++ b/lib/ModuRes/RA.v
@@ -30,15 +30,18 @@ Section Definitions.
 End Definitions.
 Arguments ra_valid {T} {_} t.
 
+Delimit Scope ra_scope with ra.
+Local Open Scope ra_scope.
+
 Notation "1" := (ra_unit _) : ra_scope.
 Notation "p · q" := (ra_op _ p q) (at level 40, left associativity) : ra_scope.
 Notation "'↓' p" := (ra_valid p) (at level 35) : ra_scope.
-Delimit Scope ra_scope with ra.
 
 (* General RA lemmas *)
 Section RAs.
   Context {T} `{RA T}.
-  Local Open Scope ra_scope.
+
+  Implicit Types (t : T).
 
   Lemma ra_op_unit2 t: t · 1 == t.
   Proof.
@@ -66,7 +69,6 @@ End RAs.
 (* RAs with cartesian products of carriers. *)
 Section Products.
   Context S T `{raS : RA S, raT : RA T}.
-  Local Open Scope ra_scope.
 
   Global Instance ra_unit_prod : RA_unit (S * T) := (ra_unit S, ra_unit T).
   Global Instance ra_op_prod : RA_op (S * T) :=
@@ -96,7 +98,6 @@ End Products.
 
 Section ProductLemmas.
   Context {S T} `{raS : RA S, raT : RA T}.
-  Local Open Scope ra_scope.
     
   Lemma ra_op_prod_fst (st1 st2: S*T):
     fst (st1 · st2) = fst st1 · fst st2.
@@ -121,10 +122,11 @@ End ProductLemmas.
 
 Section PositiveCarrier.
   Context {T} `{raT : RA T}.
-  Local Open Scope ra_scope.
 
-  Definition ra_pos: Type := { r | ↓ r }.
-  Coercion ra_proj (t:ra_pos): T := proj1_sig t.
+  Definition ra_pos: Type := { t : T | ↓ t }.
+  Coercion ra_proj (r:ra_pos): T := proj1_sig r.
+
+  Implicit Types (t u : T) (r : ra_pos).
 
   Global Instance ra_pos_type : Setoid ra_pos := _.
 
@@ -135,26 +137,26 @@ Section PositiveCarrier.
     now eapply ra_valid_unit; apply _.
   Qed.
 
-  Lemma ra_proj_cancel r (VAL: ↓r):
-    ra_proj (ra_mk_pos r (VAL:=VAL)) = r.
+  Lemma ra_proj_cancel t (VAL: ↓t):
+    ra_proj (ra_mk_pos t (VAL:=VAL)) = t.
   Proof.
     reflexivity.
   Qed.
 
-  Lemma ra_op_pos_valid t1 t2 t:
-    t1 · t2 == ra_proj t -> ↓ t1.
+  Lemma ra_op_pos_valid t1 t2 r:
+    t1 · t2 == ra_proj r -> ↓ t1.
   Proof.
-    destruct t as [t Hval]; simpl. intros Heq. rewrite <-Heq in Hval.
+    destruct r as [t Hval]; simpl. intros Heq. rewrite <-Heq in Hval.
     eapply ra_op_valid; eassumption.
   Qed.
 
-  Lemma ra_op_pos_valid2 t1 t2 t:
-    t1 · t2 == ra_proj t -> ↓ t2.
+  Lemma ra_op_pos_valid2 t1 t2 r:
+    t1 · t2 == ra_proj r -> ↓ t2.
   Proof.
     rewrite comm. now eapply ra_op_pos_valid.
   Qed.
 
-  Lemma ra_pos_valid (r : ra_pos):
+  Lemma ra_pos_valid r:
     ↓(ra_proj r).
   Proof.
     destruct r as [r VAL].
@@ -163,6 +165,7 @@ Section PositiveCarrier.
 
 End PositiveCarrier.
 Global Arguments ra_pos T {_}.
+Notation "'⁺' T" := (ra_pos T) (at level 75) : type_scope.
 
 (** Validity automation tactics *)
 Ltac auto_valid_inner :=
@@ -186,10 +189,11 @@ Tactic Notation "pose↓" ident(name) ":=" constr(t) "by" tactic(tac) := pose_va
 
 Section Order.
   Context T `{raT : RA T}.
-  Local Open Scope ra_scope.
 
-  Definition pra_ord (t1 t2 : ra_pos T) :=
-    exists td, td · (ra_proj t1) == (ra_proj t2).
+  Implicit Types (t : T) (r s : ra_pos T).
+
+  Definition pra_ord r1 r2 :=
+    exists t, t · (ra_proj r1) == (ra_proj r2).
 
   Global Instance pra_ord_preo: PreOrder pra_ord.
   Proof.
@@ -215,7 +219,7 @@ Section Order.
     now erewrite ra_op_unit2 by apply _.
   Qed.
 
-  Definition ra_ord (t1 t2 : T) :=
+  Definition ra_ord t1 t2 :=
     exists t, t · t1 == t2.
 
   Global Instance ra_ord_preo: PreOrder ra_ord.
@@ -253,16 +257,17 @@ End Order.
 
 Section Exclusive.
   Context (T: Type) `{Setoid T}.
-  Local Open Scope ra_scope.
-
+  
   Inductive ra_res_ex: Type :=
   | ex_own: T -> ra_res_ex
   | ex_unit: ra_res_ex
   | ex_bot: ra_res_ex.
 
-  Definition ra_res_ex_eq e1 e2: Prop :=
-    match e1, e2 with
-      | ex_own s1, ex_own s2 => s1 == s2
+  Implicit Types (r s : ra_res_ex).
+
+  Definition ra_res_ex_eq r s: Prop :=
+    match r, s with
+      | ex_own t1, ex_own t2 => t1 == t2
       | ex_unit, ex_unit => True
       | ex_bot, ex_bot => True
       | _, _ => False
@@ -282,15 +287,15 @@ Section Exclusive.
       
   Global Instance ra_unit_ex : RA_unit ra_res_ex := ex_unit.
   Global Instance ra_op_ex : RA_op ra_res_ex :=
-    fun e1 e2 =>
-      match e1, e2 with
-        | ex_own s1, ex_unit => ex_own s1
-        | ex_unit, ex_own s2 => ex_own s2
+    fun r s =>
+      match r, s with
+        | ex_own _, ex_unit => r
+        | ex_unit, ex_own _ => s
         | ex_unit, ex_unit   => ex_unit
         | _, _               => ex_bot
       end.
   Global Instance ra_valid_ex : RA_valid ra_res_ex :=
-    fun e => match e with
+    fun r => match r with
                | ex_bot => False
                | _      => True
              end.
@@ -310,7 +315,7 @@ Section Exclusive.
 End Exclusive.
 
 
-(* Package of a monoid as a module type (for use with other modules). *)
+(* Package a RA as a module type (for use with other modules). *)
 Module Type RA_T.
 
   Parameter res : Type.
diff --git a/masks.v b/masks.v
index 4fc39102f..baae9e2a8 100644
--- a/masks.v
+++ b/masks.v
@@ -39,6 +39,7 @@ Qed.
 Global Instance mask_type : Setoid mask := mkType meq.
 
 Delimit Scope mask_scope with mask.
+Bind Scope mask_scope with mask.
 Notation "m1 ⊆ m2"  := (mle m1 m2)  (at level 70) : mask_scope.
 Notation "m1 ∩ m2" := (mcap m1 m2) (at level 40) : mask_scope.
 Notation "m1 \  m2"  := (mminus m1 m2) (at level 30) : mask_scope.
-- 
GitLab