diff --git a/theories/logrel/F_mu_ref_conc/examples/counter.v b/theories/logrel/F_mu_ref_conc/examples/counter.v index 6e70b2a7e4ad3203d6004c6d07ebd9f7ee7cc5a9..2c43b6ba248dede26b054637d3505e932a346443 100644 --- a/theories/logrel/F_mu_ref_conc/examples/counter.v +++ b/theories/logrel/F_mu_ref_conc/examples/counter.v @@ -165,7 +165,7 @@ Section CG_Counter. Proof. intros HNE. iIntros "[#Hspec [Hx Hj]]". unfold counter_read. iMod (step_rec _ _ j K _ Unit with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_load _ _ j K with "[Hj Hx]") as "[Hj Hx]"; eauto. { by iFrame "Hspec Hj". } iModIntro. by iFrame "Hj Hx". @@ -264,12 +264,12 @@ Section CG_Counter. iMod (steps_newlock _ _ j ((AppRCtx (RecV _)) :: K) _ with "[Hj]") as (l) "[Hj Hl]"; eauto. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. rewrite CG_locked_increment_subst /=. + iAsimpl. rewrite CG_locked_increment_subst /=. rewrite counter_read_subst /=. iMod (step_alloc _ _ j ((AppRCtx (RecV _)) :: K) _ _ _ _ with "[Hj]") as (cnt') "[Hj Hcnt']"; eauto. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. rewrite CG_locked_increment_subst /=. + iAsimpl. rewrite CG_locked_increment_subst /=. rewrite counter_read_subst /=. Unshelve. all: try match goal with |- to_val _ = _ => auto using to_of_val end. @@ -284,7 +284,7 @@ Section CG_Counter. iApply fupd_wp. iMod (inv_alloc counterN with "[Hinv]") as "#Hinv"; [iNext; iExact "Hinv"|]. (* splitting increment and read *) - iApply wp_pure_step_later; trivial. iModIntro. iNext. asimpl. + iApply wp_pure_step_later; trivial. iModIntro. iNext. iAsimpl. rewrite counter_read_subst /=. iApply wp_value; auto. iExists (PairV (CG_locked_incrementV _ _) (counter_readV _)); simpl. @@ -296,7 +296,7 @@ Section CG_Counter. rewrite CG_locked_increment_of_val /=. destruct v; iDestruct "Heq" as "[% %]"; simplify_eq/=. iLöb as "Hlat". - iApply wp_pure_step_later; trivial. asimpl. iNext. + iApply wp_pure_step_later; trivial. iAsimpl. iNext. (* fine-grained reads the counter *) iApply (wp_bind (fill [AppRCtx (RecV _)])); iApply wp_wand_l; iSplitR; [iIntros (v) "Hv"; iExact "Hv"|]. @@ -306,7 +306,7 @@ Section CG_Counter. iModIntro. iNext. iIntros "Hcnt". iMod ("Hclose" with "[Hl Hcnt Hcnt']"). { iNext. iExists _. iFrame "Hl Hcnt Hcnt'". } - iApply wp_pure_step_later; trivial. asimpl. iModIntro. iNext. + iApply wp_pure_step_later; trivial. iAsimpl. iModIntro. iNext. (* fine-grained performs increment *) iApply (wp_bind (fill [CasRCtx (LocV _) (NatV _); IfCtx _ _])); iApply wp_wand_l; iSplitR; [iIntros (v) "Hv"; iExact "Hv"|]. diff --git a/theories/logrel/F_mu_ref_conc/examples/lock.v b/theories/logrel/F_mu_ref_conc/examples/lock.v index b5f675d1a9878f8ed6a7579b4f78c4586187c8ac..6b93469422a8cfb13806da5199e4c31a7f15c9e6 100644 --- a/theories/logrel/F_mu_ref_conc/examples/lock.v +++ b/theories/logrel/F_mu_ref_conc/examples/lock.v @@ -135,24 +135,24 @@ Section proof. Proof. iIntros (HNE H1 H2) "[#Hspec [HP [Hl Hj]]]". iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. rewrite H1. + iAsimpl. rewrite H1. iMod (steps_acquire _ _ j ((AppRCtx (RecV _)) :: K) _ _ with "[Hj Hl]") as "[Hj Hl]"; eauto. { simpl. iFrame "Hspec Hj Hl"; eauto. } simpl. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. rewrite H1. + iAsimpl. rewrite H1. iMod (H2 ((AppRCtx (RecV _)) :: K) with "[Hj HP]") as "[Hj HQ]"; eauto. { simpl. iFrame "Hspec Hj HP"; eauto. } simpl. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (steps_release _ _ j ((AppRCtx (RecV _)) :: K) _ _ with "[Hj Hl]") as "[Hj Hl]"; eauto. { simpl. by iFrame. } rewrite ?fill_app /=. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. iModIntro; by iFrame. + iAsimpl. iModIntro; by iFrame. Unshelve. all: try match goal with |- to_val _ = _ => auto using to_of_val end. trivial. diff --git a/theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v b/theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v index 268515832ec9d193eb55ba1f15128de88224e15f..3aacfaf6820d65f1e3d24f5735a91f4b4c6983e6 100644 --- a/theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v +++ b/theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v @@ -83,7 +83,7 @@ Section CG_Stack. Proof. intros HNE. iIntros "[#Hspec [Hx Hj]]". unfold CG_push. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_load _ _ j (PairRCtx _ :: InjRCtx :: FoldCtx :: StoreRCtx (LocV _) :: K) _ _ _ with "[Hj Hx]") as "[Hj Hx]"; eauto. simpl. iFrame "Hspec Hj"; trivial. simpl. @@ -177,7 +177,7 @@ Section CG_Stack. Proof. intros HNE. iIntros "[#Hspec [Hx Hj]]". unfold CG_pop. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_load _ _ j (UnfoldCtx :: CaseCtx _ _ :: K) _ _ _ with "[Hj Hx]") as "[Hj Hx]"; eauto. rewrite ?fill_app. simpl. @@ -187,7 +187,7 @@ Section CG_Stack. _ _ _ _ with "[Hj]") as "Hj"; eauto. simpl. iMod (step_case_inr _ _ j K _ _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_snd _ _ j (StoreRCtx (LocV _) :: AppRCtx (RecV _) :: K) _ _ _ _ _ _ with "[Hj]") as "Hj"; eauto. simpl. @@ -197,7 +197,7 @@ Section CG_Stack. iFrame "Hspec Hj"; trivial. rewrite ?fill_app. simpl. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_fst _ _ j (InjRCtx :: K) _ _ _ _ _ _ with "[Hj]") as "Hj"; eauto. simpl. @@ -215,14 +215,14 @@ Section CG_Stack. Proof. iIntros (HNE) "[#Hspec [Hx Hj]]". unfold CG_pop. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_load _ _ j (UnfoldCtx :: CaseCtx _ _ :: K) _ _ _ with "[Hj Hx]") as "[Hj Hx]"; eauto. simpl. iFrame "Hspec Hj"; trivial. simpl. iMod (step_Fold _ _ j (CaseCtx _ _ :: K) _ _ _ _ with "[Hj]") as "Hj"; eauto. iMod (step_case_inl _ _ j K _ _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iModIntro. iFrame "Hj Hx"; trivial. Unshelve. all: try match goal with |- to_val _ = _ => simpl; by rewrite ?to_of_val end. @@ -336,7 +336,7 @@ Section CG_Stack. with "[Hj Hx Hl]") as "Hj"; last done; [|iFrame; iFrame "#"]. iIntros (K') "[#Hspec [Hx Hj]]". iMod (step_rec _ _ j K' _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_load _ _ j K' _ _ _ _ with "[Hj Hx]") as "[Hj Hx]"; eauto. - iFrame "#"; iFrame. @@ -405,12 +405,12 @@ Section CG_Stack. Proof. iIntros (HNE) "[#Hspec Hj]". unfold CG_iter. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - rewrite -CG_iter_folding. Opaque CG_iter. asimpl. + rewrite -CG_iter_folding. Opaque CG_iter. iAsimpl. iMod (step_Fold _ _ j (CaseCtx _ _ :: K) _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_case_inr _ _ j K _ _ _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_fst _ _ j (AppRCtx f :: AppRCtx (RecV _) :: K) _ _ _ _ _ _ with "[Hj]") as "Hj"; eauto. Unshelve. @@ -426,10 +426,10 @@ Section CG_Stack. Proof. iIntros (HNE) "[#Hspec Hj]". unfold CG_iter. iMod (step_rec _ _ j K _ _ _ _ with "[Hj]") as "Hj"; eauto. - rewrite -CG_iter_folding. Opaque CG_iter. asimpl. + rewrite -CG_iter_folding. Opaque CG_iter. iAsimpl. iMod (step_Fold _ _ j (CaseCtx _ _ :: K) _ _ _ with "[Hj]") as "Hj"; eauto. - asimpl. + iAsimpl. iMod (step_case_inl _ _ j K _ _ _ _ _ with "[Hj]") as "Hj"; eauto. Unshelve. all: try match goal with |- to_val _ = _ => simpl; by rewrite ?to_of_val end. diff --git a/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v b/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v index ea90117c39c4177832016f4c05d039e6e8fbe084..30fb1b77f3773347d68379f5ee06221c91a742e8 100644 --- a/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v +++ b/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v @@ -8,18 +8,6 @@ From iris.proofmode Require Import tactics. Definition stackN : namespace := nroot .@ "stack". -Ltac iAsimpl := - repeat match goal with - | |- context [ (_ ⤇ ?e)%I ] => progress ( - let e' := fresh "feed" in evar (e':expr); - assert (e = e') as ->; [asimpl; unfold e'; reflexivity|]; - unfold e'; clear e') - | |- context [ WP ?e @ _ {{ _ }}%I ] => progress ( - let e' := fresh "feed" in evar (e':expr); - assert (e = e') as ->; [asimpl; unfold e'; reflexivity|]; - unfold e'; clear e') - end. - Section Stack_refinement. Context `{heapIG Σ, cfgSG Σ, inG Σ (authR stackUR)}. Notation D := (prodC valC valC -n> iProp Σ). diff --git a/theories/logrel/F_mu_ref_conc/rules_binary.v b/theories/logrel/F_mu_ref_conc/rules_binary.v index 0289c98cc7b136af07ebf062a9ad0a9571c7a81c..f37ca0bf5206c13e4fff4e3a9e013745e95c3d1d 100644 --- a/theories/logrel/F_mu_ref_conc/rules_binary.v +++ b/theories/logrel/F_mu_ref_conc/rules_binary.v @@ -48,6 +48,18 @@ Notation "l ↦ₛ{ q } v" := (heapS_mapsto l q v) Notation "l ↦ₛ v" := (heapS_mapsto l 1 v) (at level 20) : uPred_scope. Notation "j ⤇ e" := (tpool_mapsto j e) (at level 20) : uPred_scope. +Ltac iAsimpl := + repeat match goal with + | |- context [ (_ ⤇ ?e)%I ] => progress ( + let e' := fresh in evar (e':expr); + assert (e = e') as ->; [asimpl; unfold e'; reflexivity|]; + unfold e'; clear e') + | |- context [ WP ?e @ _ {{ _ }}%I ] => progress ( + let e' := fresh in evar (e':expr); + assert (e = e') as ->; [asimpl; unfold e'; reflexivity|]; + unfold e'; clear e') + end. + Section conversions. Context `{cfgSG Σ}. diff --git a/theories/logrel/F_mu_ref_conc/soundness_binary.v b/theories/logrel/F_mu_ref_conc/soundness_binary.v index 471dd8ea52e51e5d325c8ec277cb27078075faef..70e50a65d7c0fa55ce318bb74cb63bb077cae6b6 100644 --- a/theories/logrel/F_mu_ref_conc/soundness_binary.v +++ b/theories/logrel/F_mu_ref_conc/soundness_binary.v @@ -31,7 +31,7 @@ Proof. { iApply (@logrel_binary.interp_env_nil Σ HeapΣ). } simpl. rewrite empty_env_subst empty_env_subst. iApply ("Hrel" $! 0 []). - { rewrite /tpool_mapsto. asimpl. by iFrame. } + { rewrite /tpool_mapsto. iAsimpl. by iFrame. } iModIntro. iIntros (v1); iDestruct 1 as (v2) "[Hj #Hinterp]". iInv specN as (tp σ) ">[Hown Hsteps]" "Hclose"; iDestruct "Hsteps" as %Hsteps'. rewrite /tpool_mapsto /=.