Skip to content
Snippets Groups Projects
Commit fc30ca08 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Generalize update tactics into iMod and iModIntro for modalities.

There are now two proof mode tactics for dealing with modalities:

- `iModIntro` : introduction of a modality
- `iMod pm_trm as (x1 ... xn) "ipat"` : eliminate a modality

The behavior of these tactics can be controlled by instances of the `IntroModal`
and `ElimModal` type class. We have declared instances for later, except 0,
basic updates and fancy updates. The tactic `iMod` is flexible enough that it
can also eliminate an updates around a weakest pre, and so forth.

The corresponding introduction patterns of these tactics are `!>` and `>`.

These tactics replace the tactics `iUpdIntro`, `iUpd` and `iTimeless`.

Source of backwards incompatability: the introduction pattern `!>` is used for
introduction of arbitrary modalities. It used to introduce laters by stripping
of a later of each hypotheses.
parent 9bc0a38b
No related branches found
No related tags found
No related merge requests found
Showing
with 214 additions and 207 deletions
......@@ -98,8 +98,19 @@ Separating logic specific tactics
- `iCombine "H1" "H2" as "H"` : turns `H1 : P1` and `H2 : P2` into
`H : P1 ★ P2`.
Modalities
----------
- `iModIntro` : introduction of a modality that is an instance of the
`IntoModal` type class. Instances include: later, except 0, basic update and
fancy update.
- `iMod pm_trm as (x1 ... xn) "ipat"` : eliminate a modality `pm_trm` that is
an instance of the `ElimModal` type class. Instances include: later, except 0,
basic update and fancy update.
The later modality
------------------
- `iNext` : introduce a later by stripping laters from all hypotheses.
- `iLöb as "IH" forall (x1 ... xn)` : perform Löb induction by generating a
hypothesis `IH : ▷ goal`. The tactic generalizes over the Coq level variables
......@@ -108,6 +119,7 @@ The later modality
Induction
---------
- `iInduction x as cpat "IH" forall (x1 ... xn) "selpat"` : perform induction on
the Coq term `x`. The Coq introduction pattern is used to name the introduced
variables. The induction hypotheses are inserted into the persistent context
......@@ -124,13 +136,8 @@ Rewriting
Iris
----
- `iUpdIntro` : introduction of an update modality.
- `iUpd pm_trm as (x1 ... xn) "ipat"` : run an update modality `pm_trm` (if the
goal permits, i.e. it can be expanded to an update modality.
- `iInv N as (x1 ... xn) "ipat"` : open the invariant `N`.
- `iTimeless "H"` : strip a later of a timeless hypothesis `H` (if the goal
permits, i.e. it is a later, True now, update modality, or a weakest
precondition).
- `iInv N as (x1 ... xn) "ipat" "Hclose"` : open the invariant `N`, the update
for closing the invariant is put in a hypothesis named `Hclose`.
Miscellaneous
-------------
......@@ -170,8 +177,7 @@ _introduction patterns_:
- `[]` : false elimination.
- `%` : move the hypothesis to the pure Coq context (anonymously).
- `# ipat` : move the hypothesis to the persistent context.
- `> ipat` : remove a later of a timeless hypothesis (if the goal permits).
- `==> ipat` : run an update modality (if the goal permits).
- `> ipat` : eliminate a modality (if the goal permits).
Apart from this, there are the following introduction patterns that can only
appear at the top level:
......@@ -181,8 +187,7 @@ appear at the top level:
previous pattern, e.g., `{$H1 H2 $H3}`).
- `!%` : introduce a pure goal (and leave the proof mode).
- `!#` : introduce an always modality (given that the spatial context is empty).
- `!>` : introduce a later (which strips laters from all hypotheses).
- `!==>` : introduce an update modality
- `!>` : introduce a modality.
- `/=` : perform `simpl`.
- `*` : introduce all universal quantifiers.
- `**` : introduce all universal quantifiers, as well as all arrows and wands.
......@@ -222,9 +227,9 @@ _specification patterns_ to express splitting of hypotheses:
framed in the generated goal.
- `[-H1 ... Hn]` : negated form of the above pattern. This pattern does not
accept hypotheses prefixed with a `$`.
- `==>[H1 ... Hn]` : same as the above pattern, but can only be used if the goal
is an update modality, in which case the update modality will be kept in the
goal of the premise too.
- `>[H1 ... Hn]` : same as the above pattern, but can only be used if the goal
is a modality, in which case the modality will be kept in the generated goal
for the premise will be wrapped into the modality.
- `[#]` : This pattern can be used when eliminating `P -★ Q` with `P`
persistent. Using this pattern, all hypotheses are available in the goal for
`P`, as well the remaining goal.
......
......@@ -20,7 +20,7 @@ Definition heap_adequacy Σ `{heapPreG Σ} e σ φ :
adequate e σ φ.
Proof.
intros Hwp; eapply (wp_adequacy Σ); iIntros (?) "Hσ".
iUpd (auth_alloc to_heap _ heapN _ σ with "[Hσ]") as (γ) "[Hh _]";[|by iNext|].
iMod (auth_alloc to_heap _ heapN _ σ with "[Hσ]") as (γ) "[Hh _]";[|by iNext|].
{ exact: to_heap_valid. }
set (Hheap := HeapG _ _ _ γ).
iApply (Hwp _). by rewrite /heap_ctx.
......
......@@ -109,10 +109,10 @@ Section heap.
heap_ctx ( l, l v ={E}=★ Φ (LitV (LitLoc l))) WP Alloc e @ E {{ Φ }}.
Proof.
iIntros (<-%of_to_val ?) "[#Hinv HΦ]". rewrite /heap_ctx.
iUpd (auth_empty heap_name) as "Ha".
iUpd (auth_open with "[$Hinv $Ha]") as (σ) "(%&Hσ&Hcl)"; first done.
iApply wp_alloc_pst. iFrame "Hσ". iNext. iIntros (l) "[% Hσ] !==>".
iUpd ("Hcl" with "* [Hσ]") as "Ha".
iMod (auth_empty heap_name) as "Ha".
iMod (auth_open with "[$Hinv $Ha]") as (σ) "(%&Hσ&Hcl)"; first done.
iApply wp_alloc_pst. iFrame "Hσ". iNext. iIntros (l) "[% Hσ] !>".
iMod ("Hcl" with "* [Hσ]") as "Ha".
{ iFrame. iPureIntro. rewrite to_heap_insert.
eapply alloc_singleton_local_update; by auto using lookup_to_heap_None. }
iApply "HΦ". by rewrite heap_mapsto_eq /heap_mapsto_def.
......@@ -125,10 +125,10 @@ Section heap.
Proof.
iIntros (?) "[#Hinv [>Hl HΦ]]".
rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
iUpd (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iMod (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iApply (wp_load_pst _ σ); first eauto using heap_singleton_included.
iIntros "{$Hσ} !> Hσ !==>". iUpd ("Hcl" with "* [Hσ]") as "Ha"; first eauto.
by iApply "HΦ".
iIntros "{$Hσ}"; iNext; iIntros "Hσ !>".
iMod ("Hcl" with "* [Hσ]") as "Ha"; first eauto. by iApply "HΦ".
Qed.
Lemma wp_store E l v' e v Φ :
......@@ -138,9 +138,9 @@ Section heap.
Proof.
iIntros (<-%of_to_val ?) "[#Hinv [>Hl HΦ]]".
rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
iUpd (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iMod (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iApply (wp_store_pst _ σ); first eauto using heap_singleton_included.
iIntros "{$Hσ} !> Hσ !==>". iUpd ("Hcl" with "* [Hσ]") as "Ha".
iIntros "{$Hσ}"; iNext; iIntros "Hσ !>". iMod ("Hcl" with "* [Hσ]") as "Ha".
{ iFrame. iPureIntro. rewrite to_heap_insert.
eapply singleton_local_update, exclusive_local_update; last done.
by eapply heap_singleton_included'. }
......@@ -154,10 +154,10 @@ Section heap.
Proof.
iIntros (<-%of_to_val <-%of_to_val ??) "[#Hinv [>Hl HΦ]]".
rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
iUpd (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iMod (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iApply (wp_cas_fail_pst _ σ); [eauto using heap_singleton_included|done|].
iIntros "{$Hσ} !> Hσ !==>". iUpd ("Hcl" with "* [Hσ]") as "Ha"; first eauto.
by iApply "HΦ".
iIntros "{$Hσ}"; iNext; iIntros "Hσ !>".
iMod ("Hcl" with "* [Hσ]") as "Ha"; first eauto. by iApply "HΦ".
Qed.
Lemma wp_cas_suc E l e1 v1 e2 v2 Φ :
......@@ -167,9 +167,9 @@ Section heap.
Proof.
iIntros (<-%of_to_val <-%of_to_val ?) "[#Hinv [>Hl HΦ]]".
rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
iUpd (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iMod (auth_open with "[$Hinv $Hl]") as (σ) "(%&Hσ&Hcl)"; first done.
iApply (wp_cas_suc_pst _ σ); first eauto using heap_singleton_included.
iIntros "{$Hσ} !> Hσ !==>". iUpd ("Hcl" with "* [Hσ]") as "Ha".
iIntros "{$Hσ}"; iNext; iIntros "Hσ !>". iMod ("Hcl" with "* [Hσ]") as "Ha".
{ iFrame. iPureIntro. rewrite to_heap_insert.
eapply singleton_local_update, exclusive_local_update; last done.
by eapply heap_singleton_included'. }
......
......@@ -97,22 +97,22 @@ Lemma newbarrier_spec (P : iProp Σ) (Φ : val → iProp Σ) :
Proof.
iIntros (HN) "[#? HΦ]".
rewrite /newbarrier /=. wp_seq. wp_alloc l as "Hl".
iApply ("HΦ" with "==>[-]").
iUpd (saved_prop_alloc (F:=idCF) P) as (γ) "#?".
iUpd (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]")
iApply ("HΦ" with ">[-]").
iMod (saved_prop_alloc (F:=idCF) P) as (γ) "#?".
iMod (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]")
as (γ') "[#? Hγ']"; eauto.
{ iNext. rewrite /barrier_inv /=. iFrame.
iExists (const P). rewrite !big_sepS_singleton /=. eauto. }
iAssert (barrier_ctx γ' l P)%I as "#?".
{ rewrite /barrier_ctx. by repeat iSplit. }
iAssert (sts_ownS γ' (i_states γ) {[Change γ]}
sts_ownS γ' low_states {[Send]})%I with "==>[-]" as "[Hr Hs]".
sts_ownS γ' low_states {[Send]})%I with ">[-]" as "[Hr Hs]".
{ iApply sts_ownS_op; eauto using i_states_closed, low_states_closed.
- set_solver.
- iApply (sts_own_weaken with "Hγ'");
auto using sts.closed_op, i_states_closed, low_states_closed;
abstract set_solver. }
iUpdIntro. rewrite /recv /send. iSplitL "Hr".
iModIntro. rewrite /recv /send. iSplitL "Hr".
- iExists γ', P, P, γ. iFrame. auto.
- auto.
Qed.
......@@ -122,14 +122,14 @@ Lemma signal_spec l P (Φ : val → iProp Σ) :
Proof.
rewrite /signal /send /barrier_ctx /=.
iIntros "(Hs&HP&HΦ)"; iDestruct "Hs" as (γ) "[#(%&Hh&Hsts) Hγ]". wp_let.
iUpd (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
destruct p; [|done]. wp_store. iFrame "HΦ".
iUpd ("Hclose" $! (State High I) ( : set token) with "[-]"); last done.
iMod ("Hclose" $! (State High I) ( : set token) with "[-]"); last done.
iSplit; [iPureIntro; by eauto using signal_step|].
iNext. rewrite {2}/barrier_inv /ress /=; iFrame "Hl".
iDestruct "Hr" as (Ψ) "[Hr Hsp]"; iExists Ψ; iFrame "Hsp".
iIntros "!> _"; by iApply "Hr".
iNext. iIntros "_"; by iApply "Hr".
Qed.
Lemma wait_spec l P (Φ : val iProp Σ) :
......@@ -138,14 +138,14 @@ Proof.
rename P into R; rewrite /recv /barrier_ctx.
iIntros "[Hr HΦ]"; iDestruct "Hr" as (γ P Q i) "(#(%&Hh&Hsts)&Hγ&#HQ&HQR)".
iLöb as "IH". wp_rec. wp_bind (! _)%E.
iUpd (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
wp_load. destruct p.
- iUpd ("Hclose" $! (State Low I) {[ Change i ]} with "[Hl Hr]") as "Hγ".
- iMod ("Hclose" $! (State Low I) {[ Change i ]} with "[Hl Hr]") as "Hγ".
{ iSplit; first done. iNext. rewrite {2}/barrier_inv /=. by iFrame. }
iAssert (sts_ownS γ (i_states i) {[Change i]})%I with "==>[Hγ]" as "Hγ".
iAssert (sts_ownS γ (i_states i) {[Change i]})%I with ">[Hγ]" as "Hγ".
{ iApply (sts_own_weaken with "Hγ"); eauto using i_states_closed. }
iUpdIntro. wp_if.
iModIntro. wp_if.
iApply ("IH" with "Hγ [HQR] HΦ"). auto.
- (* a High state: the comparison succeeds, and we perform a transition and
return to the client *)
......@@ -153,12 +153,12 @@ Proof.
iDestruct (big_sepS_delete _ _ i with "Hsp") as "[#HΨi Hsp]"; first done.
iAssert ( Ψ i [ set] j I {[i]}, Ψ j)%I with "[HΨ]" as "[HΨ HΨ']".
{ iNext. iApply (big_sepS_delete _ _ i); first done. by iApply "HΨ". }
iUpd ("Hclose" $! (State High (I {[ i ]})) ( : set token) with "[HΨ' Hl Hsp]").
iMod ("Hclose" $! (State High (I {[ i ]})) ( : set token) with "[HΨ' Hl Hsp]").
{ iSplit; [iPureIntro; by eauto using wait_step|].
iNext. rewrite {2}/barrier_inv /=; iFrame "Hl". iExists Ψ; iFrame. auto. }
iPoseProof (saved_prop_agree i Q (Ψ i) with "[#]") as "Heq"; first by auto.
iUpdIntro. wp_if.
iUpdIntro. iApply "HΦ". iApply "HQR". by iRewrite "Heq".
iModIntro. wp_if.
iModIntro. iApply "HΦ". iApply "HQR". by iRewrite "Heq".
Qed.
Lemma recv_split E l P1 P2 :
......@@ -166,25 +166,25 @@ Lemma recv_split E l P1 P2 :
Proof.
rename P1 into R1; rename P2 into R2. rewrite {1}/recv /barrier_ctx.
iIntros (?). iDestruct 1 as (γ P Q i) "(#(%&Hh&Hsts)&Hγ&#HQ&HQR)".
iUpd (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
iUpd (saved_prop_alloc_strong (R1: ∙%CF (iProp Σ)) I) as (i1) "[% #Hi1]".
iUpd (saved_prop_alloc_strong (R2: ∙%CF (iProp Σ)) (I {[i1]}))
iMod (saved_prop_alloc_strong (R1: ∙%CF (iProp Σ)) I) as (i1) "[% #Hi1]".
iMod (saved_prop_alloc_strong (R2: ∙%CF (iProp Σ)) (I {[i1]}))
as (i2) "[Hi2' #Hi2]"; iDestruct "Hi2'" as %Hi2.
rewrite ->not_elem_of_union, elem_of_singleton in Hi2; destruct Hi2.
iUpd ("Hclose" $! (State p ({[i1; i2]} I {[i]}))
iMod ("Hclose" $! (State p ({[i1; i2]} I {[i]}))
{[Change i1; Change i2 ]} with "[-]") as "Hγ".
{ iSplit; first by eauto using split_step.
iNext. rewrite {2}/barrier_inv /=. iFrame "Hl".
iApply (ress_split _ _ _ Q R1 R2); eauto. iFrame; auto. }
iAssert (sts_ownS γ (i_states i1) {[Change i1]}
sts_ownS γ (i_states i2) {[Change i2]})%I with "==>[-]" as "[Hγ1 Hγ2]".
sts_ownS γ (i_states i2) {[Change i2]})%I with ">[-]" as "[Hγ1 Hγ2]".
{ iApply sts_ownS_op; eauto using i_states_closed, low_states_closed.
- abstract set_solver.
- iApply (sts_own_weaken with "Hγ");
eauto using sts.closed_op, i_states_closed.
abstract set_solver. }
iUpdIntro; iSplitL "Hγ1"; rewrite /recv /barrier_ctx.
iModIntro; iSplitL "Hγ1"; rewrite /recv /barrier_ctx.
- iExists γ, P, R1, i1. iFrame; auto.
- iExists γ, P, R2, i2. iFrame; auto.
Qed.
......@@ -194,7 +194,7 @@ Proof.
rewrite /recv.
iIntros "HP HP1"; iDestruct "HP1" as (γ P Q i) "(#Hctx&Hγ&Hi&HP1)".
iExists γ, P, Q, i. iFrame "Hctx Hγ Hi".
iIntros "!> HQ". by iApply "HP"; iApply "HP1".
iNext. iIntros "HQ". by iApply "HP"; iApply "HP1".
Qed.
Lemma recv_mono l P1 P2 : (P1 P2) recv l P1 recv l P2.
......
......@@ -38,10 +38,10 @@ Section mono_proof.
heap_ctx ( l, mcounter l 0 -★ Φ #l) WP newcounter #() {{ Φ }}.
Proof.
iIntros (?) "[#Hh HΦ]". rewrite /newcounter /=. wp_seq. wp_alloc l as "Hl".
iUpd (own_alloc ( (O:mnat) (O:mnat))) as (γ) "[Hγ Hγ']"; first done.
iUpd (inv_alloc N _ (mcounter_inv γ l) with "[Hl Hγ]").
iMod (own_alloc ( (O:mnat) (O:mnat))) as (γ) "[Hγ Hγ']"; first done.
iMod (inv_alloc N _ (mcounter_inv γ l) with "[Hl Hγ]").
{ iNext. iExists 0%nat. by iFrame. }
iUpdIntro. iApply "HΦ". rewrite /mcounter; eauto 10.
iModIntro. iApply "HΦ". rewrite /mcounter; eauto 10.
Qed.
Lemma inc_mono_spec l n (Φ : val iProp Σ) :
......@@ -50,22 +50,22 @@ Section mono_proof.
iIntros "[Hl HΦ]". iLöb as "IH". wp_rec.
iDestruct "Hl" as (γ) "(% & #? & #Hinv & Hγf)".
wp_bind (! _)%E. iInv N as (c) ">[Hγ Hl]" "Hclose".
wp_load. iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iUpdIntro. wp_let. wp_op.
wp_load. iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iModIntro. wp_let. wp_op.
wp_bind (CAS _ _ _). iInv N as (c') ">[Hγ Hl]" "Hclose".
destruct (decide (c' = c)) as [->|].
- iDestruct (own_valid_2 with "[$Hγ $Hγf]")
as %[?%mnat_included _]%auth_valid_discrete_2.
iUpd (own_update_2 with "[$Hγ $Hγf]") as "[Hγ Hγf]".
iMod (own_update_2 with "[$Hγ $Hγf]") as "[Hγ Hγf]".
{ apply auth_update, (mnat_local_update _ _ (S c)); auto. }
wp_cas_suc. iUpd ("Hclose" with "[Hl Hγ]") as "_".
wp_cas_suc. iMod ("Hclose" with "[Hl Hγ]") as "_".
{ iNext. iExists (S c). rewrite Nat2Z.inj_succ Z.add_1_l. by iFrame. }
iUpdIntro. wp_if. iApply "HΦ"; iExists γ; repeat iSplit; eauto.
iModIntro. wp_if. iApply "HΦ"; iExists γ; repeat iSplit; eauto.
iApply (own_mono with "Hγf"). apply: auth_frag_mono.
by apply mnat_included, le_n_S.
- wp_cas_fail; first (by intros [= ?%Nat2Z.inj]).
iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c'; by iFrame|].
iUpdIntro. wp_if. iApply ("IH" with "[Hγf] HΦ").
iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c'; by iFrame|].
iModIntro. wp_if. iApply ("IH" with "[Hγf] HΦ").
rewrite {3}/mcounter; eauto 10.
Qed.
......@@ -77,9 +77,9 @@ Section mono_proof.
rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
iDestruct (own_valid_2 with "[$Hγ $Hγf]")
as %[?%mnat_included _]%auth_valid_discrete_2.
iUpd (own_update_2 with "[$Hγ $Hγf]") as "[Hγ Hγf]".
iMod (own_update_2 with "[$Hγ $Hγf]") as "[Hγ Hγf]".
{ apply auth_update, (mnat_local_update _ _ c); auto. }
iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iApply ("HΦ" with "[%]"); rewrite /mcounter; eauto 10.
Qed.
End mono_proof.
......@@ -116,11 +116,11 @@ Section contrib_spec.
WP newcounter #() {{ Φ }}.
Proof.
iIntros (?) "[#Hh HΦ]". rewrite /newcounter /=. wp_seq. wp_alloc l as "Hl".
iUpd (own_alloc ( (Some (1%Qp, O%nat)) (Some (1%Qp, 0%nat))))
iMod (own_alloc ( (Some (1%Qp, O%nat)) (Some (1%Qp, 0%nat))))
as (γ) "[Hγ Hγ']"; first done.
iUpd (inv_alloc N _ (ccounter_inv γ l) with "[Hl Hγ]").
iMod (inv_alloc N _ (ccounter_inv γ l) with "[Hl Hγ]").
{ iNext. iExists 0%nat. by iFrame. }
iUpdIntro. iApply "HΦ". rewrite /ccounter_ctx /ccounter; eauto 10.
iModIntro. iApply "HΦ". rewrite /ccounter_ctx /ccounter; eauto 10.
Qed.
Lemma inc_contrib_spec γ l q n (Φ : val iProp Σ) :
......@@ -129,19 +129,19 @@ Section contrib_spec.
Proof.
iIntros "(#(%&?&?) & Hγf & HΦ)". iLöb as "IH". wp_rec.
wp_bind (! _)%E. iInv N as (c) ">[Hγ Hl]" "Hclose".
wp_load. iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iUpdIntro. wp_let. wp_op.
wp_load. iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iModIntro. wp_let. wp_op.
wp_bind (CAS _ _ _). iInv N as (c') ">[Hγ Hl]" "Hclose".
destruct (decide (c' = c)) as [->|].
- iUpd (own_update_2 with "[$Hγ $Hγf]") as "[Hγ Hγf]".
- iMod (own_update_2 with "[$Hγ $Hγf]") as "[Hγ Hγf]".
{ apply auth_update, option_local_update, prod_local_update_2.
apply (nat_local_update _ _ (S c) (S n)); omega. }
wp_cas_suc. iUpd ("Hclose" with "[Hl Hγ]") as "_".
wp_cas_suc. iMod ("Hclose" with "[Hl Hγ]") as "_".
{ iNext. iExists (S c). rewrite Nat2Z.inj_succ Z.add_1_l. by iFrame. }
iUpdIntro. wp_if. by iApply "HΦ".
iModIntro. wp_if. by iApply "HΦ".
- wp_cas_fail; first (by intros [= ?%Nat2Z.inj]).
iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c'; by iFrame|].
iUpdIntro. wp_if. by iApply ("IH" with "[Hγf] HΦ").
iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c'; by iFrame|].
iModIntro. wp_if. by iApply ("IH" with "[Hγf] HΦ").
Qed.
Lemma read_contrib_spec γ l q n (Φ : val iProp Σ) :
......@@ -153,7 +153,7 @@ Section contrib_spec.
rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
iDestruct (own_valid_2 with "[$Hγ $Hγf]")
as %[[? ?%nat_included]%Some_pair_included_total_2 _]%auth_valid_discrete_2.
iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iApply ("HΦ" with "[%]"); rewrite /ccounter; eauto 10.
Qed.
......@@ -165,7 +165,7 @@ Section contrib_spec.
rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
iDestruct (own_valid_2 with "[$Hγ $Hγf]") as %[Hn _]%auth_valid_discrete_2.
apply (Some_included_exclusive _) in Hn as [= ->]%leibniz_equiv; last done.
iUpd ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
iMod ("Hclose" with "[Hl Hγ]") as "_"; [iNext; iExists c; by iFrame|].
by iApply "HΦ".
Qed.
End contrib_spec.
......@@ -23,7 +23,7 @@ Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp
WP par e {{ Φ }}.
Proof.
iIntros (?) "(#Hh&Hf1&Hf2&HΦ)".
rewrite /par. wp_value. iUpdIntro. wp_let. wp_proj.
rewrite /par. wp_value. iModIntro. wp_let. wp_proj.
wp_apply (spawn_spec parN); try wp_done; try solve_ndisj; iFrame "Hf1 Hh".
iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _).
iApply wp_wand_l; iFrame "Hf2"; iIntros (v) "H2". wp_let.
......
......@@ -54,11 +54,11 @@ Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) (Φ : val → iProp Σ) :
Proof.
iIntros (<-%of_to_val ?) "(#Hh & Hf & HΦ)". rewrite /spawn /=.
wp_let. wp_alloc l as "Hl". wp_let.
iUpd (own_alloc (Excl ())) as (γ) "Hγ"; first done.
iUpd (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?".
iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done.
iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?".
{ iNext. iExists NONEV. iFrame; eauto. }
wp_apply wp_fork; simpl. iSplitR "Hf".
- iUpdIntro. wp_seq. iUpdIntro. iApply "HΦ". rewrite /join_handle. eauto.
- iModIntro. wp_seq. iModIntro. iApply "HΦ". rewrite /join_handle. eauto.
- wp_bind (f _). iApply wp_wand_l. iFrame "Hf"; iIntros (v) "Hv".
iInv N as (v') "[Hl _]" "Hclose".
wp_store. iApply "Hclose". iNext. iExists (SOMEV v). iFrame. eauto.
......@@ -70,11 +70,11 @@ Proof.
rewrite /join_handle; iIntros "[[% H] Hv]". iDestruct "H" as (γ) "(#?&Hγ&#?)".
iLöb as "IH". wp_rec. wp_bind (! _)%E. iInv N as (v) "[Hl Hinv]" "Hclose".
wp_load. iDestruct "Hinv" as "[%|Hinv]"; subst.
- iUpd ("Hclose" with "[Hl]"); [iNext; iExists _; iFrame; eauto|].
iUpdIntro. wp_match. iApply ("IH" with "Hγ Hv").
- iMod ("Hclose" with "[Hl]"); [iNext; iExists _; iFrame; eauto|].
iModIntro. wp_match. iApply ("IH" with "Hγ Hv").
- iDestruct "Hinv" as (v') "[% [HΨ|Hγ']]"; simplify_eq/=.
+ iUpd ("Hclose" with "[Hl Hγ]"); [iNext; iExists _; iFrame; eauto|].
iUpdIntro. wp_match. by iApply "Hv".
+ iMod ("Hclose" with "[Hl Hγ]"); [iNext; iExists _; iFrame; eauto|].
iModIntro. wp_match. by iApply "Hv".
+ iCombine "Hγ" "Hγ'" as "Hγ". iDestruct (own_valid with "Hγ") as %[].
Qed.
End proof.
......
......@@ -51,10 +51,10 @@ Section proof.
Proof.
iIntros (?) "(#Hh & HR & HΦ)". rewrite /newlock /=.
wp_seq. wp_alloc l as "Hl".
iUpd (own_alloc (Excl ())) as (γ) "Hγ"; first done.
iUpd (inv_alloc N _ (lock_inv γ l R) with "[-HΦ]") as "#?".
iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done.
iMod (inv_alloc N _ (lock_inv γ l R) with "[-HΦ]") as "#?".
{ iIntros "!>". iExists false. by iFrame. }
iUpdIntro. iApply "HΦ". iExists l. eauto.
iModIntro. iApply "HΦ". iExists l. eauto.
Qed.
Lemma try_acquire_spec γ lk R (Φ: val iProp Σ) :
......@@ -63,11 +63,11 @@ Section proof.
Proof.
iIntros "[#Hl HΦ]". iDestruct "Hl" as (l) "(% & #? & % & #?)". subst.
wp_rec. iInv N as ([]) "[Hl HR]" "Hclose".
- wp_cas_fail. iUpd ("Hclose" with "[Hl]"); first (iNext; iExists true; eauto).
iUpdIntro. iDestruct "HΦ" as "[_ HΦ]". iApply "HΦ".
- wp_cas_fail. iMod ("Hclose" with "[Hl]"); first (iNext; iExists true; eauto).
iModIntro. iDestruct "HΦ" as "[_ HΦ]". iApply "HΦ".
- wp_cas_suc. iDestruct "HR" as "[Hγ HR]".
iUpd ("Hclose" with "[Hl]"); first (iNext; iExists true; eauto).
iUpdIntro. iDestruct "HΦ" as "[HΦ _]". rewrite /locked. by iApply ("HΦ" with "Hγ HR").
iMod ("Hclose" with "[Hl]"); first (iNext; iExists true; eauto).
iModIntro. iDestruct "HΦ" as "[HΦ _]". rewrite /locked. by iApply ("HΦ" with "Hγ HR").
Qed.
Lemma acquire_spec γ lk R (Φ : val iProp Σ) :
......@@ -75,7 +75,7 @@ Section proof.
Proof.
iIntros "[#Hl HΦ]". iLöb as "IH". wp_rec. wp_bind (try_acquire _).
iApply try_acquire_spec. iFrame "#". iSplit.
- iIntros "Hlked HR". wp_if. iUpdIntro. iApply ("HΦ" with "Hlked HR").
- iIntros "Hlked HR". wp_if. iModIntro. iApply ("HΦ" with "Hlked HR").
- wp_if. iApply ("IH" with "HΦ").
Qed.
......
......@@ -80,12 +80,12 @@ Section proof.
Proof.
iIntros (HN) "(#Hh & HR & HΦ)". rewrite /newlock /=.
wp_seq. wp_alloc lo as "Hlo". wp_alloc ln as "Hln".
iUpd (own_alloc ( (Excl' 0%nat, ) (Excl' 0%nat, ))) as (γ) "[Hγ Hγ']".
iMod (own_alloc ( (Excl' 0%nat, ) (Excl' 0%nat, ))) as (γ) "[Hγ Hγ']".
{ by rewrite -auth_both_op. }
iUpd (inv_alloc _ _ (lock_inv γ lo ln R) with "[-HΦ]").
iMod (inv_alloc _ _ (lock_inv γ lo ln R) with "[-HΦ]").
{ iNext. rewrite /lock_inv.
iExists 0%nat, 0%nat. iFrame. iLeft. by iFrame. }
iUpdIntro. iApply ("HΦ" $! (#lo, #ln)%V γ). iExists lo, ln. eauto.
iModIntro. iApply ("HΦ" $! (#lo, #ln)%V γ). iExists lo, ln. eauto.
Qed.
Lemma wait_loop_spec γ lk x R (Φ : val iProp Σ) :
......@@ -96,16 +96,16 @@ Section proof.
iInv N as (o n) "(Hlo & Hln & Ha)" "Hclose".
wp_load. destruct (decide (x = o)) as [->|Hneq].
- iDestruct "Ha" as "[Hainv [[Ho HR] | Haown]]".
+ iUpd ("Hclose" with "[Hlo Hln Hainv Ht]") as "_".
+ iMod ("Hclose" with "[Hlo Hln Hainv Ht]") as "_".
{ iNext. iExists o, n. iFrame. eauto. }
iUpdIntro. wp_let. wp_op=>[_|[]] //.
wp_if. iUpdIntro.
iModIntro. wp_let. wp_op=>[_|[]] //.
wp_if. iModIntro.
iApply ("HΦ" with "[-HR] HR"). rewrite /locked; eauto.
+ iDestruct (own_valid_2 with "[$Ht $Haown]") as % [_ ?%gset_disj_valid_op].
set_solver.
- iUpd ("Hclose" with "[Hlo Hln Ha]").
- iMod ("Hclose" with "[Hlo Hln Ha]").
{ iNext. iExists o, n. by iFrame. }
iUpdIntro. wp_let. wp_op=>[[/Nat2Z.inj //]|?].
iModIntro. wp_let. wp_op=>[[/Nat2Z.inj //]|?].
wp_if. iApply ("IH" with "Ht"). by iExact "HΦ".
Qed.
......@@ -115,28 +115,28 @@ Section proof.
iIntros "[Hl HΦ]". iDestruct "Hl" as (lo ln) "(% & #? & % & #?)".
iLöb as "IH". wp_rec. wp_bind (! _)%E. subst. wp_proj.
iInv N as (o n) "[Hlo [Hln Ha]]" "Hclose".
wp_load. iUpd ("Hclose" with "[Hlo Hln Ha]") as "_".
wp_load. iMod ("Hclose" with "[Hlo Hln Ha]") as "_".
{ iNext. iExists o, n. by iFrame. }
iUpdIntro. wp_let. wp_proj. wp_op.
iModIntro. wp_let. wp_proj. wp_op.
wp_bind (CAS _ _ _).
iInv N as (o' n') "(>Hlo' & >Hln' & >Hauth & Haown)" "Hclose".
destruct (decide (#n' = #n))%V as [[= ->%Nat2Z.inj] | Hneq].
- wp_cas_suc.
iUpd (own_update with "Hauth") as "[Hauth Hofull]".
iMod (own_update with "Hauth") as "[Hauth Hofull]".
{ eapply auth_update_alloc, prod_local_update_2.
eapply (gset_disj_alloc_empty_local_update _ {[ n ]}).
apply (seq_set_S_disjoint 0). }
rewrite -(seq_set_S_union_L 0).
iUpd ("Hclose" with "[Hlo' Hln' Haown Hauth]") as "_".
iMod ("Hclose" with "[Hlo' Hln' Haown Hauth]") as "_".
{ iNext. iExists o', (S n).
rewrite Nat2Z.inj_succ -Z.add_1_r. by iFrame. }
iUpdIntro. wp_if.
iModIntro. wp_if.
iApply (wait_loop_spec γ (#lo, #ln)).
iFrame "HΦ". rewrite /issued; eauto 10.
- wp_cas_fail.
iUpd ("Hclose" with "[Hlo' Hln' Hauth Haown]") as "_".
iMod ("Hclose" with "[Hlo' Hln' Hauth Haown]") as "_".
{ iNext. iExists o', n'. by iFrame. }
iUpdIntro. wp_if. by iApply "IH".
iModIntro. wp_if. by iApply "IH".
Qed.
Lemma release_spec γ lk R (Φ : val iProp Σ):
......@@ -150,19 +150,19 @@ Section proof.
wp_load.
iDestruct (own_valid_2 with "[$Hauth $Hγo]") as
%[[<-%Excl_included%leibniz_equiv _]%prod_included _]%auth_valid_discrete_2.
iUpd ("Hclose" with "[Hlo Hln Hauth Haown]") as "_".
iMod ("Hclose" with "[Hlo Hln Hauth Haown]") as "_".
{ iNext. iExists o, n. by iFrame. }
iUpdIntro. wp_op.
iModIntro. wp_op.
iInv N as (o' n') "(>Hlo & >Hln & >Hauth & Haown)" "Hclose".
wp_store.
iDestruct (own_valid_2 with "[$Hauth $Hγo]") as
%[[<-%Excl_included%leibniz_equiv _]%prod_included _]%auth_valid_discrete_2.
iDestruct "Haown" as "[[Hγo' _]|?]".
{ iDestruct (own_valid_2 with "[$Hγo $Hγo']") as %[[] ?]. }
iUpd (own_update_2 with "[$Hauth $Hγo]") as "[Hauth Hγo]".
iMod (own_update_2 with "[$Hauth $Hγo]") as "[Hauth Hγo]".
{ apply auth_update, prod_local_update_1.
by apply option_local_update, (exclusive_local_update _ (Excl (S o))). }
iUpd ("Hclose" with "[Hlo Hln Hauth Haown Hγo HR]") as "_"; last auto.
iMod ("Hclose" with "[Hlo Hln Hauth Haown Hγo HR]") as "_"; last auto.
iNext. iExists (S o), n'.
rewrite Nat2Z.inj_succ -Z.add_1_r. iFrame. iLeft. by iFrame.
Qed.
......
......@@ -46,9 +46,9 @@ Proof.
rewrite {1}wp_unfold /wp_pre. iIntros (Hstep) "[(Hw & HE & Hσ) [H|[_ H]]]".
{ iDestruct "H" as (v) "[% _]". apply val_stuck in Hstep; simplify_eq. }
rewrite fupd_eq /fupd_def.
iUpd ("H" $! σ1 with "Hσ [Hw HE]") as ">(Hw & HE & _ & H)"; first by iFrame.
iUpdIntro; iNext.
iUpd ("H" $! e2 σ2 efs with "[%] [Hw HE]")
iMod ("H" $! σ1 with "Hσ [Hw HE]") as ">(Hw & HE & _ & H)"; first by iFrame.
iModIntro; iNext.
iMod ("H" $! e2 σ2 efs with "[%] [Hw HE]")
as ">($ & $ & $ & $)"; try iFrame; eauto.
Qed.
......@@ -76,8 +76,8 @@ Proof.
revert e1 t1 t2 σ1 σ2; simpl; induction n as [|n IH]=> e1 t1 t2 σ1 σ2 /=.
{ inversion_clear 1; iIntros "?"; eauto 10. }
iIntros (Hsteps) "H". inversion_clear Hsteps as [|?? [t1' σ1']].
iUpd (wptp_step with "H") as (e1' t1'') "[% H]"; first eauto; simplify_eq.
iUpdIntro; iNext; iUpd "H" as ">?". by iApply IH.
iMod (wptp_step with "H") as (e1' t1'') "[% H]"; first eauto; simplify_eq.
iModIntro; iNext; iMod "H" as ">?". by iApply IH.
Qed.
Instance bupd_iter_mono n : Proper (() ==> ()) (Nat.iter n (λ P, |==> P)%I).
......@@ -99,7 +99,7 @@ Proof.
rewrite (Nat_iter_S_r (S n)). apply bupd_iter_mono.
iDestruct 1 as (e2 t2') "(% & (Hw & HE & _) & H & _)"; simplify_eq.
iDestruct (wp_value_inv with "H") as "H". rewrite fupd_eq /fupd_def.
iUpd ("H" with "[Hw HE]") as ">(_ & _ & $)"; iFrame; auto.
iMod ("H" with "[Hw HE]") as ">(_ & _ & $)"; iFrame; auto.
Qed.
Lemma wp_safe e σ Φ :
......@@ -107,7 +107,7 @@ Lemma wp_safe e σ Φ :
Proof.
rewrite wp_unfold /wp_pre. iIntros "[(Hw&HE&Hσ) [H|[_ H]]]".
{ iDestruct "H" as (v) "[% _]"; eauto 10. }
rewrite fupd_eq. iUpd ("H" with "* Hσ [-]") as ">(?&?&%&?)"; first by iFrame.
rewrite fupd_eq. iMod ("H" with "* Hσ [-]") as ">(?&?&%&?)"; first by iFrame.
eauto 10.
Qed.
......@@ -133,7 +133,7 @@ Proof.
iIntros "[HI H]".
iDestruct "H" as (e2' t2') "(% & (Hw&HE&Hσ) & _)"; subst.
rewrite fupd_eq in HI;
iUpd (HI with "HI [Hw HE]") as "> (_ & _ & H)"; first by iFrame.
iMod (HI with "HI [Hw HE]") as "> (_ & _ & H)"; first by iFrame.
iDestruct "H" as (σ2') "[Hσf %]".
iDestruct (ownP_agree σ2 σ2' with "[-]") as %<-. by iFrame. eauto.
Qed.
......@@ -146,13 +146,13 @@ Proof.
intros Hwp; split.
- intros t2 σ2 v2 [n ?]%rtc_nsteps.
eapply (soundness (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
rewrite Nat_iter_S. iUpd (iris_alloc σ) as (?) "(?&?&?&Hσ)".
iUpdIntro. iNext. iApply wptp_result; eauto.
rewrite Nat_iter_S. iMod (iris_alloc σ) as (?) "(?&?&?&Hσ)".
iModIntro. iNext. iApply wptp_result; eauto.
iFrame. iSplitL. by iApply Hwp. by iApply big_sepL_nil.
- intros t2 σ2 e2 [n ?]%rtc_nsteps ?.
eapply (soundness (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
rewrite Nat_iter_S. iUpd (iris_alloc σ) as (?) "(Hw & HE & Hσ & Hσf)".
iUpdIntro. iNext. iApply wptp_safe; eauto.
rewrite Nat_iter_S. iMod (iris_alloc σ) as (?) "(Hw & HE & Hσ & Hσf)".
iModIntro. iNext. iApply wptp_safe; eauto.
iFrame "Hw HE Hσ". iSplitL. by iApply Hwp. by iApply big_sepL_nil.
Qed.
......@@ -164,8 +164,8 @@ Theorem wp_invariance Σ `{irisPreG Λ Σ} e σ1 t2 σ2 I φ Φ :
Proof.
intros Hwp HI [n ?]%rtc_nsteps.
eapply (soundness (M:=iResUR Σ) _ (S (S (S n)))); iIntros "".
rewrite Nat_iter_S. iUpd (iris_alloc σ1) as (?) "(Hw & HE & ? & Hσ)".
rewrite Nat_iter_S. iMod (iris_alloc σ1) as (?) "(Hw & HE & ? & Hσ)".
rewrite fupd_eq in Hwp.
iUpd (Hwp _ with "Hσ [Hw HE]") as ">(? & ? & ? & ?)"; first by iFrame.
iUpdIntro. iNext. iApply wptp_invariance; eauto. iFrame. by iApply big_sepL_nil.
iMod (Hwp _ with "Hσ [Hw HE]") as ">(? & ? & ? & ?)"; first by iFrame.
iModIntro. iNext. iApply wptp_invariance; eauto. iFrame. by iApply big_sepL_nil.
Qed.
......@@ -95,9 +95,9 @@ Section auth.
(f t) φ t ={E}=★ γ, (γ G) auth_ctx γ N f φ auth_own γ (f t).
Proof.
iIntros (?) "Hφ". rewrite /auth_own /auth_ctx.
iUpd (own_alloc_strong (Auth (Excl' (f t)) (f t)) G) as (γ) "[% Hγ]"; first done.
iMod (own_alloc_strong (Auth (Excl' (f t)) (f t)) G) as (γ) "[% Hγ]"; first done.
iRevert "Hγ"; rewrite auth_both_op; iIntros "[Hγ Hγ']".
iUpd (inv_alloc N _ (auth_inv γ f φ) with "[-Hγ']") as "#?".
iMod (inv_alloc N _ (auth_inv γ f φ) with "[-Hγ']") as "#?".
{ iNext. rewrite /auth_inv. iExists t. by iFrame. }
eauto.
Qed.
......@@ -106,7 +106,7 @@ Section auth.
(f t) φ t ={E}=★ γ, auth_ctx γ N f φ auth_own γ (f t).
Proof.
iIntros (?) "Hφ".
iUpd (auth_alloc_strong N E t with "Hφ") as (γ) "[_ ?]"; eauto.
iMod (auth_alloc_strong N E t with "Hφ") as (γ) "[_ ?]"; eauto.
Qed.
Lemma auth_empty γ : True ==★ auth_own γ ∅.
......@@ -119,12 +119,12 @@ Section auth.
Proof.
iIntros "(Hinv & Hγf)". rewrite /auth_inv /auth_own.
iDestruct "Hinv" as (t) "[>Hγa Hφ]".
iUpdIntro. iExists t.
iModIntro. iExists t.
iDestruct (own_valid_2 with "[$Hγa $Hγf]") as % [? ?]%auth_valid_discrete_2.
iSplit; first done. iFrame. iIntros (u b) "[% Hφ]".
iUpd (own_update_2 with "[$Hγa $Hγf]") as "[Hγa Hγf]".
iMod (own_update_2 with "[$Hγa $Hγf]") as "[Hγa Hγf]".
{ eapply auth_update; eassumption. }
iUpdIntro. iFrame. iExists u. iFrame.
iModIntro. iFrame. iExists u. iFrame.
Qed.
Lemma auth_open E N γ a :
......@@ -140,9 +140,9 @@ Section auth.
to unpack and repack various proofs.
TODO: Make this mostly automatic, by supporting "opening accessors
around accessors". *)
iUpd (auth_acc with "[$Hinv $Hγf]") as (t) "(?&?&HclAuth)".
iUpdIntro. iExists t. iFrame. iIntros (u b) "H".
iUpd ("HclAuth" $! u b with "H") as "(Hinv & ?)". by iUpd ("Hclose" with "Hinv").
iMod (auth_acc with "[$Hinv $Hγf]") as (t) "(?&?&HclAuth)".
iModIntro. iExists t. iFrame. iIntros (u b) "H".
iMod ("HclAuth" $! u b with "H") as "(Hinv & ?)". by iMod ("Hclose" with "Hinv").
Qed.
End auth.
......
......@@ -75,7 +75,7 @@ Lemma box_own_agree γ Q1 Q2 :
Proof.
rewrite /box_own_prop own_valid_2 prod_validI /= and_elim_r.
rewrite option_validI /= agree_validI agree_equivI later_equivI /=.
iIntros "#HQ !>". rewrite -{2}(iProp_fold_unfold Q1).
iIntros "#HQ". iNext. rewrite -{2}(iProp_fold_unfold Q1).
iRewrite "HQ". by rewrite iProp_fold_unfold.
Qed.
......@@ -91,14 +91,14 @@ Lemma box_insert f P Q :
slice N γ Q box N (<[γ:=false]> f) (Q P).
Proof.
iDestruct 1 as (Φ) "[#HeqP Hf]".
iUpd (own_alloc_strong ( Excl' false Excl' false,
iMod (own_alloc_strong ( Excl' false Excl' false,
Some (to_agree (Next (iProp_unfold Q)))) (dom _ f))
as (γ) "[Hdom Hγ]"; first done.
rewrite pair_split. iDestruct "Hγ" as "[[Hγ Hγ'] #HγQ]".
iDestruct "Hdom" as % ?%not_elem_of_dom.
iUpd (inv_alloc N _ (slice_inv γ Q) with "[Hγ]") as "#Hinv".
iMod (inv_alloc N _ (slice_inv γ Q) with "[Hγ]") as "#Hinv".
{ iNext. iExists false; eauto. }
iUpdIntro; iExists γ; repeat iSplit; auto.
iModIntro; iExists γ; repeat iSplit; auto.
iNext. iExists (<[γ:=Q]> Φ); iSplit.
- iNext. iRewrite "HeqP". by rewrite big_sepM_fn_insert'.
- rewrite (big_sepM_fn_insert (λ _ _ P', _ _ _ P' _ _ (_ _ P')))%I //.
......@@ -113,7 +113,7 @@ Proof.
iIntros (?) "[#Hinv H]"; iDestruct "H" as (Φ) "[#HeqP Hf]".
iExists ([ map] γ'↦_ delete γ f, Φ γ')%I.
iInv N as (b) "(Hγ & #HγQ &_)" "Hclose".
iApply fupd_trans_frame; iFrame "Hclose"; iUpdIntro; iNext.
iApply fupd_trans_frame; iFrame "Hclose"; iModIntro; iNext.
iDestruct (big_sepM_delete _ f _ false with "Hf")
as "[[Hγ' #[HγΦ ?]] ?]"; first done.
iDestruct (box_own_agree γ Q (Φ γ) with "[#]") as "HeqQ"; first by eauto.
......@@ -133,10 +133,10 @@ Proof.
iDestruct (big_sepM_later _ f with "Hf") as "Hf".
iDestruct (big_sepM_delete _ f _ false with "Hf")
as "[[>Hγ' #[HγΦ Hinv']] ?]"; first done.
iUpd (box_own_auth_update γ b' false true with "[Hγ Hγ']")
iMod (box_own_auth_update γ b' false true with "[Hγ Hγ']")
as "[Hγ Hγ']"; first by iFrame.
iUpd ("Hclose" with "[Hγ HQ]"); first (iNext; iExists true; by iFrame).
iUpdIntro; iNext; iExists Φ; iSplit.
iMod ("Hclose" with "[Hγ HQ]"); first (iNext; iExists true; by iFrame).
iModIntro; iNext; iExists Φ; iSplit.
- by rewrite big_sepM_insert_override.
- rewrite -insert_delete big_sepM_insert ?lookup_delete //.
iFrame; eauto.
......@@ -153,9 +153,9 @@ Proof.
as "[[>Hγ' #[HγΦ Hinv']] ?]"; first done.
iDestruct (box_own_auth_agree γ b true with "[-]") as %->; first by iFrame.
iFrame "HQ".
iUpd (box_own_auth_update γ with "[Hγ Hγ']") as "[Hγ Hγ']"; first by iFrame.
iUpd ("Hclose" with "[Hγ]"); first (iNext; iExists false; by repeat iSplit).
iUpdIntro; iNext; iExists Φ; iSplit.
iMod (box_own_auth_update γ with "[Hγ Hγ']") as "[Hγ Hγ']"; first by iFrame.
iMod ("Hclose" with "[Hγ]"); first (iNext; iExists false; by repeat iSplit).
iModIntro; iNext; iExists Φ; iSplit.
- by rewrite big_sepM_insert_override.
- rewrite -insert_delete big_sepM_insert ?lookup_delete //.
iFrame; eauto.
......@@ -171,7 +171,7 @@ Proof.
iApply (big_sepM_impl _ _ f); iFrame "Hf".
iAlways; iIntros (γ b' ?) "[(Hγ' & #$ & #$) HΦ]".
iInv N as (b) "[>Hγ _]" "Hclose".
iUpd (box_own_auth_update γ with "[Hγ Hγ']") as "[Hγ $]"; first by iFrame.
iMod (box_own_auth_update γ with "[Hγ Hγ']") as "[Hγ $]"; first by iFrame.
iApply "Hclose". iNext; iExists true. by iFrame.
Qed.
......@@ -181,17 +181,17 @@ Lemma box_empty_all f P Q :
Proof.
iDestruct 1 as (Φ) "[#HeqP Hf]".
iAssert ([ map] γb f, Φ γ box_own_auth γ ( Excl' false)
box_own_prop γ (Φ γ) inv N (slice_inv γ (Φ γ)))%I with "==>[Hf]" as "[HΦ ?]".
box_own_prop γ (Φ γ) inv N (slice_inv γ (Φ γ)))%I with ">[Hf]" as "[HΦ ?]".
{ iApply (fupd_big_sepM _ _ f); iApply (big_sepM_impl _ _ f); iFrame "Hf".
iAlways; iIntros (γ b ?) "(Hγ' & #$ & #$)".
assert (true = b) as <- by eauto.
iInv N as (b) "(>Hγ & _ & HΦ)" "Hclose".
iDestruct (box_own_auth_agree γ b true with "[-]") as %->; first by iFrame.
iUpd (box_own_auth_update γ true true false with "[Hγ Hγ']")
iMod (box_own_auth_update γ true true false with "[Hγ Hγ']")
as "[Hγ $]"; first by iFrame.
iUpd ("Hclose" with "[Hγ]"); first (iNext; iExists false; iFrame; eauto).
iMod ("Hclose" with "[Hγ]"); first (iNext; iExists false; iFrame; eauto).
by iApply "HΦ". }
iUpdIntro; iSplitL "HΦ".
iModIntro; iSplitL "HΦ".
- rewrite eq_iff later_iff big_sepM_later. by iApply "HeqP".
- iExists Φ; iSplit; by rewrite big_sepM_fmap.
Qed.
......
......@@ -47,8 +47,8 @@ Section proofs.
Lemma cinv_alloc E N P : P ={E}=★ γ, cinv N γ P cinv_own γ 1.
Proof.
rewrite /cinv /cinv_own. iIntros "HP".
iUpd (own_alloc 1%Qp) as (γ) "H1"; first done.
iUpd (inv_alloc N _ (P own γ 1%Qp)%I with "[HP]"); eauto.
iMod (own_alloc 1%Qp) as (γ) "H1"; first done.
iMod (inv_alloc N _ (P own γ 1%Qp)%I with "[HP]"); eauto.
Qed.
Lemma cinv_cancel E N γ P :
......@@ -65,7 +65,7 @@ Section proofs.
Proof.
rewrite /cinv. iIntros (?) "#Hinv Hγ".
iInv N as "[$|>Hγ']" "Hclose".
- iIntros "!==> {$Hγ} HP". iApply "Hclose"; eauto.
- iIntros "!> {$Hγ} HP". iApply "Hclose"; eauto.
- iDestruct (cinv_own_1_l with "[Hγ Hγ']") as %[]. by iFrame.
Qed.
End proofs.
......@@ -40,9 +40,9 @@ Module savedprop. Section savedprop.
Lemma contradiction : False.
Proof.
apply (@soundness M False 1); simpl.
iIntros "". iUpd A_alloc as (i) "#H".
iIntros "". iMod A_alloc as (i) "#H".
iPoseProof (saved_NA with "H") as "HN".
iUpdIntro. iNext.
iModIntro. iNext.
iApply "HN". iApply saved_A. done.
Qed.
......@@ -113,12 +113,12 @@ Module inv. Section inv.
Lemma fupd_frame_r E P Q : (fupd E P Q) fupd E (P Q).
Proof. by rewrite comm fupd_frame_l comm. Qed.
Global Instance elim_fupd_fupd E P Q : ElimUpd (fupd E P) P (fupd E Q) (fupd E Q).
Proof. by rewrite /ElimUpd fupd_frame_r uPred.wand_elim_r fupd_fupd. Qed.
Global Instance elim_fupd_fupd E P Q : ElimModal (fupd E P) P (fupd E Q) (fupd E Q).
Proof. by rewrite /ElimModal fupd_frame_r uPred.wand_elim_r fupd_fupd. Qed.
Global Instance elim_fupd0_fupd1 P Q : ElimUpd (fupd M0 P) P (fupd M1 Q) (fupd M1 Q).
Global Instance elim_fupd0_fupd1 P Q : ElimModal (fupd M0 P) P (fupd M1 Q) (fupd M1 Q).
Proof.
by rewrite /ElimUpd fupd_frame_r uPred.wand_elim_r fupd_mask_mono fupd_fupd.
by rewrite /ElimModal fupd_frame_r uPred.wand_elim_r fupd_mask_mono fupd_fupd.
Qed.
Global Instance exists_split_fupd0 {A} E P (Φ : A iProp) :
......@@ -135,8 +135,8 @@ Module inv. Section inv.
Lemma saved_alloc (P : gname iProp) : True fupd M1 ( γ, saved γ (P γ)).
Proof.
iIntros "". iUpd (sts_alloc) as (γ) "Hs".
iUpd (inv_alloc (start γ (finished γ (P γ))) with "[Hs]") as (i) "#Hi".
iIntros "". iMod (sts_alloc) as (γ) "Hs".
iMod (inv_alloc (start γ (finished γ (P γ))) with "[Hs]") as (i) "#Hi".
{ auto. }
iApply fupd_intro. by iExists γ, i.
Qed.
......@@ -145,7 +145,7 @@ Module inv. Section inv.
Proof.
iIntros "(#HsP & #HsQ & #HP)". iDestruct "HsP" as (i) "HiP".
iApply (inv_open' i). iSplit; first done.
iIntros "HaP". iAssert (fupd M0 (finished γ)) with "[HaP]" as "==> Hf".
iIntros "HaP". iAssert (fupd M0 (finished γ)) with "[HaP]" as "> Hf".
{ iDestruct "HaP" as "[Hs | [Hf _]]".
- by iApply start_finish.
- by iApply fupd_intro. }
......@@ -173,7 +173,7 @@ Module inv. Section inv.
Proof.
iIntros "#Hi !# #HA". iPoseProof "HA" as "HA'".
iDestruct "HA'" as (P) "#[HNP Hi']".
iUpd (saved_cast i (A i) P with "[]") as "HP".
iMod (saved_cast i (A i) P with "[]") as "HP".
{ eauto. }
by iApply "HNP".
Qed.
......@@ -187,7 +187,7 @@ Module inv. Section inv.
Lemma contradiction : False.
Proof.
apply consistency. iIntros "".
iUpd A_alloc as (i) "#H".
iMod A_alloc as (i) "#H".
iPoseProof (saved_NA with "H") as "HN".
iApply "HN". iApply saved_A. done.
Qed.
......
......@@ -23,7 +23,7 @@ Lemma wp_lift_head_step E Φ e1 :
WP e1 @ E {{ Φ }}.
Proof.
iIntros "H". iApply (wp_lift_step E); try done.
iUpd "H" as (σ1) "(%&Hσ1&Hwp)". iUpdIntro. iExists σ1.
iMod "H" as (σ1) "(%&Hσ1&Hwp)". iModIntro. iExists σ1.
iSplit; first by eauto. iFrame. iNext. iIntros (e2 σ2 efs) "[% ?]".
iApply "Hwp". by eauto.
Qed.
......
......@@ -51,20 +51,20 @@ Proof. apply ne_proper, _. Qed.
Lemma fupd_intro_mask E1 E2 P : E2 E1 P |={E1,E2}=> |={E2,E1}=> P.
Proof.
intros (E1''&->&?)%subseteq_disjoint_union_L.
rewrite fupd_eq /fupd_def ownE_op //. iIntros "H ($ & $ & HE) !==>".
iApply except_0_intro. iIntros "[$ $] !==>" . iApply except_0_intro.
rewrite fupd_eq /fupd_def ownE_op //. iIntros "H ($ & $ & HE) !>".
iApply except_0_intro. iIntros "[$ $] !>" . iApply except_0_intro.
by iFrame.
Qed.
Lemma except_0_fupd E1 E2 P : (|={E1,E2}=> P) ={E1,E2}=★ P.
Proof.
rewrite fupd_eq. iIntros "H [Hw HE]". iTimeless "H". iApply "H"; by iFrame.
rewrite fupd_eq. iIntros "H [Hw HE]". iMod "H". iApply "H"; by iFrame.
Qed.
Lemma bupd_fupd E P : (|==> P) ={E}=★ P.
Proof.
rewrite fupd_eq /fupd_def. iIntros "H [$ $]"; iUpd "H".
iUpdIntro. by iApply except_0_intro.
rewrite fupd_eq /fupd_def. iIntros "H [$ $]"; iMod "H".
iModIntro. by iApply except_0_intro.
Qed.
Lemma fupd_mono E1 E2 P Q : (P Q) (|={E1,E2}=> P) ={E1,E2}=★ Q.
......@@ -76,16 +76,16 @@ Qed.
Lemma fupd_trans E1 E2 E3 P : (|={E1,E2}=> |={E2,E3}=> P) ={E1,E3}=★ P.
Proof.
rewrite fupd_eq /fupd_def. iIntros "HP HwE".
iUpd ("HP" with "HwE") as ">(Hw & HE & HP)". iApply "HP"; by iFrame.
iMod ("HP" with "HwE") as ">(Hw & HE & HP)". iApply "HP"; by iFrame.
Qed.
Lemma fupd_mask_frame_r' E1 E2 Ef P :
E1 Ef (|={E1,E2}=> E2 Ef P) ={E1 Ef,E2 Ef}=★ P.
Proof.
intros. rewrite fupd_eq /fupd_def ownE_op //. iIntros "Hvs (Hw & HE1 &HEf)".
iUpd ("Hvs" with "[Hw HE1]") as ">($ & HE2 & HP)"; first by iFrame.
iMod ("Hvs" with "[Hw HE1]") as ">($ & HE2 & HP)"; first by iFrame.
iDestruct (ownE_op' with "[HE2 HEf]") as "[? $]"; first by iFrame.
iUpdIntro; iApply except_0_intro. by iApply "HP".
iModIntro; iApply except_0_intro. by iApply "HP".
Qed.
Lemma fupd_frame_r E1 E2 P Q : (|={E1,E2}=> P) Q ={E1,E2}=★ P Q.
......@@ -183,16 +183,18 @@ Section proofmode_classes.
Global Instance is_except_0_fupd E1 E2 P : IsExcept0 (|={E1,E2}=> P).
Proof. by rewrite /IsExcept0 except_0_fupd. Qed.
Global Instance from_upd_fupd E P : FromUpd (|={E}=> P) P.
Proof. by rewrite /FromUpd -bupd_fupd. Qed.
Global Instance into_modal_fupd E P : IntoModal P (|={E}=> P).
Proof. rewrite /IntoModal. apply fupd_intro. Qed.
Global Instance elim_upd_bupd_fupd E1 E2 P Q :
ElimUpd (|==> P) P (|={E1,E2}=> Q) (|={E1,E2}=> Q).
Proof. by rewrite /ElimUpd (bupd_fupd E1) fupd_frame_r wand_elim_r fupd_trans. Qed.
Global Instance elim_upd_fupd_fupd E1 E2 E3 P Q :
ElimUpd (|={E1,E2}=> P) P (|={E1,E3}=> Q) (|={E2,E3}=> Q).
Proof. by rewrite /ElimUpd fupd_frame_r wand_elim_r fupd_trans. Qed.
Global Instance elim_modal_bupd_fupd E1 E2 P Q :
ElimModal (|==> P) P (|={E1,E2}=> Q) (|={E1,E2}=> Q).
Proof.
by rewrite /ElimModal (bupd_fupd E1) fupd_frame_r wand_elim_r fupd_trans.
Qed.
Global Instance elim_modal_fupd_fupd E1 E2 E3 P Q :
ElimModal (|={E1,E2}=> P) P (|={E1,E3}=> Q) (|={E2,E3}=> Q).
Proof. by rewrite /ElimModal fupd_frame_r wand_elim_r fupd_trans. Qed.
End proofmode_classes.
Hint Extern 2 (coq_tactics.of_envs _ _) =>
match goal with |- _ |={_}=> _ => iUpdIntro end.
match goal with |- _ |={_}=> _ => iModIntro end.
......@@ -62,7 +62,7 @@ Lemma ht_vs E P P' Φ Φ' e :
(P ={E}=> P') {{ P' }} e @ E {{ Φ' }} ( v, Φ' v ={E}=> Φ v)
{{ P }} e @ E {{ Φ }}.
Proof.
iIntros "(#Hvs & #Hwp & #HΦ) !# HP". iUpd ("Hvs" with "HP") as "HP".
iIntros "(#Hvs & #Hwp & #HΦ) !# HP". iMod ("Hvs" with "HP") as "HP".
iApply wp_fupd; iApply wp_wand_r; iSplitL; [by iApply "Hwp"|].
iIntros (v) "Hv". by iApply "HΦ".
Qed.
......@@ -73,7 +73,7 @@ Lemma ht_atomic E1 E2 P P' Φ Φ' e :
{{ P }} e @ E1 {{ Φ }}.
Proof.
iIntros (?) "(#Hvs & #Hwp & #HΦ) !# HP". iApply (wp_atomic _ E2); auto.
iUpd ("Hvs" with "HP") as "HP". iUpdIntro.
iMod ("Hvs" with "HP") as "HP". iModIntro.
iApply wp_wand_r; iSplitL; [by iApply "Hwp"|].
iIntros (v) "Hv". by iApply "HΦ".
Qed.
......
......@@ -32,7 +32,7 @@ Proof. rewrite inv_eq /inv; apply _. Qed.
Lemma inv_alloc N E P : P ={E}=★ inv N P.
Proof.
rewrite inv_eq /inv_def fupd_eq /fupd_def. iIntros "HP [Hw $]".
iUpd (ownI_alloc ( nclose N) P with "[HP Hw]") as (i) "(% & $ & ?)"; auto.
iMod (ownI_alloc ( nclose N) P with "[HP Hw]") as (i) "(% & $ & ?)"; auto.
- intros Ef. exists (coPpick (nclose N coPset.of_gset Ef)).
rewrite -coPset.elem_of_of_gset comm -elem_of_difference.
apply coPpick_elem_of=> Hfin.
......@@ -49,16 +49,16 @@ Proof.
iDestruct "Hi" as % ?%elem_of_subseteq_singleton.
rewrite {1 4}(union_difference_L (nclose N) E) // ownE_op; last set_solver.
rewrite {1 5}(union_difference_L {[ i ]} (nclose N)) // ownE_op; last set_solver.
iIntros "(Hw & [HE $] & $)"; iUpdIntro; iApply except_0_intro.
iIntros "(Hw & [HE $] & $)"; iModIntro; iApply except_0_intro.
iDestruct (ownI_open i P with "[Hw HE]") as "($ & $ & HD)"; first by iFrame.
iIntros "HP [Hw $] !==>"; iApply except_0_intro. iApply ownI_close; by iFrame.
iIntros "HP [Hw $] !> !>". iApply ownI_close; by iFrame.
Qed.
Lemma inv_open_timeless E N P `{!TimelessP P} :
nclose N E inv N P ={E,EN}=★ P (P ={EN,E}=★ True).
Proof.
iIntros (?) "Hinv". iUpd (inv_open with "Hinv") as "[>HP Hclose]"; auto.
iIntros "!==> {$HP} HP". iApply "Hclose"; auto.
iIntros (?) "Hinv". iMod (inv_open with "Hinv") as "[>HP Hclose]"; auto.
iIntros "!> {$HP} HP". iApply "Hclose"; auto.
Qed.
End inv.
......@@ -66,7 +66,7 @@ Tactic Notation "iInvCore" constr(N) "as" tactic(tac) constr(Hclose) :=
let Htmp := iFresh in
let patback := intro_pat.parse_one Hclose in
let pat := constr:(IList [[IName Htmp; patback]]) in
iUpd (inv_open _ N with "[#]") as pat;
iMod (inv_open _ N with "[#]") as pat;
[idtac|iAssumption || fail "iInv: invariant" N "not found"|idtac];
[solve_ndisj || match goal with |- ?P => fail "iInv: cannot solve" P end
|tac Htmp].
......
......@@ -19,13 +19,13 @@ Lemma wp_lift_step E Φ e1 :
Proof.
iIntros "H". rewrite wp_unfold /wp_pre.
destruct (to_val e1) as [v|] eqn:EQe1.
- iLeft. iExists v. iSplit. done. iUpd "H" as (σ1) "[% _]".
- iLeft. iExists v. iSplit. done. iMod "H" as (σ1) "[% _]".
by erewrite reducible_not_val in EQe1.
- iRight; iSplit; eauto.
iIntros (σ1) "Hσ". iUpd "H" as (σ1') "(% & >Hσf & H)".
iIntros (σ1) "Hσ". iMod "H" as (σ1') "(% & >Hσf & H)".
iDestruct (ownP_agree σ1 σ1' with "[-]") as %<-; first by iFrame.
iUpdIntro; iSplit; [done|]; iNext; iIntros (e2 σ2 efs Hstep).
iUpd (ownP_update σ1 σ2 with "[-H]") as "[$ ?]"; first by iFrame.
iModIntro; iSplit; [done|]; iNext; iIntros (e2 σ2 efs Hstep).
iMod (ownP_update σ1 σ2 with "[-H]") as "[$ ?]"; first by iFrame.
iApply "H"; eauto.
Qed.
......@@ -38,10 +38,10 @@ Lemma wp_lift_pure_step `{Inhabited (state Λ)} E Φ e1 :
Proof.
iIntros (Hsafe Hstep) "H". rewrite wp_unfold /wp_pre; iRight; iSplit; auto.
{ iPureIntro. eapply reducible_not_val, (Hsafe inhabitant). }
iIntros (σ1) "Hσ". iUpd (fupd_intro_mask' E ) as "Hclose"; first set_solver.
iUpdIntro. iSplit; [done|]; iNext; iIntros (e2 σ2 efs ?).
iIntros (σ1) "Hσ". iMod (fupd_intro_mask' E ) as "Hclose"; first set_solver.
iModIntro. iSplit; [done|]; iNext; iIntros (e2 σ2 efs ?).
destruct (Hstep σ1 e2 σ2 efs); auto; subst.
iUpd "Hclose"; iUpdIntro. iFrame "Hσ". iApply "H"; auto.
iMod "Hclose"; iModIntro. iFrame "Hσ". iApply "H"; auto.
Qed.
(** Derived lifting lemmas. *)
......@@ -53,12 +53,12 @@ Lemma wp_lift_atomic_step {E Φ} e1 σ1 :
WP e1 @ E {{ Φ }}.
Proof.
iIntros (Hatomic ?) "[Hσ H]". iApply (wp_lift_step E _ e1).
iUpd (fupd_intro_mask' E ) as "Hclose"; first set_solver. iUpdIntro.
iMod (fupd_intro_mask' E ) as "Hclose"; first set_solver. iModIntro.
iExists σ1. iFrame "Hσ"; iSplit; eauto.
iNext; iIntros (e2 σ2 efs) "[% Hσ]".
edestruct (Hatomic σ1 e2 σ2 efs) as [v2 <-%of_to_val]; eauto.
iDestruct ("H" $! v2 σ2 efs with "[Hσ]") as "[HΦ $]"; first by eauto.
iUpd "Hclose". iUpd "HΦ". iApply wp_value; auto using to_of_val.
iMod "Hclose". iMod "HΦ". iApply wp_value; auto using to_of_val.
Qed.
Lemma wp_lift_atomic_det_step {E Φ e1} σ1 v2 σ2 efs :
......
......@@ -85,10 +85,10 @@ Section sts.
φ s ={E}=★ γ, sts_ctx γ N φ sts_own γ s ( sts.tok s).
Proof.
iIntros "Hφ". rewrite /sts_ctx /sts_own.
iUpd (own_alloc (sts_auth s ( sts.tok s))) as (γ) "Hγ".
iMod (own_alloc (sts_auth s ( sts.tok s))) as (γ) "Hγ".
{ apply sts_auth_valid; set_solver. }
iExists γ; iRevert "Hγ"; rewrite -sts_op_auth_frag_up; iIntros "[Hγ $]".
iUpd (inv_alloc N _ (sts_inv γ φ) with "[Hφ Hγ]") as "#?"; auto.
iMod (inv_alloc N _ (sts_inv γ φ) with "[Hφ Hγ]") as "#?"; auto.
rewrite /sts_inv. iNext. iExists s. by iFrame.
Qed.
......@@ -103,11 +103,11 @@ Section sts.
assert (s S) by eauto using sts_auth_frag_valid_inv.
assert ( sts_frag S T) as [??] by eauto using cmra_valid_op_r.
rewrite sts_op_auth_frag //.
iUpdIntro; iExists s; iSplit; [done|]; iFrame "Hφ".
iModIntro; iExists s; iSplit; [done|]; iFrame "Hφ".
iIntros (s' T') "[% Hφ]".
iUpd (own_update with "Hγ") as "Hγ"; first eauto using sts_update_auth.
iMod (own_update with "Hγ") as "Hγ"; first eauto using sts_update_auth.
iRevert "Hγ"; rewrite -sts_op_auth_frag_up; iIntros "[Hγ $]".
iUpdIntro. iNext. iExists s'; by iFrame.
iModIntro. iNext. iExists s'; by iFrame.
Qed.
Lemma sts_acc E γ s0 T :
......@@ -129,9 +129,9 @@ Section sts.
to unpack and repack various proofs.
TODO: Make this mostly automatic, by supporting "opening accessors
around accessors". *)
iUpd (sts_accS with "[Hinv Hγf]") as (s) "(?&?& HclSts)"; first by iFrame.
iUpdIntro. iExists s. iFrame. iIntros (s' T') "H".
iUpd ("HclSts" $! s' T' with "H") as "(Hinv & ?)". by iUpd ("Hclose" with "Hinv").
iMod (sts_accS with "[Hinv Hγf]") as (s) "(?&?& HclSts)"; first by iFrame.
iModIntro. iExists s. iFrame. iIntros (s' T') "H".
iMod ("HclSts" $! s' T' with "H") as "(Hinv & ?)". by iMod ("Hclose" with "Hinv").
Qed.
Lemma sts_open E N γ s0 T :
......
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