diff --git a/barrier/barrier.v b/barrier/barrier.v index 08a059318f9f2f6ffda360c11d486203d860e692..75a35c82a3f185778e79076b5b422b6cbad26952 100644 --- a/barrier/barrier.v +++ b/barrier/barrier.v @@ -1,10 +1,10 @@ From heap_lang Require Export substitution notation. -Definition newchan : val := λ: "", ref '0. +Definition newbarrier : val := λ: "", ref '0. Definition signal : val := λ: "x", "x" <- '1. Definition wait : val := rec: "wait" "x" := if: !"x" = '1 then '() else "wait" "x". -Instance newchan_closed : Closed newchan. Proof. solve_closed. Qed. +Instance newbarrier_closed : Closed newbarrier. Proof. solve_closed. Qed. Instance signal_closed : Closed signal. Proof. solve_closed. Qed. -Instance wait_closed : Closed wait. Proof. solve_closed. Qed. \ No newline at end of file +Instance wait_closed : Closed wait. Proof. solve_closed. Qed. diff --git a/barrier/client.v b/barrier/client.v index b54796f95ef56bde2b1bc57645bbcdcc80265ed1..a4183e64b93cc1bdf6ad2d7c34d9ce30a8fb0d68 100644 --- a/barrier/client.v +++ b/barrier/client.v @@ -2,7 +2,7 @@ From barrier Require Import proof. From program_logic Require Import auth sts saved_prop hoare ownership. Import uPred. -Definition client := (let: "b" := newchan '() in wait "b")%L. +Definition client := (let: "b" := newbarrier '() in wait "b")%L. Section client. Context {Σ : iFunctorG} `{!heapG Σ, !barrierG Σ} (heapN N : namespace). @@ -12,7 +12,7 @@ Section client. heapN ⊥ N → heap_ctx heapN ⊑ || client {{ λ _, True }}. Proof. intros ?. rewrite /client. - ewp eapply (newchan_spec heapN N True%I); last done. + ewp eapply (newbarrier_spec heapN N True%I); last done. apply sep_intro_True_r; first done. apply forall_intro=>l. apply wand_intro_l. rewrite right_id. wp_let. etrans; last eapply wait_spec. diff --git a/barrier/proof.v b/barrier/proof.v index 2cb4a39f858e6a39953e66d18c9aee12e1970cb8..4e128f74c07f9c6d5610f2ac2a003ed1037a6a1c 100644 --- a/barrier/proof.v +++ b/barrier/proof.v @@ -118,12 +118,12 @@ Proof. Qed. (** Actual proofs *) -Lemma newchan_spec (P : iProp) (Φ : val → iProp) : +Lemma newbarrier_spec (P : iProp) (Φ : val → iProp) : heapN ⊥ N → (heap_ctx heapN ★ ∀ l, recv l P ★ send l P -★ Φ (LocV l)) - ⊑ || newchan '() {{ Φ }}. + ⊑ || newbarrier '() {{ Φ }}. Proof. - intros HN. rewrite /newchan. wp_seq. + intros HN. rewrite /newbarrier. wp_seq. rewrite -wp_pvs. wp eapply wp_alloc; eauto with I ndisj. apply forall_intro=>l. rewrite (forall_elim l). apply wand_intro_l. rewrite !assoc. apply pvs_wand_r. diff --git a/barrier/specification.v b/barrier/specification.v index 3bd06c46da4e084951ab21ba60c8ae4a610c2ea2..82d1d95d1cffb201293314acfc4e33007c863492 100644 --- a/barrier/specification.v +++ b/barrier/specification.v @@ -12,7 +12,7 @@ Local Notation iProp := (iPropG heap_lang Σ). Lemma barrier_spec (heapN N : namespace) : heapN ⊥ N → ∃ recv send : loc → iProp -n> iProp, - (∀ P, heap_ctx heapN ⊑ {{ True }} newchan '() {{ λ v, ∃ l, v = LocV l ★ recv l P ★ send l P }}) ∧ + (∀ P, heap_ctx heapN ⊑ {{ True }} newbarrier '() {{ λ v, ∃ l, v = LocV l ★ recv l P ★ send l P }}) ∧ (∀ l P, {{ send l P ★ P }} signal (LocV l) {{ λ _, True }}) ∧ (∀ l P, {{ recv l P }} wait (LocV l) {{ λ _, P }}) ∧ (∀ l P Q, {{ recv l (P ★ Q) }} Skip {{ λ _, recv l P ★ recv l Q }}) ∧ @@ -22,7 +22,7 @@ Proof. exists (λ l, CofeMor (recv heapN N l)), (λ l, CofeMor (send heapN N l)). split_and?; simpl. - intros P. apply: always_intro. apply impl_intro_r. - rewrite -(newchan_spec heapN N P) // always_and_sep_r. + rewrite -(newbarrier_spec heapN N P) // always_and_sep_r. apply sep_mono_r, forall_intro=>l; apply wand_intro_l. by rewrite right_id -(exist_intro l) const_equiv // left_id. - intros l P. apply ht_alt. by rewrite -signal_spec right_id.