diff --git a/barrier/barrier.v b/barrier/barrier.v index fd34e8f9b14bb3294637bad55eaa991bad59776a..65c0af7aa966a11a4e26501665ea5df95ab63953 100644 --- a/barrier/barrier.v +++ b/barrier/barrier.v @@ -210,7 +210,7 @@ Section proof. (* I think some evars here are better than repeating *everything* *) eapply (sts_fsaS _ (wp_fsa _)) with (N0:=N) (γ0:=γ); simpl; eauto with I ndisj. - rewrite [(_ ★ sts_ownS _ _ _)%I]comm -!assoc /wp_fsa. apply sep_mono_r. + rewrite [(_ ★ sts_ownS _ _ _)%I]comm -!assoc. apply sep_mono_r. apply forall_intro=>-[p I]. apply wand_intro_l. rewrite -!assoc. apply const_elim_sep_l=>Hs. destruct p; last done. rewrite {1}/barrier_inv =>/={Hs}. rewrite later_sep. @@ -251,7 +251,7 @@ Section proof. (* I think some evars here are better than repeating *everything* *) eapply (sts_fsaS _ (wp_fsa _)) with (N0:=N) (γ0:=γ); simpl; eauto with I ndisj. - rewrite !assoc [(_ ★ sts_ownS _ _ _)%I]comm -!assoc /wp_fsa. apply sep_mono_r. + rewrite !assoc [(_ ★ sts_ownS _ _ _)%I]comm -!assoc. apply sep_mono_r. apply forall_intro=>-[p I]. apply wand_intro_l. rewrite -!assoc. apply const_elim_sep_l=>Hs. rewrite {1}/barrier_inv =>/=. rewrite later_sep. diff --git a/heap_lang/heap.v b/heap_lang/heap.v index 69596f9b642d3e81f5c2969f44b63f3eb5dede0e..9314bf11ff003900fe529a90b389dc2817d1f656 100644 --- a/heap_lang/heap.v +++ b/heap_lang/heap.v @@ -109,8 +109,7 @@ Section heap. with N heap_name ∅; simpl; eauto with I. rewrite -later_intro. apply sep_mono_r,forall_intro=> h; apply wand_intro_l. rewrite -assoc left_id; apply const_elim_sep_l=> ?. - rewrite {1}[(▷ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs. - rewrite /wp_fsa -(wp_alloc_pst _ (of_heap h)) //. + rewrite -(wp_alloc_pst _ (of_heap h)) //. apply sep_mono_r; rewrite HP; apply later_mono. apply forall_mono=> l; apply wand_intro_l. rewrite always_and_sep_l -assoc; apply const_elim_sep_l=> ?. @@ -134,7 +133,6 @@ Section heap. with N heap_name {[ l := Excl v ]}; simpl; eauto with I. rewrite HPΦ{HPΦ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l. rewrite -assoc; apply const_elim_sep_l=> ?. - rewrite {1}[(▷ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs. rewrite -(wp_load_pst _ (<[l:=v]>(of_heap h))) ?lookup_insert //. rewrite const_equiv // left_id. rewrite -(map_insert_singleton_op h); last by eapply heap_singleton_inv_l. @@ -153,7 +151,6 @@ Section heap. with N heap_name {[ l := Excl v' ]}; simpl; eauto with I. rewrite HPΦ{HPΦ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l. rewrite -assoc; apply const_elim_sep_l=> ?. - rewrite {1}[(▷ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs. rewrite -(wp_store_pst _ (<[l:=v']>(of_heap h))) ?lookup_insert //. rewrite /heap_inv alter_singleton insert_insert. rewrite -!(map_insert_singleton_op h); try by eapply heap_singleton_inv_l. @@ -174,7 +171,6 @@ Section heap. with N heap_name {[ l := Excl v' ]}; simpl; eauto 10 with I. rewrite HPΦ{HPΦ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l. rewrite -assoc; apply const_elim_sep_l=> ?. - rewrite {1}[(▷ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs. rewrite -(wp_cas_fail_pst _ (<[l:=v']>(of_heap h))) ?lookup_insert //. rewrite const_equiv // left_id. rewrite -(map_insert_singleton_op h); last by eapply heap_singleton_inv_l. @@ -194,7 +190,6 @@ Section heap. with N heap_name {[ l := Excl v1 ]}; simpl; eauto 10 with I. rewrite HPΦ{HPΦ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l. rewrite -assoc; apply const_elim_sep_l=> ?. - rewrite {1}[(▷ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs. rewrite -(wp_cas_suc_pst _ (<[l:=v1]>(of_heap h))) ?lookup_insert //. rewrite /heap_inv alter_singleton insert_insert. rewrite -!(map_insert_singleton_op h); try by eapply heap_singleton_inv_l. diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v index 5359b62ba1a0990e528c9c4e7e4a4836710291ae..276d910510cc886ad39227ac2d0789b41365a752 100644 --- a/heap_lang/lifting.v +++ b/heap_lang/lifting.v @@ -22,7 +22,7 @@ Proof. apply weakestpre.wp_bind. Qed. (** Base axioms for core primitives of the language: Stateful reductions. *) Lemma wp_alloc_pst E σ e v Φ : to_val e = Some v → - (ownP σ ★ ▷ (∀ l, σ !! l = None ∧ ownP (<[l:=v]>σ) -★ Φ (LocV l))) + (▷ ownP σ ★ ▷ (∀ l, σ !! l = None ∧ ownP (<[l:=v]>σ) -★ Φ (LocV l))) ⊑ || Alloc e @ E {{ Φ }}. Proof. (* TODO RJ: This works around ssreflect bug #22. *) @@ -40,7 +40,7 @@ Qed. Lemma wp_load_pst E σ l v Φ : σ !! l = Some v → - (ownP σ ★ ▷ (ownP σ -★ Φ v)) ⊑ || Load (Loc l) @ E {{ Φ }}. + (▷ ownP σ ★ ▷ (ownP σ -★ Φ v)) ⊑ || Load (Loc l) @ E {{ Φ }}. Proof. intros. rewrite -(wp_lift_atomic_det_step σ v σ None) ?right_id //; last by intros; inv_step; eauto using to_of_val. @@ -48,7 +48,7 @@ Qed. Lemma wp_store_pst E σ l e v v' Φ : to_val e = Some v → σ !! l = Some v' → - (ownP σ ★ ▷ (ownP (<[l:=v]>σ) -★ Φ (LitV LitUnit))) + (▷ ownP σ ★ ▷ (ownP (<[l:=v]>σ) -★ Φ (LitV LitUnit))) ⊑ || Store (Loc l) e @ E {{ Φ }}. Proof. intros. rewrite -(wp_lift_atomic_det_step σ (LitV LitUnit) (<[l:=v]>σ) None) @@ -57,7 +57,7 @@ Qed. Lemma wp_cas_fail_pst E σ l e1 v1 e2 v2 v' Φ : to_val e1 = Some v1 → to_val e2 = Some v2 → σ !! l = Some v' → v' ≠v1 → - (ownP σ ★ ▷ (ownP σ -★ Φ (LitV $ LitBool false))) + (▷ ownP σ ★ ▷ (ownP σ -★ Φ (LitV $ LitBool false))) ⊑ || Cas (Loc l) e1 e2 @ E {{ Φ }}. Proof. intros. rewrite -(wp_lift_atomic_det_step σ (LitV $ LitBool false) σ None) @@ -66,7 +66,7 @@ Qed. Lemma wp_cas_suc_pst E σ l e1 v1 e2 v2 Φ : to_val e1 = Some v1 → to_val e2 = Some v2 → σ !! l = Some v1 → - (ownP σ ★ ▷ (ownP (<[l:=v2]>σ) -★ Φ (LitV $ LitBool true))) + (▷ ownP σ ★ ▷ (ownP (<[l:=v2]>σ) -★ Φ (LitV $ LitBool true))) ⊑ || Cas (Loc l) e1 e2 @ E {{ Φ }}. Proof. intros. rewrite -(wp_lift_atomic_det_step σ (LitV $ LitBool true) diff --git a/program_logic/hoare_lifting.v b/program_logic/hoare_lifting.v index cda1417822f8691af3df00fb0d59526dc2faa6c5..b5c38de20820c96d92a7b03e5585e84ad8347e7e 100644 --- a/program_logic/hoare_lifting.v +++ b/program_logic/hoare_lifting.v @@ -22,7 +22,7 @@ Lemma ht_lift_step E1 E2 E1 ⊆ E2 → to_val e1 = None → reducible e1 σ1 → (∀ e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef → φ e2 σ2 ef) → - ((P ={E2,E1}=> ownP σ1 ★ ▷ P') ∧ ∀ e2 σ2 ef, + ((P ={E2,E1}=> ▷ ownP σ1 ★ ▷ P') ∧ ∀ e2 σ2 ef, (■φ e2 σ2 ef ★ ownP σ2 ★ P' ={E1,E2}=> Φ1 e2 σ2 ef ★ Φ2 e2 σ2 ef) ∧ {{ Φ1 e2 σ2 ef }} e2 @ E2 {{ Ψ }} ∧ {{ Φ2 e2 σ2 ef }} ef ?@ ⊤ {{ λ _, True }}) @@ -54,7 +54,7 @@ Lemma ht_lift_atomic_step reducible e1 σ1 → (∀ e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef → φ e2 σ2 ef) → (∀ e2 σ2 ef, {{ ■φ e2 σ2 ef ★ P }} ef ?@ ⊤ {{ λ _, True }}) ⊑ - {{ ownP σ1 ★ ▷ P }} e1 @ E {{ λ v, ∃ σ2 ef, ownP σ2 ★ ■φ (of_val v) σ2 ef }}. + {{ ▷ ownP σ1 ★ ▷ P }} e1 @ E {{ λ v, ∃ σ2 ef, ownP σ2 ★ ■φ (of_val v) σ2 ef }}. Proof. intros ? Hsafe Hstep; set (φ' e σ ef := is_Some (to_val e) ∧ φ e σ ef). rewrite -(ht_lift_step E E φ' _ P diff --git a/program_logic/lifting.v b/program_logic/lifting.v index 77cd393d0259e98ada8f51bfeec4be4fc6e0648d..ada372ac10f4b2b28a9094d33fd143cb915e9763 100644 --- a/program_logic/lifting.v +++ b/program_logic/lifting.v @@ -23,7 +23,7 @@ Lemma wp_lift_step E1 E2 E1 ⊆ E2 → to_val e1 = None → reducible e1 σ1 → (∀ e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef → φ e2 σ2 ef) → - (|={E2,E1}=> ownP σ1 ★ ▷ ∀ e2 σ2 ef, + (|={E2,E1}=> ▷ ownP σ1 ★ ▷ ∀ e2 σ2 ef, (■φ e2 σ2 ef ∧ ownP σ2) -★ |={E1,E2}=> || e2 @ E2 {{ Φ }} ★ wp_fork ef) ⊑ || e1 @ E2 {{ Φ }}. Proof. @@ -31,7 +31,7 @@ Proof. intros rf k Ef σ1' ???; destruct (Hvs rf (S k) Ef σ1') as (r'&(r1&r2&?&?&Hwp)&Hws); auto; clear Hvs; cofe_subst r'. destruct (wsat_update_pst k (E1 ∪ Ef) σ1 σ1' r1 (r2 ⋅ rf)) as [-> Hws']. - { by apply ownP_spec; auto. } + { apply equiv_dist. rewrite -(ownP_spec k); auto. } { by rewrite assoc. } constructor; [done|intros e2 σ2 ef ?; specialize (Hws' σ2)]. destruct (λ H1 H2 H3, Hwp e2 σ2 ef k (update_pst σ2 r1) H1 H2 H3 rf k Ef σ2) @@ -64,7 +64,7 @@ Lemma wp_lift_atomic_step {E Φ} e1 reducible e1 σ1 → (∀ e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef → ∃ v2, to_val e2 = Some v2 ∧ φ v2 σ2 ef) → - (ownP σ1 ★ ▷ ∀ v2 σ2 ef, ■φ v2 σ2 ef ∧ ownP σ2 -★ Φ v2 ★ wp_fork ef) + (▷ ownP σ1 ★ ▷ ∀ v2 σ2 ef, ■φ v2 σ2 ef ∧ ownP σ2 -★ Φ v2 ★ wp_fork ef) ⊑ || e1 @ E {{ Φ }}. Proof. intros. rewrite -(wp_lift_step E E (λ e2 σ2 ef, ∃ v2, @@ -84,7 +84,7 @@ Lemma wp_lift_atomic_det_step {E Φ e1} σ1 v2 σ2 ef : reducible e1 σ1 → (∀ e2' σ2' ef', prim_step e1 σ1 e2' σ2' ef' → σ2 = σ2' ∧ to_val e2' = Some v2 ∧ ef = ef') → - (ownP σ1 ★ ▷ (ownP σ2 -★ Φ v2 ★ wp_fork ef)) ⊑ || e1 @ E {{ Φ }}. + (▷ ownP σ1 ★ ▷ (ownP σ2 -★ Φ v2 ★ wp_fork ef)) ⊑ || e1 @ E {{ Φ }}. Proof. intros. rewrite -(wp_lift_atomic_step _ (λ v2' σ2' ef', σ2 = σ2' ∧ v2 = v2' ∧ ef = ef') σ1) //; last naive_solver. diff --git a/program_logic/ownership.v b/program_logic/ownership.v index 2cbb11caef8aceb0b98da3358866c469c8920777..ce69c6dff8e94af692e1c21995f8934c8fc566de 100644 --- a/program_logic/ownership.v +++ b/program_logic/ownership.v @@ -71,7 +71,7 @@ Global Instance ownG_unit_always_stable m : AlwaysStable (ownG (unit m)). Proof. by rewrite /AlwaysStable always_ownG_unit. Qed. (* inversion lemmas *) -Lemma ownI_spec r n i P : +Lemma ownI_spec n r i P : ✓{n} r → (ownI i P) n r ↔ wld r !! i ≡{n}≡ Some (to_agree (Next (iProp_unfold P))). Proof. @@ -81,12 +81,12 @@ Proof. (cmra_included_includedN _ P'),HP; apply map_lookup_validN with (wld r) i. - intros ?; split_and?; try apply cmra_empty_leastN; eauto. Qed. -Lemma ownP_spec r n σ : ✓{n} r → (ownP σ) n r ↔ pst r ≡{n}≡ Excl σ. +Lemma ownP_spec n r σ : ✓{n} r → (ownP σ) n r ↔ pst r ≡ Excl σ. Proof. intros (?&?&?); rewrite /uPred_holds /= res_includedN /= Excl_includedN //. - naive_solver (apply cmra_empty_leastN). + rewrite (timeless_iff n). naive_solver (apply cmra_empty_leastN). Qed. -Lemma ownG_spec r n m : (ownG m) n r ↔ Some m ≼{n} gst r. +Lemma ownG_spec n r m : (ownG m) n r ↔ Some m ≼{n} gst r. Proof. rewrite /uPred_holds /= res_includedN; naive_solver (apply cmra_empty_leastN). Qed. diff --git a/program_logic/weakestpre.v b/program_logic/weakestpre.v index 28f4f2ac7352672af8f34d624c4656a7c7069536..4ff71923da3025016bd98fc5b60e3cc8673b3a87 100644 --- a/program_logic/weakestpre.v +++ b/program_logic/weakestpre.v @@ -247,6 +247,7 @@ Proof. auto using wp_mask_frame_mono. Qed. (** * Weakest-pre is a FSA. *) Definition wp_fsa (e : expr Λ) : FSA Λ Σ (val Λ) := λ E, wp E e. +Global Arguments wp_fsa _ _ / _. Global Instance wp_fsa_prf : FrameShiftAssertion (atomic e) (wp_fsa e). Proof. rewrite /wp_fsa; split; auto using wp_mask_frame_mono, wp_atomic, wp_frame_r.