diff --git a/heap_lang/derived.v b/heap_lang/derived.v index 9a95ee5c896a723d668ce27a506ab672fcae09cb..367f34eaaa077df0ceb77cb5b2d3e7b52297feac 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 2c9969e7654bdb805af67fd06a0294cc15668989..fa9b8cd257b635200a8e54644faecd74c775e0c5 100644 --- a/heap_lang/lib/assert.v +++ b/heap_lang/lib/assert.v @@ -4,13 +4,13 @@ From iris.proofmode Require Import tactics. From iris.heap_lang Require Import proofmode notation. Definition assert : val := - locked (λ: "v", if: "v" #() then #() else #0 #0)%V. (* #0 #0 is unsafe *) + λ: "v", if: "v" #() then #() else #0 #0. (* #0 #0 is unsafe *) (* just below ;; *) Notation "'assert:' e" := (assert (λ: <>, e))%E (at level 99) : expr_scope. Lemma wp_assert `{heapG Σ} E (Φ : val → iProp Σ) e `{!Closed [] e} : WP e @ E {{ v, ⌜v = #true⌠∧ â–· Φ #() }} ⊢ WP assert: e @ E {{ Φ }}. Proof. - iIntros "HΦ". rewrite /assert -lock. wp_let. wp_seq. + iIntros "HΦ". rewrite /assert. wp_let. wp_seq. iApply (wp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if. Qed. diff --git a/heap_lang/lib/barrier/barrier.v b/heap_lang/lib/barrier/barrier.v index 78ba65bcb82a5362b94fc033f01255df1e8bec99..61f8f72524eb58e09a144bf3e9683b80fea76657 100644 --- a/heap_lang/lib/barrier/barrier.v +++ b/heap_lang/lib/barrier/barrier.v @@ -1,6 +1,6 @@ From iris.heap_lang Require Export notation. -Definition newbarrier : val := locked (λ: <>, ref #false)%V. -Definition signal : val := locked (λ: "x", "x" <- #true)%V. +Definition newbarrier : val := λ: <>, ref #false. +Definition signal : val := λ: "x", "x" <- #true. Definition wait : val := - locked (rec: "wait" "x" := if: !"x" then #() else "wait" "x")%V. + rec: "wait" "x" := if: !"x" then #() else "wait" "x". diff --git a/heap_lang/lib/barrier/proof.v b/heap_lang/lib/barrier/proof.v index 8b57fb7a71e43c9d5fffc60a84e5a347ab48fba5..2b65b88c7d2d736a7d60b52d142c86e17e0acbdf 100644 --- a/heap_lang/lib/barrier/proof.v +++ b/heap_lang/lib/barrier/proof.v @@ -95,7 +95,7 @@ Lemma newbarrier_spec (P : iProp Σ) : {{{ heap_ctx }}} newbarrier #() {{{ l, RET #l; recv l P ∗ send l P }}}. Proof. iIntros (HN Φ) "#? HΦ". - rewrite -wp_fupd /newbarrier -lock /=. wp_seq. wp_alloc l as "Hl". + rewrite -wp_fupd /newbarrier /=. wp_seq. wp_alloc l as "Hl". iApply ("HΦ" with ">[-]"). iMod (saved_prop_alloc (F:=idCF) P) as (γ) "#?". iMod (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]") @@ -119,7 +119,7 @@ Qed. Lemma signal_spec l P : {{{ send l P ∗ P }}} signal #l {{{ RET #(); True }}}. Proof. - rewrite /signal /send /barrier_ctx -lock /=. + rewrite /signal /send /barrier_ctx /=. iIntros (Φ) "(Hs&HP) HΦ"; iDestruct "Hs" as (γ) "[#(%&Hh&Hsts) Hγ]". wp_let. iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]") as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto. diff --git a/heap_lang/lib/counter.v b/heap_lang/lib/counter.v index 097bd7fb667f65e328c4c264466f9ca0f83c6915..72d16f8300607cf998e00ff036993724af1cec15 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 ea9a11aed2749ecaf0391c18990bb6a1cddb3bed..3c15c0cd533eb5e4c627b3aa1ffcd1d4c4c6c75e 100644 --- a/heap_lang/lib/par.v +++ b/heap_lang/lib/par.v @@ -5,11 +5,11 @@ Import uPred. Definition parN : namespace := nroot .@ "par". Definition par : val := - locked (λ: "fs", + λ: "fs", let: "handle" := spawn (Fst "fs") in let: "v2" := Snd "fs" #() in let: "v1" := join "handle" in - ("v1", "v2"))%V. + ("v1", "v2"). Notation "e1 ||| e2" := (par (Pair (λ: <>, e1) (λ: <>, e2)))%E : expr_scope. Section proof. @@ -26,7 +26,7 @@ Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp ⊢ WP par e {{ Φ }}. Proof. iIntros (?) "(#Hh&Hf1&Hf2&HΦ)". - rewrite /par -lock. wp_value. wp_let. wp_proj. + rewrite /par. wp_value. wp_let. wp_proj. wp_apply (spawn_spec parN with "[$Hh $Hf1]"); try wp_done; try solve_ndisj. iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _). iApply (wp_wand with "Hf2"); iIntros (v) "H2". wp_let. @@ -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 46bebd09a94c24c547a842d2265b06fee12863f7..20cbb59b4246f6d43a6b99aa5451d934fbdf74f0 100644 --- a/heap_lang/lib/spawn.v +++ b/heap_lang/lib/spawn.v @@ -5,15 +5,15 @@ From iris.heap_lang Require Import proofmode notation. From iris.algebra Require Import excl. Definition spawn : val := - locked (λ: "f", + λ: "f", let: "c" := ref NONE in - Fork ("c" <- SOME ("f" #())) ;; "c")%V. + Fork ("c" <- SOME ("f" #())) ;; "c". Definition join : val := - locked (rec: "join" "c" := + rec: "join" "c" := match: !"c" with SOME "x" => "x" | NONE => "join" "c" - end)%V. + end. (** The CMRA & functor we need. *) (* Not bundling heapG, as it may be shared with other users. *) @@ -50,7 +50,7 @@ Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) : heapN ⊥ N → {{{ heap_ctx ∗ WP f #() {{ Ψ }} }}} spawn e {{{ l, RET #l; join_handle l Ψ }}}. Proof. - iIntros (<-%of_to_val ? Φ) "(#Hh & Hf) HΦ". rewrite /spawn -lock /=. + iIntros (<-%of_to_val ? Φ) "(#Hh & Hf) HΦ". rewrite /spawn /=. wp_let. wp_alloc l as "Hl". wp_let. iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done. iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?". diff --git a/heap_lang/lib/spin_lock.v b/heap_lang/lib/spin_lock.v index f6d0e66fbc3edc7f8a73ce789e44d5ad3acd2248..e9bcad0c6674cf9e939615a7a49124267b7a14ea 100644 --- a/heap_lang/lib/spin_lock.v +++ b/heap_lang/lib/spin_lock.v @@ -5,11 +5,11 @@ From iris.heap_lang Require Import proofmode notation. From iris.algebra Require Import excl. From iris.heap_lang.lib Require Import lock. -Definition newlock : val := ssreflect.locked (λ: <>, ref #false)%V. -Definition try_acquire : val := ssreflect.locked (λ: "l", CAS "l" #false #true)%V. +Definition newlock : val := λ: <>, ref #false. +Definition try_acquire : val := λ: "l", CAS "l" #false #true. Definition acquire : val := - ssreflect.locked (rec: "acquire" "l" := if: try_acquire "l" then #() else "acquire" "l")%V. -Definition release : val := ssreflect.locked (λ: "l", "l" <- #false)%V. + rec: "acquire" "l" := if: try_acquire "l" then #() else "acquire" "l". +Definition release : val := λ: "l", "l" <- #false. (** The CMRA we need. *) (* Not bundling heapG, as it may be shared with other users. *) @@ -48,7 +48,7 @@ Section proof. heapN ⊥ N → {{{ heap_ctx ∗ R }}} newlock #() {{{ lk γ, RET lk; is_lock γ lk R }}}. Proof. - iIntros (? Φ) "[#Hh HR] HΦ". rewrite -wp_fupd /newlock. unlock. + iIntros (? Φ) "[#Hh HR] HΦ". rewrite -wp_fupd /newlock /=. wp_seq. wp_alloc l as "Hl". iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done. iMod (inv_alloc N _ (lock_inv γ l R) with "[-HΦ]") as "#?". @@ -83,7 +83,7 @@ Section proof. Proof. iIntros (Φ) "(Hlock & Hlocked & HR) HΦ". iDestruct "Hlock" as (l) "(% & #? & % & #?)". subst. - rewrite /release. unlock. wp_let. iInv N as (b) "[Hl _]" "Hclose". + rewrite /release /=. wp_let. iInv N as (b) "[Hl _]" "Hclose". wp_store. iApply "HΦ". iApply "Hclose". iNext. iExists false. by iFrame. Qed. End proof. diff --git a/heap_lang/lib/ticket_lock.v b/heap_lang/lib/ticket_lock.v index 8e16f24ece3feb0bf022cd93ab2b97c15032008d..e4c691e6bdf45e225bbb70d9300424ed55aa0253 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 7d45d55010816376fda543e59271b8d1ac3c1a8c..eb07aaa0b4cf6098df4bf2b2bc9627c857287836 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 040c5ae21dd112a6e83b96a04f3fc27e9ffc92a2..bd2bb6ff4cf2c9a3344136759983785d871f46ec 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 bc647d8e44c8726d4107761f3905b3db0713d988..65fff20b5bbeec940954bcb3045ac84071b5f477 100644 --- a/heap_lang/wp_tactics.v +++ b/heap_lang/wp_tactics.v @@ -44,13 +44,21 @@ Tactic Notation "wp_value" := | _ => fail "wp_value: not a wp" end. +Lemma of_val_unlock v e : of_val v = e → of_val (locked v) = e. +Proof. by unlock. Qed. + +(* 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 solve_of_val_unlock := try apply of_val_unlock; 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; [solve_of_val_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 +68,8 @@ 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; [solve_of_val_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 1800ad32cce040545e4fcbc44022634cf6552087..d9f173c180f0dabcb15ac7e8dc3bc0b0cc091cbd 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 4b852ad6d87237bfba33030ed24f7260f47c0032..a97e466e26e575398b00388fe2e1231692878290 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 1ba0ee98bf1d45f839cf552dfd0f766980fec9c8..60160c224cbf04d9404d613d54ea3e13b5d7efbc 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 7000179672758fafb898385ab0bb3f54c10f74e1..2a143a107b42abb8241c6e9ef4ba234a04057767 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 64565e93fa8e9047cf8fe087829c52ed5851227c..10cc5bf35b46470162ce0ac0734aa1f89c46f7e0 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 323f8896aec983c45a528aa2f90f8ff70e0f28f2..0f7b99c73d6626db07d8240972e63209f0676d69 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 bc216d6538ab7a62120fb5aaf23b54e6077c7ff0..0f38392204e601f68c806ec3e93f897827f03030 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'. -