Skip to content
Snippets Groups Projects

Shorter barrier proof that uses auth instead of sts.

Merged Robbert Krebbers requested to merge robbert/barrier_auth into master
All threads resolved!
3 files
+ 87
224
Compare changes
  • Side-by-side
  • Inline
Files
3
+ 87
138
From iris.program_logic Require Export weakestpre.
From iris.program_logic Require Export weakestpre.
 
From iris.base_logic Require Import invariants lib.saved_prop.
From iris.heap_lang Require Export lang.
From iris.heap_lang Require Export lang.
From stdpp Require Import functions.
From iris.base_logic Require Import lib.saved_prop lib.sts.
From iris.heap_lang Require Import proofmode.
From iris.heap_lang Require Import proofmode.
 
From iris.algebra Require Import auth gset.
From iris_examples.barrier Require Export barrier.
From iris_examples.barrier Require Export barrier.
From iris_examples.barrier Require Import protocol.
Set Default Proof Using "Type".
Set Default Proof Using "Type".
(** The CMRAs/functors we need. *)
(** The CMRAs/functors we need. *)
Class barrierG Σ := BarrierG {
Class barrierG Σ := BarrierG {
barrier_stsG :> stsG Σ sts;
barrier_inG :> inG Σ (authR (gset_disjUR gname));
barrier_savedPropG :> savedPropG Σ;
barrier_savedPropG :> savedPropG Σ;
}.
}.
Definition barrierΣ : gFunctors := #[stsΣ sts; savedPropΣ].
Definition barrierΣ : gFunctors :=
 
#[ GFunctor (authRF (gset_disjUR gname)); savedPropΣ ].
Instance subG_barrierΣ {Σ} : subG barrierΣ Σ barrierG Σ.
Instance subG_barrierΣ {Σ} : subG barrierΣ Σ barrierG Σ.
Proof. solve_inG. Qed.
Proof. solve_inG. Qed.
@@ -20,181 +20,130 @@ Proof. solve_inG. Qed.
@@ -20,181 +20,130 @@ Proof. solve_inG. Qed.
(** Now we come to the Iris part of the proof. *)
(** Now we come to the Iris part of the proof. *)
Section proof.
Section proof.
Context `{!heapG Σ, !barrierG Σ} (N : namespace).
Context `{!heapG Σ, !barrierG Σ} (N : namespace).
Implicit Types I : gset gname.
Definition ress (P : iProp Σ) (I : gset gname) : iProp Σ :=
Definition barrier_inv (l : loc) (γ : gname) (P : iProp Σ) : iProp Σ :=
( Ψ : gname iProp Σ,
( (b : bool) (γsps : gset gname),
(P -∗ [ set] i I, Ψ i) [ set] i I, saved_prop_own i (Ψ i))%I.
l #b
own γ ( (GSet γsps))
Coercion state_to_val (s : state) : val :=
((if b then True else P) -∗
match s with State Low _ => #false | State High _ => #true end.
+1
([ set] γsp γsps, R, saved_prop_own γsp R R)))%I.
Arguments state_to_val !_ / : simpl nomatch.
Definition state_to_prop (s : state) (P : iProp Σ) : iProp Σ :=
match s with State Low _ => P | State High _ => True%I end.
Arguments state_to_prop !_ _ / : simpl nomatch.
Definition barrier_inv (l : loc) (P : iProp Σ) (s : state) : iProp Σ :=
(l s ress (state_to_prop s P) (state_I s))%I.
Definition barrier_ctx (γ : gname) (l : loc) (P : iProp Σ) : iProp Σ :=
sts_ctx γ N (barrier_inv l P).
Definition send (l : loc) (P : iProp Σ) : iProp Σ :=
( γ, barrier_ctx γ l P sts_ownS γ low_states {[ Send ]})%I.
Definition recv (l : loc) (R : iProp Σ) : iProp Σ :=
Definition recv (l : loc) (R : iProp Σ) : iProp Σ :=
( γ P Q i,
( γ P R' γsp,
barrier_ctx γ l P sts_ownS γ (i_states i) {[ Change i ]}
inv N (barrier_inv l γ P)
saved_prop_own i Q (Q -∗ R))%I.
(R' -∗ R)
 
own γ ( GSet {[ γsp ]})
 
saved_prop_own γsp R')%I.
Global Instance barrier_ctx_persistent (γ : gname) (l : loc) (P : iProp Σ) :
Definition send (l : loc) (P : iProp Σ) : iProp Σ :=
Persistent (barrier_ctx γ l P).
( γ, inv N (barrier_inv l γ P))%I.
Proof. apply _. Qed.
(** Setoids *)
(** Setoids *)
Global Instance ress_ne n : Proper (dist n ==> (=) ==> dist n) ress.
Instance barrier_inv_ne l γ : NonExpansive (barrier_inv l γ).
Proof. solve_proper. Qed.
Global Instance state_to_prop_ne s :
NonExpansive (state_to_prop s).
Proof. solve_proper. Qed.
Global Instance barrier_inv_ne n l :
Proper (dist n ==> eq ==> dist n) (barrier_inv l).
Proof. solve_proper. Qed.
Global Instance barrier_ctx_ne γ l : NonExpansive (barrier_ctx γ l).
Proof. solve_proper. Qed.
Proof. solve_proper. Qed.
Global Instance send_ne l : NonExpansive (send l).
Global Instance send_ne l : NonExpansive (send l).
Proof. solve_proper. Qed.
Proof. solve_proper. Qed.
Global Instance recv_ne l : NonExpansive (recv l).
Global Instance recv_ne l : NonExpansive (recv l).
Proof. solve_proper. Qed.
Proof. solve_proper. Qed.
(** Helper lemmas *)
Lemma ress_split i i1 i2 Q R1 R2 P I :
i I i1 I i2 I i1 i2
saved_prop_own i Q -∗ saved_prop_own i1 R1 -∗ saved_prop_own i2 R2 -∗
(Q -∗ R1 R2) -∗ ress P I -∗
ress P ({[i1;i2]} I {[i]}).
Proof.
iIntros (????) "#HQ #H1 #H2 HQR"; iDestruct 1 as (Ψ) "[HPΨ HΨ]".
iDestruct (big_opS_delete _ _ i with "HΨ") as "[#HΨi HΨ]"; first done.
iExists (<[i1:=R1]> (<[i2:=R2]> Ψ)). iSplitL "HQR HPΨ".
- iPoseProof (saved_prop_agree with "HQ HΨi") as "#Heq".
iNext. iRewrite "Heq" in "HQR". iIntros "HP". iSpecialize ("HPΨ" with "HP").
iDestruct (big_opS_delete _ _ i with "HPΨ") as "[HΨ HPΨ]"; first done.
iDestruct ("HQR" with "HΨ") as "[HR1 HR2]".
rewrite -assoc_L !big_opS_fn_insert'; [|abstract set_solver ..].
by iFrame.
- rewrite -assoc_L !big_opS_fn_insert; [|abstract set_solver ..]. eauto.
Qed.
(** Actual proofs *)
(** Actual proofs *)
Lemma newbarrier_spec (P : iProp Σ) :
Lemma newbarrier_spec (P : iProp Σ) :
{{{ True }}} newbarrier #() {{{ l, RET #l; recv l P send l P }}}.
{{{ True }}} newbarrier #() {{{ l, RET #l; recv l P send l P }}}.
Proof.
Proof.
iIntros (Φ) "_ HΦ".
iIntros (Φ) "_ HΦ". iApply wp_fupd. wp_lam. wp_alloc l as "Hl".
rewrite -wp_fupd /newbarrier /=. wp_lam. wp_alloc l as "Hl".
iApply ("HΦ" with "[> -]").
iApply ("HΦ" with "[> -]").
iMod (saved_prop_alloc P) as (γ) "#?".
iMod (saved_prop_alloc P) as (γsp) "#Hsp".
iMod (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]")
iMod (own_alloc ( GSet {[ γsp ]} GSet {[ γsp ]})) as (γ) "[H● H◯]".
as (γ') "[#? Hγ']"; eauto.
{ by apply auth_valid_discrete_2. }
{ iNext. rewrite /barrier_inv /=. iFrame.
iMod (inv_alloc N _ (barrier_inv l γ P) with "[Hl H●]") as "#Hinv".
iExists (const P). rewrite !big_opS_singleton /=. eauto. }
{ iExists false, {[ γsp ]}. iIntros "{$Hl $H●} !> HP".
iAssert (barrier_ctx γ' l P)%I as "#?".
rewrite big_sepS_singleton; eauto. }
{ done. }
iModIntro; iSplitL "H◯".
iAssert (sts_ownS γ' (i_states γ) {[Change γ]}
- iExists γ, P, P, γsp. iFrame; auto.
sts_ownS γ' low_states {[Send]})%I with "[> -]" as "[Hr Hs]".
- by iExists γ.
{ 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. }
iModIntro. iSplitL "Hr".
- iExists γ', P, P, γ. iFrame. auto.
- rewrite /send. auto.
Qed.
Qed.
Lemma signal_spec l P :
Lemma signal_spec l P :
{{{ send l P P }}} signal #l {{{ RET #(); True }}}.
{{{ send l P P }}} signal #l {{{ RET #(); True }}}.
Proof.
Proof.
rewrite /signal /=.
iIntros (Φ) "[Hs HP] HΦ". iDestruct "Hs" as (γ) "#Hinv". wp_lam.
iIntros (Φ) "[Hs HP] HΦ". iDestruct "Hs" as (γ) "[#Hsts Hγ]". wp_lam.
iInv N as ([] γsps) "(>Hl & H● & HRs)".
iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
{ wp_store. iModIntro. iSplitR "HΦ"; last by iApply "HΦ".
as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
iExists true, γsps. iFrame. }
destruct p; [|done]. wp_store.
wp_store. iDestruct ("HRs" with "HP") as "HRs".
iSpecialize ("HΦ" with "[#]") => //. iFrame "HΦ".
iModIntro. iSplitR "HΦ"; last by iApply "HΦ".
iMod ("Hclose" $! (State High I) ( : propset token) with "[-]"); last done.
iExists true, γsps. iFrame; eauto.
iSplit; [iPureIntro; by eauto using signal_step|].
rewrite /barrier_inv /ress /=. iNext. iFrame "Hl".
iDestruct "Hr" as (Ψ) "[Hr Hsp]"; iExists Ψ; iFrame "Hsp".
iNext. iIntros "_"; by iApply "Hr".
Qed.
Qed.
Lemma wait_spec l P:
Lemma wait_spec l P:
{{{ recv l P }}} wait #l {{{ RET #(); P }}}.
{{{ recv l P }}} wait #l {{{ RET #(); P }}}.
Proof.
Proof.
rename P into R.
rename P into R.
iIntros (Φ) "Hr HΦ"; iDestruct "Hr" as (γ P Q i) "(#Hsts & Hγ & #HQ & HQR)".
iIntros (Φ) "HR HΦ". iDestruct "HR" as (γ P R' γsp) "(#Hinv & HR & H◯ & #Hsp)".
iLöb as "IH". wp_rec. wp_bind (! _)%E.
iLöb as "IH". wp_rec. wp_bind (! _)%E.
iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
iInv N as ([] γsps) "(>Hl & >H● & HRs)"; last first.
as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
{ wp_load. iModIntro. iSplitL "Hl H● HRs".
wp_load. destruct p.
{ iExists false, γsps. iFrame. }
- iMod ("Hclose" $! (State Low I) {[ Change i ]} with "[Hl Hr]") as "Hγ".
by wp_apply ("IH" with "HΦ HR"). }
{ iSplit; first done. rewrite /barrier_inv /=. by iFrame. }
iSpecialize ("HRs" with "[//]"). wp_load.
iAssert (sts_ownS γ (i_states i) {[Change i]})%I with "[> Hγ]" as "Hγ".
iDestruct (own_valid_2 with "H● H◯")
{ iApply (sts_own_weaken with "Hγ"); eauto using i_states_closed. }
as %[Hvalid%gset_disj_included%elem_of_subseteq_singleton _]%auth_valid_discrete_2.
iModIntro. wp_if.
iDestruct (big_sepS_delete with "HRs") as "[HR'' HRs]"; first done.
iApply ("IH" with "Hγ [HQR] [HΦ]"); auto.
iDestruct "HR''" as (R'') "[#Hsp' HR'']".
- (* a High state: the comparison succeeds, and we perform a transition and
iDestruct (saved_prop_agree with "Hsp Hsp'") as "#Heq".
return to the client *)
iMod (own_update_2 with "H● H◯") as "H●".
iDestruct "Hr" as (Ψ) "[HΨ Hsp]".
{ apply (auth_update_dealloc _ _ (GSet (γsps {[ γsp ]}))).
iDestruct (big_opS_delete _ _ i with "Hsp") as "[#HΨi Hsp]"; first done.
apply gset_disj_dealloc_local_update. }
iAssert ( (Ψ i [ set] j I {[i]}, Ψ j))%I with "[HΨ]" as "[HΨ HΨ']".
iIntros "!>". iSplitL "Hl H● HRs".
{ iNext. iApply (big_opS_delete _ _ i); first done. by iApply "HΨ". }
{ iDestruct (bi.later_intro with "HRs") as "HRs".
iMod ("Hclose" $! (State High (I {[ i ]})) with "[HΨ' Hl Hsp]").
iModIntro. iExists true, (γsps {[ γsp ]}). iFrame; eauto. }
{ iSplit; [iPureIntro; by eauto using wait_step|].
wp_if. iApply "HΦ". iApply "HR". by iRewrite "Heq".
rewrite /barrier_inv /=. iNext. iFrame "Hl". iExists Ψ; iFrame. auto. }
iPoseProof (saved_prop_agree with "HQ HΨi") as "#Heq".
iModIntro. wp_if.
iApply "HΦ". iApply "HQR". by iRewrite "Heq".
Qed.
Qed.
Lemma recv_split E l P1 P2 :
Lemma recv_split E l P1 P2 :
N E recv l (P1 P2) ={E}=∗ recv l P1 recv l P2.
N E recv l (P1 P2) ={E}=∗ recv l P1 recv l P2.
Proof.
Proof.
rename P1 into R1; rename P2 into R2.
rename P1 into R1; rename P2 into R2.
iIntros (?). iDestruct 1 as (γ P Q i) "(#Hsts & Hγ & #HQ & HQR)".
iIntros (?). iDestruct 1 as (γ P R' γsp) "(#Hinv & HR & H◯ & #Hsp)".
iMod (sts_openS (barrier_inv l P) _ _ γ with "[Hγ]")
iInv N as (b γsps) "(>Hl & >H● & HRs)".
as ([p I]) "(% & [Hl Hr] & Hclose)"; eauto.
iDestruct (own_valid_2 with "H● H◯")
iMod (saved_prop_alloc_cofinite I) as (i1) "[% #Hi1]".
as %[Hvalid%gset_disj_included%elem_of_subseteq_singleton _]%auth_valid_discrete_2.
iMod (saved_prop_alloc_cofinite (I {[i1]}))
iMod (own_update_2 with "H● H◯") as "H●".
as (i2) "[Hi2' #Hi2]"; iDestruct "Hi2'" as %Hi2.
{ apply (auth_update_dealloc _ _ (GSet (γsps {[ γsp ]}))).
rewrite ->not_elem_of_union, elem_of_singleton in Hi2; destruct Hi2.
apply gset_disj_dealloc_local_update. }
iMod ("Hclose" $! (State p ({[i1; i2]} I {[i]}))
set (γsps' := γsps {[γsp]}).
{[Change i1; Change i2 ]} with "[-]") as "Hγ".
iMod (saved_prop_alloc_cofinite γsps' R1) as (γsp1 Hγsp1) "#Hsp1".
{ iSplit; first by eauto using split_step.
iMod (saved_prop_alloc_cofinite (γsps' {[ γsp1 ]}) R2)
rewrite /barrier_inv /=. iNext. iFrame "Hl".
as (γsp2 [? ?%not_elem_of_singleton]%not_elem_of_union) "#Hsp2".
by iApply (ress_split with "HQ Hi1 Hi2 HQR"). }
iMod (own_update _ _ ( _ ( GSet {[ γsp1 ]} (GSet {[ γsp2 ]})))
iAssert (sts_ownS γ (i_states i1) {[Change i1]}
with "H●") as "(H● & H◯1 & H◯2)".
sts_ownS γ (i_states i2) {[Change i2]})%I with "[> -]" as "[Hγ1 Hγ2]".
{ rewrite -auth_frag_op gset_disj_union; last set_solver.
{ iApply sts_ownS_op; eauto using i_states_closed, low_states_closed.
apply auth_update_alloc, (gset_disj_alloc_empty_local_update _ {[ γsp1; γsp2 ]}).
- abstract set_solver.
set_solver. }
- iApply (sts_own_weaken with "Hγ");
iModIntro. iSplitL "HR Hl HRs H●".
eauto using sts.closed_op, i_states_closed.
{ iModIntro. iExists b, ({[γsp1; γsp2]} γsps').
abstract set_solver. }
iIntros "{$Hl $H●} HP". iSpecialize ("HRs" with "HP").
iModIntro; iSplitL "Hγ1".
iDestruct (big_sepS_delete with "HRs") as "[HR'' HRs]"; first done.
- iExists γ, P, R1, i1. iFrame; auto.
iDestruct "HR''" as (R'') "[#Hsp' HR'']".
- iExists γ, P, R2, i2. iFrame; auto.
iDestruct (saved_prop_agree with "Hsp Hsp'") as "#Heq".
 
iAssert ( R')%I with "[HR'']" as "HR'"; [iNext; by iRewrite "Heq"|].
 
iDestruct ("HR" with "HR'") as "[HR1 HR2]".
 
iApply big_sepS_union; [set_solver|iFrame "HRs"].
 
iApply big_sepS_union; [set_solver|].
 
iSplitL "HR1"; rewrite big_sepS_singleton; eauto. }
 
iModIntro; iSplitL "H◯1".
 
- iExists γ, P, R1, γsp1. iFrame; auto.
 
- iExists γ, P, R2, γsp2. iFrame; auto.
Qed.
Qed.
Lemma recv_weaken l P1 P2 : (P1 -∗ P2) -∗ recv l P1 -∗ recv l P2.
Lemma recv_weaken l P1 P2 : (P1 -∗ P2) -∗ recv l P1 -∗ recv l P2.
Proof.
Proof.
iIntros "HP". iDestruct 1 as (γ P Q i) "(#Hctx&Hγ&Hi&HP1)".
iIntros "HP". iDestruct 1 as (γ P R' i) "(#Hinv & HR & H◯)".
iExists γ, P, Q, i. iFrame "Hctx Hγ Hi".
iExists γ, P, R', i. iIntros "{$Hinv $H◯} !> HQ". iApply "HP". by iApply "HR".
iNext. iIntros "HQ". by iApply "HP"; iApply "HP1".
Qed.
Qed.
Lemma recv_mono l P1 P2 : (P1 P2) recv l P1 recv l P2.
Lemma recv_mono l P1 P2 : (P1 P2) recv l P1 recv l P2.
Proof. iIntros (HP) "H". iApply (recv_weaken with "[] H"). iApply HP. Qed.
Proof. iIntros (HP) "H". iApply (recv_weaken with "[] H"). iApply HP. Qed.
End proof.
End proof.
Typeclasses Opaque barrier_ctx send recv.
Typeclasses Opaque send recv.
Loading