diff --git a/heap_lang/lib/assert.v b/heap_lang/lib/assert.v index fa9b8cd257b635200a8e54644faecd74c775e0c5..511460f281cd053009262df410b36f115fd80b18 100644 --- a/heap_lang/lib/assert.v +++ b/heap_lang/lib/assert.v @@ -9,7 +9,7 @@ Definition assert : val := 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 {{ Φ }}. + WP e @ E {{ v, ⌜v = #true⌠∧ â–· Φ #() }} -∗ WP assert: e @ E {{ Φ }}. Proof. iIntros "HΦ". rewrite /assert. wp_let. wp_seq. iApply (wp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if. diff --git a/heap_lang/lib/barrier/proof.v b/heap_lang/lib/barrier/proof.v index 7beca266c9367b7c6a8ef77d91e93411cab6e04f..ee7d8d10c932c0c7d1d244b69707b2106be46208 100644 --- a/heap_lang/lib/barrier/proof.v +++ b/heap_lang/lib/barrier/proof.v @@ -73,11 +73,11 @@ Proof. solve_proper. Qed. (** Helper lemmas *) Lemma ress_split i i1 i2 Q R1 R2 P I : i ∈ I → i1 ∉ I → i2 ∉ I → i1 ≠i2 → - saved_prop_own i Q ∗ saved_prop_own i1 R1 ∗ saved_prop_own i2 R2 ∗ - (Q -∗ R1 ∗ R2) ∗ ress P I - ⊢ ress P ({[i1;i2]} ∪ I ∖ {[i]}). + saved_prop_own i Q -∗ saved_prop_own i1 R1 -∗ saved_prop_own i2 R2 -∗ + (Q -∗ R1 ∗ R2) -∗ ress P I -∗ + ress P ({[i1;i2]} ∪ I ∖ {[i]}). Proof. - iIntros (????) "(#HQ&#H1&#H2&HQR&H)"; iDestruct "H" as (Ψ) "[HPΨ HΨ]". + iIntros (????) "#HQ #H1 #H2 HQR"; iDestruct 1 as (Ψ) "[HPΨ HΨ]". iDestruct (big_sepS_delete _ _ i with "HΨ") as "[#HΨi HΨ]"; first done. iExists (<[i1:=R1]> (<[i2:=R2]> Ψ)). iSplitL "HQR HPΨ". - iPoseProof (saved_prop_agree i Q (Ψ i) with "[#]") as "Heq"; first by iSplit. @@ -175,7 +175,7 @@ Proof. {[Change i1; Change i2 ]} with "[-]") as "Hγ". { iSplit; first by eauto using split_step. rewrite {2}/barrier_inv /=. iNext. iFrame "Hl". - iApply (ress_split _ _ _ Q R1 R2); eauto. iFrame; auto. } + by iApply (ress_split with "HQ Hi1 Hi2 HQR"). } iAssert (sts_ownS γ (i_states i1) {[Change i1]} ∗ sts_ownS γ (i_states i2) {[Change i2]})%I with ">[-]" as "[Hγ1 Hγ2]". { iApply sts_ownS_op; eauto using i_states_closed, low_states_closed. @@ -190,8 +190,7 @@ Qed. Lemma recv_weaken l P1 P2 : (P1 -∗ P2) -∗ recv l P1 -∗ recv l P2. Proof. - rewrite /recv. - iIntros "HP HP1"; iDestruct "HP1" as (γ P Q i) "(#Hctx&Hγ&Hi&HP1)". + rewrite /recv. iIntros "HP". iDestruct 1 as (γ P Q i) "(#Hctx&Hγ&Hi&HP1)". iExists γ, P, Q, i. iFrame "Hctx Hγ Hi". iNext. iIntros "HQ". by iApply "HP"; iApply "HP1". Qed. diff --git a/heap_lang/lib/barrier/specification.v b/heap_lang/lib/barrier/specification.v index 7e34c5d858dd3c05fff548a64640b93e1a7e09c9..a733cb1e10e6a0806448e5278aac0068379d317d 100644 --- a/heap_lang/lib/barrier/specification.v +++ b/heap_lang/lib/barrier/specification.v @@ -14,7 +14,7 @@ Lemma barrier_spec (N : namespace) : (∀ l P, {{ send l P ∗ P }} signal #l {{ _, True }}) ∧ (∀ l P, {{ recv l P }} wait #l {{ _, P }}) ∧ (∀ l P Q, recv l (P ∗ Q) ={↑N}=> recv l P ∗ recv l Q) ∧ - (∀ l P Q, (P -∗ Q) ⊢ recv l P -∗ recv l Q). + (∀ l P Q, (P -∗ Q) -∗ recv l P -∗ recv l Q). Proof. exists (λ l, CofeMor (recv N l)), (λ l, CofeMor (send N l)). split_and?; simpl. diff --git a/heap_lang/lib/lock.v b/heap_lang/lib/lock.v index ba10b131213c847cca5032a88bd1b2c41a12f834..e9d37f8ab0fc405759253601fb9de74459e32bbc 100644 --- a/heap_lang/lib/lock.v +++ b/heap_lang/lib/lock.v @@ -15,7 +15,7 @@ Structure lock Σ `{!heapG Σ} := Lock { is_lock_ne N γ lk n: Proper (dist n ==> dist n) (is_lock N γ lk); is_lock_persistent N γ lk R : PersistentP (is_lock N γ lk R); locked_timeless γ : TimelessP (locked γ); - locked_exclusive γ : locked γ ∗ locked γ ⊢ False; + locked_exclusive γ : locked γ -∗ locked γ -∗ False; (* -- operation specs -- *) newlock_spec N (R : iProp Σ) : {{{ R }}} newlock #() {{{ lk γ, RET lk; is_lock N γ lk R }}}; diff --git a/heap_lang/lib/par.v b/heap_lang/lib/par.v index 7124875995026745bba7661a593e4cf158fda3a4..2fcad149b3ce6198c9a4fe01ff6ebabf7a296fd1 100644 --- a/heap_lang/lib/par.v +++ b/heap_lang/lib/par.v @@ -21,11 +21,11 @@ Context `{!heapG Σ, !spawnG Σ}. This is why these are not Texan triples. *) Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp Σ) : to_val e = Some (f1,f2)%V → - (WP f1 #() {{ Ψ1 }} ∗ WP f2 #() {{ Ψ2 }} ∗ - â–· ∀ v1 v2, Ψ1 v1 ∗ Ψ2 v2 -∗ â–· Φ (v1,v2)%V) - ⊢ WP par e {{ Φ }}. + WP f1 #() {{ Ψ1 }} -∗ WP f2 #() {{ Ψ2 }} -∗ + (â–· ∀ v1 v2, Ψ1 v1 ∗ Ψ2 v2 -∗ â–· Φ (v1,v2)%V) -∗ + WP par e {{ Φ }}. Proof. - iIntros (?) "(Hf1 & Hf2 & HΦ)". + iIntros (?) "Hf1 Hf2 HΦ". rewrite /par. wp_value. wp_let. wp_proj. wp_apply (spawn_spec parN with "Hf1"); try wp_done; try solve_ndisj. iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _). @@ -36,11 +36,11 @@ Qed. Lemma wp_par (Ψ1 Ψ2 : val → iProp Σ) (e1 e2 : expr) `{!Closed [] e1, Closed [] e2} (Φ : val → iProp Σ) : - (WP e1 {{ Ψ1 }} ∗ WP e2 {{ Ψ2 }} ∗ - ∀ v1 v2, Ψ1 v1 ∗ Ψ2 v2 -∗ â–· Φ (v1,v2)%V) - ⊢ WP e1 ||| e2 {{ Φ }}. + WP e1 {{ Ψ1 }} -∗ WP e2 {{ Ψ2 }} -∗ + (∀ v1 v2, Ψ1 v1 ∗ Ψ2 v2 -∗ â–· Φ (v1,v2)%V) -∗ + WP e1 ||| e2 {{ Φ }}. Proof. - iIntros "(H1 & H2 & H)". iApply (par_spec Ψ1 Ψ2 with "[- $H]"); try wp_done. - iSplitL "H1"; by wp_let. + iIntros "H1 H2 H". iApply (par_spec Ψ1 Ψ2 with "[H1] [H2] [H]"); try wp_done. + by wp_let. by wp_let. auto. Qed. End proof. diff --git a/heap_lang/lib/spin_lock.v b/heap_lang/lib/spin_lock.v index 78e5504c793749d25cc22c7ab8f04145ef74cff5..77a7dbbaa16d56830a44c18a39c57e053cbf79fb 100644 --- a/heap_lang/lib/spin_lock.v +++ b/heap_lang/lib/spin_lock.v @@ -30,8 +30,8 @@ Section proof. Definition locked (γ : gname): iProp Σ := own γ (Excl ()). - Lemma locked_exclusive (γ : gname) : locked γ ∗ locked γ ⊢ False. - Proof. rewrite /locked -own_op own_valid. by iIntros (?). Qed. + Lemma locked_exclusive (γ : gname) : locked γ -∗ locked γ -∗ False. + Proof. iIntros "H1 H2". by iDestruct (own_valid_2 with "H1 H2") as %?. Qed. Global Instance lock_inv_ne n γ l : Proper (dist n ==> dist n) (lock_inv γ l). Proof. solve_proper. Qed. diff --git a/heap_lang/lib/ticket_lock.v b/heap_lang/lib/ticket_lock.v index 386121291b69c6853906df52eeac8afd692bb2b6..477a28442946089bc8ba37292949537839283448 100644 --- a/heap_lang/lib/ticket_lock.v +++ b/heap_lang/lib/ticket_lock.v @@ -46,11 +46,11 @@ Section proof. Definition is_lock (γ : gname) (lk : val) (R : iProp Σ) : iProp Σ := (∃ lo ln : loc, - ⌜lk = (#lo, #ln)%V⌠∧ inv N (lock_inv γ lo ln R))%I. + ⌜lk = (#lo, #ln)%V⌠∗ inv N (lock_inv γ lo ln R))%I. Definition issued (γ : gname) (lk : val) (x : nat) (R : iProp Σ) : iProp Σ := (∃ lo ln: loc, - ⌜lk = (#lo, #ln)%V⌠∧ inv N (lock_inv γ lo ln R) ∧ + ⌜lk = (#lo, #ln)%V⌠∗ inv N (lock_inv γ lo ln R) ∗ own γ (â—¯ (∅, GSet {[ x ]})))%I. Definition locked (γ : gname) : iProp Σ := (∃ o, own γ (â—¯ (Excl' o, ∅)))%I. @@ -65,10 +65,10 @@ Section proof. Global Instance locked_timeless γ : TimelessP (locked γ). Proof. apply _. Qed. - Lemma locked_exclusive (γ : gname) : (locked γ ∗ locked γ ⊢ False)%I. + Lemma locked_exclusive (γ : gname) : locked γ -∗ locked γ -∗ False. Proof. - iIntros "[H1 H2]". iDestruct "H1" as (o1) "H1". iDestruct "H2" as (o2) "H2". - iCombine "H1" "H2" as "H". iDestruct (own_valid with "H") as %[[] _]. + iDestruct 1 as (o1) "H1". iDestruct 1 as (o2) "H2". + iDestruct (own_valid_2 with "H1 H2") as %[[] _]. Qed. Lemma newlock_spec (R : iProp Σ) : diff --git a/tests/barrier_client.v b/tests/barrier_client.v index e531b726662dc6a943620668afd3b3d912126e93..d27cddd6a1241f5f463ee5b580f5eb840e537cc2 100644 --- a/tests/barrier_client.v +++ b/tests/barrier_client.v @@ -18,39 +18,38 @@ Section client. Definition y_inv (q : Qp) (l : loc) : iProp Σ := (∃ f : val, l ↦{q} f ∗ â–¡ ∀ n : Z, WP f #n {{ v, ⌜v = #(n + 42)⌠}})%I. - Lemma y_inv_split q l : y_inv q l ⊢ (y_inv (q/2) l ∗ y_inv (q/2) l). + Lemma y_inv_split q l : y_inv q l -∗ (y_inv (q/2) l ∗ y_inv (q/2) l). Proof. iDestruct 1 as (f) "[[Hl1 Hl2] #Hf]". iSplitL "Hl1"; iExists f; by iSplitL; try iAlways. Qed. Lemma worker_safe q (n : Z) (b y : loc) : - recv N b (y_inv q y) ⊢ WP worker n #b #y {{ _, True }}. + recv N b (y_inv q y) -∗ WP worker n #b #y {{ _, True }}. Proof. iIntros "Hrecv". wp_lam. wp_let. - wp_apply (wait_spec with "[- $Hrecv]"). iDestruct 1 as (f) "[Hy #Hf]". + wp_apply (wait_spec with "Hrecv"). iDestruct 1 as (f) "[Hy #Hf]". wp_seq. wp_load. iApply (wp_wand with "[]"). iApply "Hf". by iIntros (v) "_". Qed. - Lemma client_safe : True ⊢ WP client {{ _, True }}. + Lemma client_safe : WP client {{ _, True }}%I. Proof. iIntros ""; rewrite /client. wp_alloc y as "Hy". wp_let. wp_apply (newbarrier_spec N (y_inv 1 y)). iIntros (l) "[Hr Hs]". wp_let. - iApply (wp_par (λ _, True%I) (λ _, True%I)). iSplitL "Hy Hs". + iApply (wp_par (λ _, True%I) (λ _, True%I) with "[Hy Hs] [Hr]"); last auto. - (* The original thread, the sender. *) wp_store. iApply (signal_spec with "[-]"); last by iNext; auto. iSplitR "Hy"; first by eauto. iExists _; iSplitL; [done|]. iAlways; iIntros (n). wp_let. by wp_op. - (* The two spawned threads, the waiters. *) - iSplitL; [|by iIntros (_ _) "_ !>"]. iDestruct (recv_weaken with "[] Hr") as "Hr". { iIntros "Hy". by iApply (y_inv_split with "Hy"). } iMod (recv_split with "Hr") as "[H1 H2]"; first done. - iApply (wp_par (λ _, True%I) (λ _, True%I)). - iSplitL "H1"; [|iSplitL "H2"; [|by iIntros (_ _) "_ !>"]]; - by iApply worker_safe. + iApply (wp_par (λ _, True%I) (λ _, True%I) with "[H1] [H2]"); last auto. + + by iApply worker_safe. + + by iApply worker_safe. Qed. End client. @@ -60,8 +59,7 @@ Let Σ : gFunctors := #[ heapΣ ; barrierΣ ; spawnΣ ]. Lemma client_adequate σ : adequate client σ (λ _, True). Proof. - apply (heap_adequacy Σ)=> ?. - apply (client_safe (nroot .@ "barrier")); auto with ndisj. + apply (heap_adequacy Σ)=> ?. apply (client_safe (nroot .@ "barrier")). Qed. End ClosedProofs. diff --git a/tests/heap_lang.v b/tests/heap_lang.v index 4ca9412d44cb20725e5c2eb9f6e4413a6179b0b9..8e3c473c7cb22e0f4f8ec159d9703d44a1a973e2 100644 --- a/tests/heap_lang.v +++ b/tests/heap_lang.v @@ -12,7 +12,7 @@ Section LiftingTests. Definition heap_e : expr := let: "x" := ref #1 in "x" <- !"x" + #1 ;; !"x". - Lemma heap_e_spec E : True ⊢ WP heap_e @ E {{ v, ⌜v = #2⌠}}. + Lemma heap_e_spec E : WP heap_e @ E {{ v, ⌜v = #2⌠}}%I. Proof. iIntros "". rewrite /heap_e. wp_alloc l. wp_let. wp_load. wp_op. wp_store. by wp_load. @@ -23,7 +23,7 @@ Section LiftingTests. let: "y" := ref #1 in "x" <- !"x" + #1 ;; !"x". - Lemma heap_e2_spec E : True ⊢ WP heap_e2 @ E {{ v, ⌜v = #2⌠}}. + Lemma heap_e2_spec E : WP heap_e2 @ E {{ v, ⌜v = #2⌠}}%I. Proof. iIntros "". rewrite /heap_e2. wp_alloc l. wp_let. wp_alloc l'. wp_let. @@ -40,7 +40,7 @@ Section LiftingTests. Lemma FindPred_spec n1 n2 E Φ : n1 < n2 → - Φ #(n2 - 1) ⊢ WP FindPred #n2 #n1 @ E {{ Φ }}. + Φ #(n2 - 1) -∗ WP FindPred #n2 #n1 @ E {{ Φ }}. Proof. iIntros (Hn) "HΦ". iLöb as "IH" forall (n1 Hn). wp_rec. wp_let. wp_op. wp_let. wp_op=> ?; wp_if. @@ -48,7 +48,7 @@ Section LiftingTests. - by assert (n1 = n2 - 1) as -> by omega. Qed. - Lemma Pred_spec n E Φ : â–· Φ #(n - 1) ⊢ WP Pred #n @ E {{ Φ }}. + Lemma Pred_spec n E Φ : â–· Φ #(n - 1) -∗ WP Pred #n @ E {{ Φ }}. Proof. iIntros "HΦ". wp_lam. wp_op=> ?; wp_if. - wp_op. wp_op. @@ -62,5 +62,5 @@ Section LiftingTests. Proof. iIntros "". wp_apply Pred_spec. wp_let. by wp_apply Pred_spec. Qed. End LiftingTests. -Lemma heap_e_adequate σ : adequate heap_e σ (λ v, v = #2). +Lemma heap_e_adequate σ : adequate heap_e σ (= #2). Proof. eapply (heap_adequacy heapΣ)=> ?. by apply heap_e_spec. Qed. diff --git a/tests/joining_existentials.v b/tests/joining_existentials.v index 9d93e11f6190eea4805297ca678ed2a27f095b4b..ecf84402c7ea78f1c3f478f1e207165b88995893 100644 --- a/tests/joining_existentials.v +++ b/tests/joining_existentials.v @@ -31,29 +31,28 @@ Definition barrier_res γ (Φ : X → iProp Σ) : iProp Σ := (∃ x, own γ (Shot x) ∗ Φ x)%I. Lemma worker_spec e γ l (Φ Ψ : X → iProp Σ) `{!Closed [] e} : - recv N l (barrier_res γ Φ) ∗ (∀ x, {{ Φ x }} e {{ _, Ψ x }}) - ⊢ WP wait #l ;; e {{ _, barrier_res γ Ψ }}. + recv N l (barrier_res γ Φ) -∗ (∀ x, {{ Φ x }} e {{ _, Ψ x }}) -∗ + WP wait #l ;; e {{ _, barrier_res γ Ψ }}. Proof. - iIntros "[Hl #He]". wp_apply (wait_spec with "[- $Hl]"); simpl. + iIntros "Hl #He". wp_apply (wait_spec with "[- $Hl]"); simpl. iDestruct 1 as (x) "[#Hγ Hx]". wp_seq. iApply (wp_wand with "[Hx]"); [by iApply "He"|]. iIntros (v) "?"; iExists x; by iSplit. Qed. Context (P : iProp Σ) (Φ Φ1 Φ2 Ψ Ψ1 Ψ2 : X -n> iProp Σ). -Context {Φ_split : ∀ x, Φ x ⊢ (Φ1 x ∗ Φ2 x)}. -Context {Ψ_join : ∀ x, (Ψ1 x ∗ Ψ2 x) ⊢ Ψ x}. +Context {Φ_split : ∀ x, Φ x -∗ (Φ1 x ∗ Φ2 x)}. +Context {Ψ_join : ∀ x, Ψ1 x -∗ Ψ2 x -∗ Ψ x}. -Lemma P_res_split γ : barrier_res γ Φ ⊢ barrier_res γ Φ1 ∗ barrier_res γ Φ2. +Lemma P_res_split γ : barrier_res γ Φ -∗ barrier_res γ Φ1 ∗ barrier_res γ Φ2. Proof. iDestruct 1 as (x) "[#Hγ Hx]". iDestruct (Φ_split with "Hx") as "[H1 H2]". by iSplitL "H1"; iExists x; iSplit. Qed. -Lemma Q_res_join γ : barrier_res γ Ψ1 ∗ barrier_res γ Ψ2 ⊢ â–· barrier_res γ Ψ. +Lemma Q_res_join γ : barrier_res γ Ψ1 -∗ barrier_res γ Ψ2 -∗ â–· barrier_res γ Ψ. Proof. - iIntros "[Hγ Hγ']"; - iDestruct "Hγ" as (x) "[#Hγ Hx]"; iDestruct "Hγ'" as (x') "[#Hγ' Hx']". + iDestruct 1 as (x) "[#Hγ Hx]"; iDestruct 1 as (x') "[#Hγ' Hx']". iAssert (â–· (x ≡ x'))%I as "Hxx". { iCombine "Hγ" "Hγ'" as "Hγ2". iClear "Hγ Hγ'". rewrite own_valid csum_validI /= agree_validI agree_equivI uPred.later_equivI /=. @@ -62,23 +61,22 @@ Proof. { by split; intro; simpl; symmetry; apply iProp_fold_unfold. } rewrite !cFunctor_compose. iNext. by iRewrite "Hγ2". } iNext. iRewrite -"Hxx" in "Hx'". - iExists x; iFrame "Hγ". iApply Ψ_join; by iSplitL "Hx". + iExists x; iFrame "Hγ". iApply (Ψ_join with "Hx Hx'"). Qed. Lemma client_spec_new eM eW1 eW2 `{!Closed [] eM, !Closed [] eW1, !Closed [] eW2} : - P - ∗ {{ P }} eM {{ _, ∃ x, Φ x }} - ∗ (∀ x, {{ Φ1 x }} eW1 {{ _, Ψ1 x }}) - ∗ (∀ x, {{ Φ2 x }} eW2 {{ _, Ψ2 x }}) - ⊢ WP client eM eW1 eW2 {{ _, ∃ γ, barrier_res γ Ψ }}. + P -∗ + {{ P }} eM {{ _, ∃ x, Φ x }} -∗ + (∀ x, {{ Φ1 x }} eW1 {{ _, Ψ1 x }}) -∗ + (∀ x, {{ Φ2 x }} eW2 {{ _, Ψ2 x }}) -∗ + WP client eM eW1 eW2 {{ _, ∃ γ, barrier_res γ Ψ }}. Proof. - iIntros "/= (HP & #He & #He1 & #He2)"; rewrite /client. + iIntros "/= HP #He #He1 #He2"; rewrite /client. iMod (own_alloc (Pending : one_shotR Σ F)) as (γ) "Hγ"; first done. wp_apply (newbarrier_spec N (barrier_res γ Φ)); auto. iIntros (l) "[Hr Hs]". set (workers_post (v : val) := (barrier_res γ Ψ1 ∗ barrier_res γ Ψ2)%I). - wp_let. wp_apply (wp_par (λ _, True)%I workers_post). - iSplitL "HP Hs Hγ"; [|iSplitL "Hr"]. + wp_let. wp_apply (wp_par (λ _, True)%I workers_post with "[HP Hs Hγ] [Hr]"). - wp_bind eM. iApply (wp_wand with "[HP]"); [by iApply "He"|]. iIntros (v) "HP"; iDestruct "HP" as (x) "HP". wp_let. iMod (own_update with "Hγ") as "Hx". @@ -87,11 +85,11 @@ Proof. iExists x; auto. - iDestruct (recv_weaken with "[] Hr") as "Hr"; first by iApply P_res_split. iMod (recv_split with "Hr") as "[H1 H2]"; first done. - wp_apply (wp_par (λ _, barrier_res γ Ψ1)%I (λ _, barrier_res γ Ψ2)%I). - iSplitL "H1"; [|iSplitL "H2"]. - + iApply worker_spec; auto. - + iApply worker_spec; auto. + wp_apply (wp_par (λ _, barrier_res γ Ψ1)%I + (λ _, barrier_res γ Ψ2)%I with "[H1] [H2]"). + + iApply (worker_spec with "H1"); auto. + + iApply (worker_spec with "H2"); auto. + auto. - - iIntros (_ v) "[_ H]". iDestruct (Q_res_join with "H") as "?". auto. + - iIntros (_ v) "[_ [H1 H2]]". iDestruct (Q_res_join with "H1 H2") as "?". auto. Qed. End proof. diff --git a/tests/list_reverse.v b/tests/list_reverse.v index 45a6b9c39c3714dfdcebbb255d2cba4db028028f..77846047e054131907e71de989ed2cac47e9a3e7 100644 --- a/tests/list_reverse.v +++ b/tests/list_reverse.v @@ -26,11 +26,11 @@ Definition rev : val := end. Lemma rev_acc_wp hd acc xs ys (Φ : val → iProp Σ) : - is_list hd xs ∗ is_list acc ys ∗ - (∀ w, is_list w (reverse xs ++ ys) -∗ Φ w) - ⊢ WP rev hd acc {{ Φ }}. + is_list hd xs -∗ is_list acc ys -∗ + (∀ w, is_list w (reverse xs ++ ys) -∗ Φ w) -∗ + WP rev hd acc {{ Φ }}. Proof. - iIntros "(Hxs & Hys & HΦ)". + iIntros "Hxs Hys HΦ". iLöb as "IH" forall (hd acc xs ys Φ). wp_rec. wp_let. destruct xs as [|x xs]; iSimplifyEq. - wp_match. by iApply "HΦ". @@ -42,11 +42,11 @@ Proof. Qed. Lemma rev_wp hd xs (Φ : val → iProp Σ) : - is_list hd xs ∗ (∀ w, is_list w (reverse xs) -∗ Φ w) - ⊢ WP rev hd (InjL #()) {{ Φ }}. + is_list hd xs -∗ (∀ w, is_list w (reverse xs) -∗ Φ w) -∗ + WP rev hd (InjL #()) {{ Φ }}. Proof. - iIntros "[Hxs HΦ]". - iApply (rev_acc_wp hd NONEV xs [] with "[- $Hxs]"). - iSplit; first done. iIntros (w). rewrite right_id_L. iApply "HΦ". + iIntros "Hxs HΦ". + iApply (rev_acc_wp hd NONEV xs [] with "Hxs [%]")=> //. + iIntros (w). rewrite right_id_L. iApply "HΦ". Qed. End list_reverse. diff --git a/tests/proofmode.v b/tests/proofmode.v index df03e0f98087dafdf1f5008dbbb0188f1e85a676..833b8fdc1f88cfb0b2e870089be3e35f4e58f35a 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -2,7 +2,7 @@ From iris.proofmode Require Import tactics. From iris.base_logic.lib Require Import invariants. Lemma demo_0 {M : ucmraT} (P Q : uPred M) : - â–¡ (P ∨ Q) ⊢ (∀ x, ⌜x = 0⌠∨ ⌜x = 1âŒ) → (Q ∨ P). + â–¡ (P ∨ Q) -∗ (∀ x, ⌜x = 0⌠∨ ⌜x = 1âŒ) → (Q ∨ P). Proof. iIntros "#H #H2". (* should remove the disjunction "H" *) @@ -37,8 +37,9 @@ Proof. Qed. Lemma demo_2 (M : ucmraT) (P1 P2 P3 P4 Q : uPred M) (P5 : nat → uPredC M): - P2 ∗ (P3 ∗ Q) ∗ True ∗ P1 ∗ P2 ∗ (P4 ∗ (∃ x:nat, P5 x ∨ P3)) ∗ True - ⊢ P1 -∗ (True ∗ True) -∗ (((P2 ∧ False ∨ P2 ∧ ⌜0 = 0âŒ) ∗ P3) ∗ Q ∗ P1 ∗ True) ∧ + P2 ∗ (P3 ∗ Q) ∗ True ∗ P1 ∗ P2 ∗ (P4 ∗ (∃ x:nat, P5 x ∨ P3)) ∗ True -∗ + P1 -∗ (True ∗ True) -∗ + (((P2 ∧ False ∨ P2 ∧ ⌜0 = 0âŒ) ∗ P3) ∗ Q ∗ P1 ∗ True) ∧ (P2 ∨ False) ∧ (False → P5 0). Proof. (* Intro-patterns do something :) *) @@ -54,17 +55,17 @@ Proof. Qed. Lemma demo_3 (M : ucmraT) (P1 P2 P3 : uPred M) : - P1 ∗ P2 ∗ P3 ⊢ â–· P1 ∗ â–· (P2 ∗ ∃ x, (P3 ∧ ⌜x = 0âŒ) ∨ P3). + P1 ∗ P2 ∗ P3 -∗ â–· P1 ∗ â–· (P2 ∗ ∃ x, (P3 ∧ ⌜x = 0âŒ) ∨ P3). Proof. iIntros "($ & $ & H)". iFrame "H". iNext. by iExists 0. Qed. Definition foo {M} (P : uPred M) := (P → P)%I. Definition bar {M} : uPred M := (∀ P, foo P)%I. -Lemma demo_4 (M : ucmraT) : True ⊢ @bar M. +Lemma demo_4 (M : ucmraT) : True -∗ @bar M. Proof. iIntros. iIntros (P) "HP". done. Qed. Lemma demo_5 (M : ucmraT) (x y : M) (P : uPred M) : - (∀ z, P → z ≡ y) ⊢ (P -∗ (x,x) ≡ (y,x)). + (∀ z, P → z ≡ y) -∗ (P -∗ (x,x) ≡ (y,x)). Proof. iIntros "H1 H2". iRewrite (uPred.internal_eq_sym x x with "[#]"); first done. @@ -73,15 +74,15 @@ Proof. Qed. Lemma demo_6 (M : ucmraT) (P Q : uPred M) : - True ⊢ ∀ x y z : nat, - ⌜x = plus 0 x⌠→ ⌜y = 0⌠→ ⌜z = 0⌠→ P → â–¡ Q → foo (x ≡ x). + (∀ x y z : nat, + ⌜x = plus 0 x⌠→ ⌜y = 0⌠→ ⌜z = 0⌠→ P → â–¡ Q → foo (x ≡ x))%I. Proof. iIntros (a) "*". iIntros "#Hfoo **". by iIntros "# _". Qed. -Lemma demo_7 (M : ucmraT) (P Q1 Q2 : uPred M) : P ∗ (Q1 ∧ Q2) ⊢ P ∗ Q1. +Lemma demo_7 (M : ucmraT) (P Q1 Q2 : uPred M) : P ∗ (Q1 ∧ Q2) -∗ P ∗ Q1. Proof. iIntros "[H1 [H2 _]]". by iFrame. Qed. Section iris. @@ -91,7 +92,7 @@ Section iris. Lemma demo_8 N E P Q R : ↑N ⊆ E → - (True -∗ P -∗ inv N Q -∗ True -∗ R) ⊢ P -∗ â–· Q ={E}=∗ R. + (True -∗ P -∗ inv N Q -∗ True -∗ R) -∗ P -∗ â–· Q ={E}=∗ R. Proof. iIntros (?) "H HP HQ". iApply ("H" with "[#] HP >[HQ] >"). @@ -102,5 +103,5 @@ Section iris. End iris. Lemma demo_9 (M : ucmraT) (x y z : M) : - ✓ x → ⌜y ≡ z⌠⊢ (✓ x ∧ ✓ x ∧ y ≡ z : uPred M). + ✓ x → ⌜y ≡ z⌠-∗ (✓ x ∧ ✓ x ∧ y ≡ z : uPred M). Proof. iIntros (Hv) "Hxy". by iFrame (Hv Hv) "Hxy". Qed. diff --git a/tests/tree_sum.v b/tests/tree_sum.v index c5a82d666de2c1e8e82c8281d31bd28b2d38a365..ebca3662f90260bd0416c5ad814aa957d75c76da 100644 --- a/tests/tree_sum.v +++ b/tests/tree_sum.v @@ -34,10 +34,10 @@ Definition sum' : val := λ: "t", !"l". Lemma sum_loop_wp `{!heapG Σ} v t l (n : Z) (Φ : val → iProp Σ) : - l ↦ #n ∗ is_tree v t ∗ (l ↦ #(sum t + n) -∗ is_tree v t -∗ Φ #()) - ⊢ WP sum_loop v #l {{ Φ }}. + l ↦ #n -∗ is_tree v t -∗ (l ↦ #(sum t + n) -∗ is_tree v t -∗ Φ #()) -∗ + WP sum_loop v #l {{ Φ }}. Proof. - iIntros "(Hl & Ht & HΦ)". + iIntros "Hl Ht HΦ". iLöb as "IH" forall (v t l n Φ). wp_rec. wp_let. destruct t as [n'|tl tr]; simpl in *. - iDestruct "Ht" as "%"; subst. @@ -54,11 +54,11 @@ Proof. Qed. Lemma sum_wp `{!heapG Σ} v t Φ : - is_tree v t ∗ (is_tree v t -∗ Φ #(sum t)) ⊢ WP sum' v {{ Φ }}. + is_tree v t -∗ (is_tree v t -∗ Φ #(sum t)) -∗ WP sum' v {{ Φ }}. Proof. - iIntros "[Ht HΦ]". rewrite /sum' /=. + iIntros "Ht HΦ". rewrite /sum' /=. wp_let. wp_alloc l as "Hl". wp_let. - wp_apply (sum_loop_wp with "[- $Ht $Hl]"). + wp_apply (sum_loop_wp with "Hl Ht"). rewrite Z.add_0_r. iIntros "Hl Ht". wp_seq. wp_load. by iApply "HΦ". Qed.