From 4b2916c480132c36089a03b585261ad22523fd6e Mon Sep 17 00:00:00 2001
From: Ralf Jung <jung@mpi-sws.org>
Date: Tue, 6 Dec 2016 22:43:34 +0100
Subject: [PATCH] Add locking to value-scope notation for lambdas

---
 heap_lang/derived.v             | 10 ++++++----
 heap_lang/lib/assert.v          |  2 --
 heap_lang/lib/barrier/barrier.v |  1 -
 heap_lang/lib/barrier/proof.v   |  2 --
 heap_lang/lib/counter.v         | 19 +++++++++----------
 heap_lang/lib/par.v             |  2 --
 heap_lang/lib/spawn.v           |  1 -
 heap_lang/lib/spin_lock.v       |  1 -
 heap_lang/lib/ticket_lock.v     | 16 ++++++++--------
 heap_lang/lifting.v             |  6 ------
 heap_lang/notation.v            | 10 +++++-----
 heap_lang/wp_tactics.v          | 19 +++++++++++++++++--
 tests/barrier_client.v          |  2 --
 tests/counter.v                 |  2 --
 tests/heap_lang.v               |  1 -
 tests/joining_existentials.v    |  2 --
 tests/list_reverse.v            |  1 -
 tests/one_shot.v                |  2 --
 tests/tree_sum.v                |  3 ---
 19 files changed, 45 insertions(+), 57 deletions(-)

diff --git a/heap_lang/derived.v b/heap_lang/derived.v
index 9a95ee5c8..367f34eaa 100644
--- a/heap_lang/derived.v
+++ b/heap_lang/derived.v
@@ -17,15 +17,17 @@ Implicit Types P Q : iProp Σ.
 Implicit Types Φ : val → iProp Σ.
 
 (** Proof rules for the sugar *)
-Lemma wp_lam E x ef e Φ :
-  is_Some (to_val e) → Closed (x :b: []) ef →
-  ▷ WP subst' x e ef @ E {{ Φ }} ⊢ WP App (Lam x ef) e @ E {{ Φ }}.
+Lemma wp_lam E x elam e1 e2 Φ :
+  e1 = Lam x elam →
+  is_Some (to_val e2) →
+  Closed (x :b: []) elam →
+  ▷ WP subst' x e2 elam @ E {{ Φ }} ⊢ WP App e1 e2 @ E {{ Φ }}.
 Proof. intros. by rewrite -(wp_rec _ BAnon) //. Qed.
 
 Lemma wp_let E x e1 e2 Φ :
   is_Some (to_val e1) → Closed (x :b: []) e2 →
   ▷ WP subst' x e1 e2 @ E {{ Φ }} ⊢ WP Let x e1 e2 @ E {{ Φ }}.
-Proof. apply wp_lam. Qed.
+Proof. by apply wp_lam. Qed.
 
 Lemma wp_seq E e1 e2 Φ :
   is_Some (to_val e1) → Closed [] e2 →
diff --git a/heap_lang/lib/assert.v b/heap_lang/lib/assert.v
index 4de019169..fa9b8cd25 100644
--- a/heap_lang/lib/assert.v
+++ b/heap_lang/lib/assert.v
@@ -14,5 +14,3 @@ Proof.
   iIntros "HΦ". rewrite /assert. wp_let. wp_seq.
   iApply (wp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if.
 Qed.
-
-Global Opaque assert.
diff --git a/heap_lang/lib/barrier/barrier.v b/heap_lang/lib/barrier/barrier.v
index f30edf8dd..61f8f7252 100644
--- a/heap_lang/lib/barrier/barrier.v
+++ b/heap_lang/lib/barrier/barrier.v
@@ -4,4 +4,3 @@ Definition newbarrier : val := λ: <>, ref #false.
 Definition signal : val := λ: "x", "x" <- #true.
 Definition wait : val :=
   rec: "wait" "x" := if: !"x" then #() else "wait" "x".
-Global Opaque newbarrier signal wait.
diff --git a/heap_lang/lib/barrier/proof.v b/heap_lang/lib/barrier/proof.v
index b67477be2..2b65b88c7 100644
--- a/heap_lang/lib/barrier/proof.v
+++ b/heap_lang/lib/barrier/proof.v
@@ -22,8 +22,6 @@ Section proof.
 Context `{!heapG Σ, !barrierG Σ} (N : namespace).
 Implicit Types I : gset gname.
 
-Local Transparent newbarrier signal wait.
-
 Definition ress (P : iProp Σ) (I : gset gname) : iProp Σ :=
   (∃ Ψ : gname → iProp Σ,
     ▷ (P -∗ [∗ set] i ∈ I, Ψ i) ∗ [∗ set] i ∈ I, saved_prop_own i (Ψ i))%I.
diff --git a/heap_lang/lib/counter.v b/heap_lang/lib/counter.v
index 097bd7fb6..72d16f830 100644
--- a/heap_lang/lib/counter.v
+++ b/heap_lang/lib/counter.v
@@ -4,12 +4,11 @@ From iris.proofmode Require Import tactics.
 From iris.algebra Require Import frac auth.
 From iris.heap_lang Require Import proofmode notation.
 
-Definition newcounter : val := locked (λ: <>, ref #0)%V.
-Definition incr : val := locked (
-  rec: "incr" "l" :=
+Definition newcounter : val := λ: <>, ref #0.
+Definition incr : val := rec: "incr" "l" :=
     let: "n" := !"l" in
-    if: CAS "l" "n" (#1 + "n") then #() else "incr" "l")%V.
-Definition read : val := locked (λ: "l", !"l")%V.
+    if: CAS "l" "n" (#1 + "n") then #() else "incr" "l".
+Definition read : val := λ: "l", !"l".
 
 (** Monotone counter *)
 Class mcounterG Σ := MCounterG { mcounter_inG :> inG Σ (authR mnatUR) }.
@@ -36,7 +35,7 @@ Section mono_proof.
     heapN ⊥ N →
     {{{ heap_ctx }}} newcounter #() {{{ l, RET #l; mcounter l 0 }}}.
   Proof.
-    iIntros (? Φ) "#Hh HΦ". rewrite -wp_fupd /newcounter -lock /=. wp_seq. wp_alloc l as "Hl".
+    iIntros (? Φ) "#Hh HΦ". rewrite -wp_fupd /newcounter. wp_seq. wp_alloc l as "Hl".
     iMod (own_alloc (● (O:mnat) ⋅ ◯ (O:mnat))) as (γ) "[Hγ Hγ']"; first done.
     iMod (inv_alloc N _ (mcounter_inv γ l) with "[Hl Hγ]").
     { iNext. iExists 0%nat. by iFrame. }
@@ -72,7 +71,7 @@ Section mono_proof.
     {{{ mcounter l j }}} read #l {{{ i, RET #i; ⌜j ≤ i⌝%nat ∧ mcounter l i }}}.
   Proof.
     iIntros (ϕ) "Hc HΦ". iDestruct "Hc" as (γ) "(% & #? & #Hinv & Hγf)".
-    rewrite /read -lock /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
+    rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
     iDestruct (own_valid_2 with "Hγ Hγf")
       as %[?%mnat_included _]%auth_valid_discrete_2.
     iMod (own_update_2 with "Hγ Hγf") as "[Hγ Hγf]".
@@ -113,7 +112,7 @@ Section contrib_spec.
     {{{ heap_ctx }}} newcounter #()
     {{{ γ l, RET #l; ccounter_ctx γ l ∗ ccounter γ 1 0 }}}.
   Proof.
-    iIntros (? Φ) "#Hh HΦ". rewrite -wp_fupd /newcounter -lock /=. wp_seq. wp_alloc l as "Hl".
+    iIntros (? Φ) "#Hh HΦ". rewrite -wp_fupd /newcounter /=. wp_seq. wp_alloc l as "Hl".
     iMod (own_alloc (● (Some (1%Qp, O%nat)) ⋅ ◯ (Some (1%Qp, 0%nat))))
       as (γ) "[Hγ Hγ']"; first done.
     iMod (inv_alloc N _ (ccounter_inv γ l) with "[Hl Hγ]").
@@ -147,7 +146,7 @@ Section contrib_spec.
     {{{ c, RET #c; ⌜n ≤ c⌝%nat ∧ ccounter γ q n }}}.
   Proof.
     iIntros (Φ) "(#(%&?&?) & Hγf) HΦ".
-    rewrite /read -lock /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
+    rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
     iDestruct (own_valid_2 with "Hγ Hγf")
       as %[[? ?%nat_included]%Some_pair_included_total_2 _]%auth_valid_discrete_2.
     iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
@@ -159,7 +158,7 @@ Section contrib_spec.
     {{{ n, RET #n; ccounter γ 1 n }}}.
   Proof.
     iIntros (Φ) "(#(%&?&?) & Hγf) HΦ".
-    rewrite /read -lock /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
+    rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
     iDestruct (own_valid_2 with "Hγ Hγf") as %[Hn _]%auth_valid_discrete_2.
     apply (Some_included_exclusive _) in Hn as [= ->]%leibniz_equiv; last done.
     iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
diff --git a/heap_lang/lib/par.v b/heap_lang/lib/par.v
index 6b4b310aa..3c15c0cd5 100644
--- a/heap_lang/lib/par.v
+++ b/heap_lang/lib/par.v
@@ -44,5 +44,3 @@ Proof.
   iSplitL "H1"; by wp_let.
 Qed.
 End proof.
-
-Global Opaque par.
diff --git a/heap_lang/lib/spawn.v b/heap_lang/lib/spawn.v
index d424aa7fa..20cbb59b4 100644
--- a/heap_lang/lib/spawn.v
+++ b/heap_lang/lib/spawn.v
@@ -78,4 +78,3 @@ Qed.
 End proof.
 
 Typeclasses Opaque join_handle.
-Global Opaque spawn join.
diff --git a/heap_lang/lib/spin_lock.v b/heap_lang/lib/spin_lock.v
index 46e8715bc..e9bcad0c6 100644
--- a/heap_lang/lib/spin_lock.v
+++ b/heap_lang/lib/spin_lock.v
@@ -89,7 +89,6 @@ Section proof.
 End proof.
 
 Typeclasses Opaque is_lock locked.
-Global Opaque newlock try_acquire acquire release.
 
 Definition spin_lock `{!heapG Σ, !lockG Σ} : lock Σ :=
   {| lock.locked_exclusive := locked_exclusive; lock.newlock_spec := newlock_spec;
diff --git a/heap_lang/lib/ticket_lock.v b/heap_lang/lib/ticket_lock.v
index 8e16f24ec..e4c691e6b 100644
--- a/heap_lang/lib/ticket_lock.v
+++ b/heap_lang/lib/ticket_lock.v
@@ -7,24 +7,24 @@ From iris.heap_lang.lib Require Export lock.
 Import uPred.
 
 Definition wait_loop: val :=
-  ssreflect.locked (rec: "wait_loop" "x" "lk" :=
+  rec: "wait_loop" "x" "lk" :=
     let: "o" := !(Fst "lk") in
     if: "x" = "o"
       then #() (* my turn *)
-      else "wait_loop" "x" "lk")%V.
+      else "wait_loop" "x" "lk".
 
 Definition newlock : val :=
-  ssreflect.locked (λ: <>, ((* owner *) ref #0, (* next *) ref #0))%V.
+  λ: <>, ((* owner *) ref #0, (* next *) ref #0).
 
 Definition acquire : val :=
-  ssreflect.locked (rec: "acquire" "lk" :=
+  rec: "acquire" "lk" :=
     let: "n" := !(Snd "lk") in
     if: CAS (Snd "lk") "n" ("n" + #1)
       then wait_loop "n" "lk"
-      else "acquire" "lk")%V.
+      else "acquire" "lk".
 
 Definition release : val :=
-  ssreflect.locked (λ: "lk", (Fst "lk") <- !(Fst "lk") + #1)%V.
+  λ: "lk", (Fst "lk") <- !(Fst "lk") + #1.
 
 (** The CMRAs we need. *)
 Class tlockG Σ :=
@@ -77,7 +77,7 @@ Section proof.
     heapN ⊥ N →
     {{{ heap_ctx ∗ R }}} newlock #() {{{ lk γ, RET lk; is_lock γ lk R }}}.
   Proof.
-    iIntros (HN Φ) "(#Hh & HR) HΦ". rewrite -wp_fupd /newlock. unlock.
+    iIntros (HN Φ) "(#Hh & HR) HΦ". rewrite -wp_fupd /newlock.
     wp_seq. wp_alloc lo as "Hlo". wp_alloc ln as "Hln".
     iMod (own_alloc (● (Excl' 0%nat, ∅) ⋅ ◯ (Excl' 0%nat, ∅))) as (γ) "[Hγ Hγ']".
     { by rewrite -auth_both_op. }
@@ -145,7 +145,7 @@ Section proof.
     iIntros (Φ) "(Hl & Hγ & HR) HΦ".
     iDestruct "Hl" as (lo ln) "(% & #? & % & #?)"; subst.
     iDestruct "Hγ" as (o) "Hγo".
-    rewrite /release. unlock. wp_let. wp_proj. wp_proj. wp_bind (! _)%E.
+    rewrite /release. wp_let. wp_proj. wp_proj. wp_bind (! _)%E.
     iInv N as (o' n) "(>Hlo & >Hln & >Hauth & Haown)" "Hclose".
     wp_load.
     iDestruct (own_valid_2 with "Hauth Hγo") as
diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v
index 7d45d5501..eb07aaa0b 100644
--- a/heap_lang/lifting.v
+++ b/heap_lang/lifting.v
@@ -111,12 +111,6 @@ Proof.
   intros; inv_head_step; eauto.
 Qed.
 
-Lemma wp_rec_locked E f x erec e1 e2 Φ `{!Closed (f :b: x :b: []) erec} :
-  e1 = of_val $ locked (RecV f x erec) →
-  is_Some (to_val e2) →
-  ▷ WP subst' x e2 (subst' f e1 erec) @ E {{ Φ }} ⊢ WP App e1 e2 @ E {{ Φ }}.
-Proof. unlock. auto using wp_rec. Qed.
-
 Lemma wp_un_op E op e v v' Φ :
   to_val e = Some v →
   un_op_eval op v = Some v' →
diff --git a/heap_lang/notation.v b/heap_lang/notation.v
index 040c5ae21..bd2bb6ff4 100644
--- a/heap_lang/notation.v
+++ b/heap_lang/notation.v
@@ -47,7 +47,7 @@ Notation "~ e" := (UnOp NegOp e%E) (at level 75, right associativity) : expr_sco
 Notation "e1 <- e2" := (Store e1%E e2%E) (at level 80) : expr_scope.
 Notation "'rec:' f x := e" := (Rec f%bind x%bind e%E)
   (at level 102, f at level 1, x at level 1, e at level 200) : expr_scope.
-Notation "'rec:' f x := e" := (RecV f%bind x%bind e%E)
+Notation "'rec:' f x := e" := (locked (RecV f%bind x%bind e%E))
   (at level 102, f at level 1, x at level 1, e at level 200) : val_scope.
 Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E)
   (at level 200, e1, e2, e3 at level 200) : expr_scope.
@@ -58,20 +58,20 @@ defined above. This is needed because App is now a coercion, and these
 notations are otherwise not pretty printed back accordingly. *)
 Notation "'rec:' f x y := e" := (Rec f%bind x%bind (Lam y%bind e%E))
   (at level 102, f, x, y at level 1, e at level 200) : expr_scope.
-Notation "'rec:' f x y := e" := (RecV f%bind x%bind (Lam y%bind e%E))
+Notation "'rec:' f x y := e" := (locked (RecV f%bind x%bind (Lam y%bind e%E)))
   (at level 102, f, x, y at level 1, e at level 200) : val_scope.
 Notation "'rec:' f x y .. z := e" := (Rec f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..))
   (at level 102, f, x, y, z at level 1, e at level 200) : expr_scope.
-Notation "'rec:' f x y .. z := e" := (RecV f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..))
+Notation "'rec:' f x y .. z := e" := (locked (RecV f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..)))
   (at level 102, f, x, y, z at level 1, e at level 200) : val_scope.
 
 Notation "λ: x , e" := (Lam x%bind e%E)
   (at level 102, x at level 1, e at level 200) : expr_scope.
 Notation "λ: x y .. z , e" := (Lam x%bind (Lam y%bind .. (Lam z%bind e%E) ..))
   (at level 102, x, y, z at level 1, e at level 200) : expr_scope.
-Notation "λ: x , e" := (LamV x%bind e%E)
+Notation "λ: x , e" := (locked (LamV x%bind e%E))
   (at level 102, x at level 1, e at level 200) : val_scope.
-Notation "λ: x y .. z , e" := (LamV x%bind (Lam y%bind .. (Lam z%bind e%E) .. ))
+Notation "λ: x y .. z , e" := (locked (LamV x%bind (Lam y%bind .. (Lam z%bind e%E) .. )))
   (at level 102, x, y, z at level 1, e at level 200) : val_scope.
 
 Notation "'let:' x := e1 'in' e2" := (Lam x%bind e2%E e1%E)
diff --git a/heap_lang/wp_tactics.v b/heap_lang/wp_tactics.v
index bc647d8e4..e3250b0d1 100644
--- a/heap_lang/wp_tactics.v
+++ b/heap_lang/wp_tactics.v
@@ -44,13 +44,28 @@ Tactic Notation "wp_value" :=
   | _ => fail "wp_value: not a wp"
   end.
 
+(* Applied to goals that are equalities of expressions. Will try to unlock the
+   LHS once if necessary, to get rid of the lock added by the syntactic sugar. *)
+Ltac wp_unlock :=
+  solve [
+      reflexivity | (* If there are no locks, this is enough. *)
+      (* Otherwise, use unification to uncover the lock. *)
+      (* Step 1: Get the LHS into the form "of_val (locked v)" *)
+      let v := fresh "v" in
+      evar (v: val); trans (of_val (locked v)); subst v; first reflexivity;
+      (* Step 2: Remove the lock from the LHS. *)
+      etrans; first solve [ apply (f_equal of_val); symmetry; apply lock ];
+      (* Now things should be convertible. *)
+      reflexivity
+    ].
+
 Tactic Notation "wp_rec" :=
   lazymatch goal with
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match eval hnf in e' with App ?e1 _ =>
 (* hnf does not reduce through an of_val *)
 (*      match eval hnf in e1 with Rec _ _ _ => *)
-    wp_bind_core K; etrans; [|(eapply wp_rec; wp_done) || (eapply wp_rec_locked; wp_done)]; simpl_subst; wp_finish
+    wp_bind_core K; etrans; [|eapply wp_rec; [wp_unlock|wp_done..]]; simpl_subst; wp_finish
 (*      end *) end) || fail "wp_rec: cannot find 'Rec' in" e
   | _ => fail "wp_rec: not a 'wp'"
   end.
@@ -60,7 +75,7 @@ Tactic Notation "wp_lam" :=
   | |- _ ⊢ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
     match eval hnf in e' with App ?e1 _ =>
 (*    match eval hnf in e1 with Rec BAnon _ _ => *)
-    wp_bind_core K; etrans; [|eapply wp_lam; wp_done]; simpl_subst; wp_finish
+    wp_bind_core K; etrans; [|eapply wp_lam; [wp_unlock|wp_done..]]; simpl_subst; wp_finish
 (*    end *) end) || fail "wp_lam: cannot find 'Lam' in" e
   | _ => fail "wp_lam: not a 'wp'"
   end.
diff --git a/tests/barrier_client.v b/tests/barrier_client.v
index 1800ad32c..d9f173c18 100644
--- a/tests/barrier_client.v
+++ b/tests/barrier_client.v
@@ -54,8 +54,6 @@ Section client.
 Qed.
 End client.
 
-Global Opaque worker client.
-
 Section ClosedProofs.
 
 Let Σ : gFunctors := #[ heapΣ ; barrierΣ ; spawnΣ ].
diff --git a/tests/counter.v b/tests/counter.v
index 4b852ad6d..a97e466e2 100644
--- a/tests/counter.v
+++ b/tests/counter.v
@@ -136,5 +136,3 @@ Proof.
   iModIntro; rewrite /C; eauto 10 with omega.
 Qed.
 End proof.
-
-Global Opaque newcounter incr read.
diff --git a/tests/heap_lang.v b/tests/heap_lang.v
index 1ba0ee98b..60160c224 100644
--- a/tests/heap_lang.v
+++ b/tests/heap_lang.v
@@ -37,7 +37,6 @@ Section LiftingTests.
   Definition Pred : val :=
     λ: "x",
       if: "x" ≤ #0 then -FindPred (-"x" + #2) #0 else FindPred "x" #0.
-  Global Opaque FindPred Pred.
 
   Lemma FindPred_spec n1 n2 E Φ :
     n1 < n2 →
diff --git a/tests/joining_existentials.v b/tests/joining_existentials.v
index 700017967..2a143a107 100644
--- a/tests/joining_existentials.v
+++ b/tests/joining_existentials.v
@@ -97,5 +97,3 @@ Proof.
   - iIntros (_ v) "[_ H]". iDestruct (Q_res_join with "H") as "?". auto.
 Qed.
 End proof.
-
-Global Opaque client.
diff --git a/tests/list_reverse.v b/tests/list_reverse.v
index 64565e93f..10cc5bf35 100644
--- a/tests/list_reverse.v
+++ b/tests/list_reverse.v
@@ -24,7 +24,6 @@ Definition rev : val :=
        "l" <- ("tmp1", "acc");;
        "rev" "tmp2" "hd"
     end.
-Global Opaque rev.
 
 Lemma rev_acc_wp hd acc xs ys (Φ : val → iProp Σ) :
   heap_ctx ∗ is_list hd xs ∗ is_list acc ys ∗
diff --git a/tests/one_shot.v b/tests/one_shot.v
index 323f8896a..0f7b99c73 100644
--- a/tests/one_shot.v
+++ b/tests/one_shot.v
@@ -96,5 +96,3 @@ Proof.
     iApply (wp_wand with "Hf2"). by iIntros (v) "#? !# _".
 Qed.
 End proof.
-
-Global Opaque one_shot_example.
diff --git a/tests/tree_sum.v b/tests/tree_sum.v
index bc216d653..0f3839220 100644
--- a/tests/tree_sum.v
+++ b/tests/tree_sum.v
@@ -64,6 +64,3 @@ Proof.
   rewrite Z.add_0_r.
   iIntros "Hl Ht". wp_seq. wp_load. by iApply "HΦ".
 Qed.
-
-Global Opaque sum_loop sum'.
-
-- 
GitLab