Skip to content
Snippets Groups Projects
Commit bb978c7d authored by Ralf Jung's avatar Ralf Jung
Browse files

enable iMod to run an atomic update

make atomic_acc typeclasses-opaque
parent 29e3f8e6
No related branches found
No related tags found
No related merge requests found
...@@ -259,6 +259,22 @@ Section lemmas. ...@@ -259,6 +259,22 @@ Section lemmas.
iDestruct "HUpd" as (P) "(HP & Hshift)". by iApply "Hshift". iDestruct "HUpd" as (P) "(HP & Hshift)". by iApply "Hshift".
Qed. Qed.
(* This lets you eliminate atomic updates with iMod. *)
Global Instance elim_mod_aupd φ Eo Ei E α β Φ Q Q' :
( R, ElimModal φ false false (|={E,Ei}=> R) R Q Q')
ElimModal (φ Eo E) false false
(atomic_update Eo Ei α β Φ)
(.. x, α x
(α x ={Ei,E}=∗ atomic_update Eo Ei α β Φ)
(.. y, β x y ={Ei,E}=∗ Φ x y))
Q Q'.
Proof.
intros ?. rewrite /ElimModal /= =>-[??]. iIntros "[AU Hcont]".
iPoseProof (aupd_aacc with "AU") as "AC".
iMod (atomic_acc_mask_weaken with "AC"); first done.
iApply "Hcont". done.
Qed.
Global Instance aupd_laterable Eo Ei α β Φ : Global Instance aupd_laterable Eo Ei α β Φ :
Laterable (atomic_update Eo Ei α β Φ). Laterable (atomic_update Eo Ei α β Φ).
Proof. Proof.
...@@ -293,6 +309,7 @@ Section lemmas. ...@@ -293,6 +309,7 @@ Section lemmas.
- iIntros (y) "Hβ". iMod "Hclose'" as "_". iApply "Hclose". done. - iIntros (y) "Hβ". iMod "Hclose'" as "_". iApply "Hclose". done.
Qed. Qed.
(* This lets you open invariants etc. when the goal is an atomic accessor. *)
Global Instance elim_acc_aacc {X} E1 E2 Ei (α' β' : X PROP) γ' α β Pas Φ : Global Instance elim_acc_aacc {X} E1 E2 Ei (α' β' : X PROP) γ' α β Pas Φ :
ElimAcc (X:=X) (fupd E1 E2) (fupd E2 E1) α' β' γ' ElimAcc (X:=X) (fupd E1 E2) (fupd E2 E1) α' β' γ'
(atomic_acc E1 Ei α Pas β Φ) (atomic_acc E1 Ei α Pas β Φ)
...@@ -319,6 +336,16 @@ Section lemmas. ...@@ -319,6 +336,16 @@ Section lemmas.
iModIntro. destruct (γ' x'); iApply "HΦ"; done. iModIntro. destruct (γ' x'); iApply "HΦ"; done.
Qed. Qed.
(* Everything that fancy updates can eliminate without changing, atomic
accessors can eliminate as well. This is a forwarding instance needed becuase
atomic_acc is becoming opaque. *)
Global Instance elim_modal_acc p q φ P P' Eo Ei α Pas β Φ :
( Q, ElimModal φ p q P P' (|={Eo,Ei}=> Q) (|={Eo,Ei}=> Q))
ElimModal φ p q P P'
(atomic_acc Eo Ei α Pas β Φ)
(atomic_acc Eo Ei α Pas β Φ).
Proof. intros Helim. apply Helim. Qed.
Lemma aacc_aacc {TA' TB' : tele} E1 E1' E2 E3 Lemma aacc_aacc {TA' TB' : tele} E1 E1' E2 E3
α P β Φ α P β Φ
(α' : TA' PROP) P' (β' Φ' : TA' TB' PROP) : (α' : TA' PROP) P' (β' Φ' : TA' TB' PROP) :
...@@ -423,3 +450,6 @@ Tactic Notation "iAuIntro" := ...@@ -423,3 +450,6 @@ Tactic Notation "iAuIntro" :=
| iSolveTC || fail "iAuIntro: not all spatial assumptions are laterable" | iSolveTC || fail "iAuIntro: not all spatial assumptions are laterable"
| (* P = ...: make the P pretty *) pm_reflexivity | (* P = ...: make the P pretty *) pm_reflexivity
| (* the new proof mode goal *) ]. | (* the new proof mode goal *) ].
(* From here on, prevent TC search from implicitly unfolding these. *)
Typeclasses Opaque atomic_acc atomic_update.
...@@ -83,7 +83,7 @@ Section proof. ...@@ -83,7 +83,7 @@ Section proof.
<<< l {q} v, RET v >>>. <<< l {q} v, RET v >>>.
Proof. Proof.
iIntros (Q Φ) "? AU". wp_let. iIntros (Q Φ) "? AU". wp_let.
iMod (aupd_aacc with "AU") as (v q) "[H↦ [_ Hclose]]". iMod "AU" as (v q) "[H↦ [_ Hclose]]".
wp_load. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ". wp_load. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ".
Qed. Qed.
...@@ -93,7 +93,7 @@ Section proof. ...@@ -93,7 +93,7 @@ Section proof.
<<< l w, RET #() >>>. <<< l w, RET #() >>>.
Proof. Proof.
iIntros (<- Q Φ) "? AU". wp_let. wp_proj. wp_proj. iIntros (<- Q Φ) "? AU". wp_let. wp_proj. wp_proj.
iMod (aupd_aacc with "AU") as (v) "[H↦ [_ Hclose]]". iMod "AU" as (v) "[H↦ [_ Hclose]]".
wp_store. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ". wp_store. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ".
Qed. Qed.
...@@ -105,7 +105,7 @@ Section proof. ...@@ -105,7 +105,7 @@ Section proof.
RET #(if decide (v = w1) then true else false) >>>. RET #(if decide (v = w1) then true else false) >>>.
Proof. Proof.
iIntros (<- <- ? Q Φ) "? AU". wp_let. repeat wp_proj. iIntros (<- <- ? Q Φ) "? AU". wp_let. repeat wp_proj.
iMod (aupd_aacc with "AU") as (v) "[H↦ [_ Hclose]]". iMod "AU" as (v) "[H↦ [_ Hclose]]".
destruct (decide (v = w1)) as [<-|Hv]; [wp_cas_suc|wp_cas_fail]; destruct (decide (v = w1)) as [<-|Hv]; [wp_cas_suc|wp_cas_fail];
iMod ("Hclose" with "H↦") as "HΦ"; by iApply "HΦ". iMod ("Hclose" with "H↦") as "HΦ"; by iApply "HΦ".
Qed. Qed.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment