diff --git a/algebra/upred.v b/algebra/upred.v index 3910a66c6902472c8fbd5b40275010de1631e1f1..f2a8a55e548fa13af8b63a7882ca378d0f066d54 100644 --- a/algebra/upred.v +++ b/algebra/upred.v @@ -615,6 +615,10 @@ Lemma sep_intro_True_l P Q R : True ⊑ P → R ⊑ Q → R ⊑ (P ★ Q). Proof. by intros; rewrite -(left_id True%I uPred_sep R); apply sep_mono. Qed. Lemma sep_intro_True_r P Q R : R ⊑ P → True ⊑ Q → R ⊑ (P ★ Q). Proof. by intros; rewrite -(right_id True%I uPred_sep R); apply sep_mono. Qed. +Lemma sep_elim_True_l P Q R : True ⊑ P → (P ★ R) ⊑ Q → R ⊑ Q. +Proof. by intros HP; rewrite -HP left_id. Qed. +Lemma sep_elim_True_r P Q R : True ⊑ P → (R ★ P) ⊑ Q → R ⊑ Q. +Proof. by intros HP; rewrite -HP right_id. Qed. Lemma wand_intro_l P Q R : (Q ★ P) ⊑ R → P ⊑ (Q -★ R). Proof. rewrite comm; apply wand_intro_r. Qed. Lemma wand_elim_r P Q : (P ★ (P -★ Q)) ⊑ Q. diff --git a/heap_lang/heap.v b/heap_lang/heap.v index 5a85ee3b6f4b7d372db492c1bf564ab14cef82ec..3f589d43f2f07ccabff1d6ce3f5e033b050bea6c 100644 --- a/heap_lang/heap.v +++ b/heap_lang/heap.v @@ -119,8 +119,10 @@ Section heap. P ⊑ (▷ (∀ l, heap_mapsto HeapI γ l v -★ Q (LocV l))) → P ⊑ wp E (Alloc e) Q. Proof. - intros HN ? Hctx HP. rewrite -(right_id True%I (★)%I (P)) (auth_empty γ) pvs_frame_l. - apply wp_strip_pvs. eapply wp_alloc_heap with (σ:=∅); try eassumption. + intros HN ? Hctx HP. eapply sep_elim_True_r. + { eapply (auth_empty γ). } + rewrite pvs_frame_l. apply wp_strip_pvs. + eapply wp_alloc_heap with (σ:=∅); try eassumption. { eauto with I. } rewrite HP comm. apply sep_mono. - rewrite /heap_own. f_equiv. apply: map_eq=>l. by rewrite lookup_fmap !lookup_empty. @@ -166,7 +168,8 @@ Section heap. P ⊑ wp E (Store (Loc l) e) Q. Proof. rewrite /heap_ctx /heap_own. intros Hl Hval HN Hctx HP. - eapply (auth_fsa (heap_inv HeapI) (wp_fsa _) (λ _:(), alter (λ _, Excl v) l)); simpl; eauto. + eapply (auth_fsa (heap_inv HeapI) (wp_fsa _) + (λ _:(), alter (λ _, Excl v) l)); simpl; eauto. rewrite HP=>{HP Hctx HN}. apply sep_mono; first done. apply forall_intro=>hf. apply wand_intro_l. rewrite /heap_inv. rewrite -assoc. apply const_elim_sep_l=>Hv /=. @@ -189,12 +192,13 @@ Section heap. + by rewrite lookup_alter /to_heap !lookup_fmap lookup_insert Hl /=. + rewrite lookup_alter_ne // !lookup_fmap lookup_insert_ne //. } rewrite !EQ. apply sep_mono; last done. - f_equiv. apply: map_eq=>l'. move: (Hv 0%nat l'). destruct (decide (l=l')); simplify_map_equality. + f_equiv. apply: map_eq=>l'. move: (Hv 0%nat l'). + destruct (decide (l=l')); simplify_map_equality. - rewrite /from_heap /to_heap lookup_insert lookup_omap !lookup_op. rewrite !lookup_fmap lookup_insert Hl. case (hf !! l')=>[[?||]|]; auto; contradiction. - - rewrite /from_heap /to_heap lookup_insert_ne // !lookup_omap !lookup_op !lookup_fmap. - rewrite lookup_insert_ne //. + - rewrite /from_heap /to_heap lookup_insert_ne // !lookup_omap. + rewrite !lookup_op !lookup_fmap lookup_insert_ne //. Qed. Lemma wp_store N E γ l v' e v P Q : @@ -204,7 +208,8 @@ Section heap. P ⊑ (heap_mapsto HeapI γ l v' ★ ▷ (heap_mapsto HeapI γ l v -★ Q (LitV LitUnit))) → P ⊑ wp E (Store (Loc l) e) Q. Proof. - rewrite /heap_mapsto=>Hval HN Hctx HP. eapply wp_store_heap; try eassumption; last first. + rewrite /heap_mapsto=>Hval HN Hctx HP. + eapply wp_store_heap; try eassumption; last first. - rewrite HP. apply sep_mono; first done. by rewrite insert_singleton. - by rewrite lookup_insert. Qed. diff --git a/program_logic/auth.v b/program_logic/auth.v index cb5f4b8245201aac5a9bb6741ed05e4b2ede85f1..31e35703742df120fc56846e1d0eccb3d3ffb53d 100644 --- a/program_logic/auth.v +++ b/program_logic/auth.v @@ -32,8 +32,8 @@ Section auth. Lemma auth_alloc N a : ✓ a → φ a ⊑ pvs N N (∃ γ, auth_ctx AuthI γ N φ ∧ auth_own AuthI γ a). Proof. - intros Ha. rewrite -(right_id True%I (★)%I (φ _)). - rewrite (own_alloc AuthI (Auth (Excl a) a) N) //; []. + intros Ha. eapply sep_elim_True_r. + { by eapply (own_alloc AuthI (Auth (Excl a) a) N). } rewrite pvs_frame_l. apply pvs_strip_pvs. rewrite sep_exist_l. apply exist_elim=>γ. rewrite -(exist_intro γ). transitivity (▷ auth_inv AuthI γ φ ★ auth_own AuthI γ a)%I.