Skip to content
Snippets Groups Projects
Commit 584702fd authored by Lennard Gäher's avatar Lennard Gäher Committed by Ralf Jung
Browse files

Fupd soundness lemmas always generate credits

parent 051c25c9
No related branches found
No related tags found
No related merge requests found
...@@ -63,7 +63,8 @@ lemma. ...@@ -63,7 +63,8 @@ lemma.
and `HasNoLc`, that allow opting for either later credits or `BiFUpdPlainly`. and `HasNoLc`, that allow opting for either later credits or `BiFUpdPlainly`.
The soundness lemmas for the fancy update modality are available in two versions, The soundness lemmas for the fancy update modality are available in two versions,
with later credits (suffix `_lc`) and without later credits (suffix `_no_lc`). with later credits (suffix `_lc`) and without later credits (suffix `_no_lc`).
The lemma `step_fupdN_soundness_gen` is generic over this choice. The lemmas without later credits still generate credits, but they cannot be used
in any meaningful way. The lemma `step_fupdN_soundness_gen` is generic over this choice.
**Changes in `program_logic`:** **Changes in `program_logic`:**
......
...@@ -112,40 +112,44 @@ Proof. ...@@ -112,40 +112,44 @@ Proof.
by iFrame. by iFrame.
Qed. Qed.
Lemma fupd_plain_soundness_no_lc `{!invGpreS Σ} E1 E2 (P: iProp Σ) `{!Plain P} : (** Note: the [_no_lc] soundness lemmas also allow generating later credits, but
( `{Hinv: !invGS Σ} `{!HasNoLc Σ}, |={E1,E2}=> P) P. these cannot be used for anything. They are merely provided to enable making
the adequacy proof generic in whether later credits are used. *)
Lemma fupd_plain_soundness_no_lc `{!invGpreS Σ} E1 E2 (P: iProp Σ) `{!Plain P} m :
( `{Hinv: !invGS Σ} `{!HasNoLc Σ}, £ m |={E1,E2}=> P) P.
Proof. Proof.
iIntros (Hfupd). apply later_soundness. iMod wsat_alloc as (Hw) "[Hw HE]". iIntros (Hfupd). apply later_soundness. iMod wsat_alloc as (Hw) "[Hw HE]".
(* We don't actually want any credits, but we need the [lcGS]. *) (* We don't actually want any credits, but we need the [lcGS]. *)
iMod (later_credits.le_upd.lc_alloc 0) as (Hc) "[_ _]". iMod (later_credits.le_upd.lc_alloc m) as (Hc) "[_ Hc]".
set (Hi := InvG _ Hw Hc false). set (Hi := InvG _ Hw Hc false).
iAssert (|={,E2}=> P)%I as "H". iAssert (|={,E2}=> P)%I with "[Hc]" as "H" .
{ iMod (fupd_mask_subseteq E1) as "_"; first done. iApply Hfupd. by constructor. } { iMod (fupd_mask_subseteq E1) as "_"; first done. iApply Hfupd; last done. by constructor. }
rewrite uPred_fupd_unseal /uPred_fupd_def. rewrite uPred_fupd_unseal /uPred_fupd_def.
iMod ("H" with "[$]") as "[Hw [HE >H']]"; iFrame. iMod ("H" with "[$]") as "[Hw [HE >H']]"; iFrame.
Qed. Qed.
Lemma step_fupdN_soundness_no_lc `{!invGpreS Σ} φ n : Lemma step_fupdN_soundness_no_lc `{!invGpreS Σ} φ n m :
( `{Hinv: !invGS Σ} `{!HasNoLc Σ}, ⊢@{iPropI Σ} |={,}=> |={}▷=>^n φ ) ( `{Hinv: !invGS Σ} `{!HasNoLc Σ}, £ m ⊢@{iPropI Σ} |={,}=> |={}▷=>^n φ )
φ. φ.
Proof. Proof.
intros Hiter. intros Hiter.
apply (soundness (M:=iResUR Σ) _ (S n)); simpl. apply (soundness (M:=iResUR Σ) _ (S n)); simpl.
apply (fupd_plain_soundness_no_lc _)=> Hinv hc. apply (fupd_plain_soundness_no_lc _ m)=> Hinv hc. iIntros "Hc".
iPoseProof (Hiter Hinv) as "H". clear Hiter. iPoseProof (Hiter Hinv) as "H". clear Hiter.
iApply fupd_plainly_mask_empty. iMod "H". iApply fupd_plainly_mask_empty. iSpecialize ("H" with "Hc").
iMod (step_fupdN_plain with "H") as "H". iModIntro. iMod (step_fupdN_plain with "H") as "H". iMod "H". iModIntro.
rewrite -later_plainly -laterN_plainly -later_laterN laterN_later. rewrite -later_plainly -laterN_plainly -later_laterN laterN_later.
iNext. iMod "H" as %. auto. iNext. iMod "H" as %. auto.
Qed. Qed.
Lemma step_fupdN_soundness_no_lc' `{!invGpreS Σ} φ n : Lemma step_fupdN_soundness_no_lc' `{!invGpreS Σ} φ n m :
( `{Hinv: !invGS Σ} `{!HasNoLc Σ}, ⊢@{iPropI Σ} |={}[]▷=>^n φ ) ( `{Hinv: !invGS Σ} `{!HasNoLc Σ}, £ m ⊢@{iPropI Σ} |={}[]▷=>^n φ )
φ. φ.
Proof. Proof.
iIntros (Hiter). eapply (step_fupdN_soundness_no_lc _ n)=>Hinv Hc. destruct n as [|n]. iIntros (Hiter). eapply (step_fupdN_soundness_no_lc _ n m)=>Hinv Hc.
iIntros "Hcred". destruct n as [|n].
{ by iApply fupd_mask_intro_discard; [|iApply (Hiter Hinv)]. } { by iApply fupd_mask_intro_discard; [|iApply (Hiter Hinv)]. }
simpl in Hiter |- *. iMod Hiter as "H". iIntros "!>!>!>". simpl in Hiter |- *. iMod (Hiter with "Hcred") as "H". iIntros "!>!>!>".
iMod "H". clear. iInduction n as [|n] "IH"; [by iApply fupd_mask_intro_discard|]. iMod "H". clear. iInduction n as [|n] "IH"; [by iApply fupd_mask_intro_discard|].
simpl. iMod "H". iIntros "!>!>!>". iMod "H". by iApply "IH". simpl. iMod "H". iIntros "!>!>!>". iMod "H". by iApply "IH".
Qed. Qed.
...@@ -228,7 +232,7 @@ Qed. ...@@ -228,7 +232,7 @@ Qed.
Lemma step_fupdN_soundness_gen `{!invGpreS Σ} (φ : Prop) Lemma step_fupdN_soundness_gen `{!invGpreS Σ} (φ : Prop)
(use_credits : bool) (n m : nat) : (use_credits : bool) (n m : nat) :
( `{Hinv : invGS Σ} `{Hc : if use_credits then HasLc Σ else HasNoLc Σ}, ( `{Hinv : invGS Σ} `{Hc : if use_credits then HasLc Σ else HasNoLc Σ},
(if use_credits then £ m else True) ={,}=∗ |={}▷=>^n φ) £ m ={,}=∗ |={}▷=>^n φ)
φ. φ.
Proof. Proof.
destruct use_credits. destruct use_credits.
......
...@@ -134,7 +134,7 @@ Theorem twp_total Σ Λ `{!invGpreS Σ} s e σ Φ n : ...@@ -134,7 +134,7 @@ Theorem twp_total Σ Λ `{!invGpreS Σ} s e σ Φ n :
sn erased_step ([e], σ). (* i.e. ([e], σ) is strongly normalizing *) sn erased_step ([e], σ). (* i.e. ([e], σ) is strongly normalizing *)
Proof. Proof.
intros Hwp. apply (soundness (M:=iResUR Σ) _ 1); simpl. intros Hwp. apply (soundness (M:=iResUR Σ) _ 1); simpl.
apply (fupd_plain_soundness_no_lc _)=> Hinv HNC. apply (fupd_plain_soundness_no_lc _ 0)=> Hinv HNC. iIntros "_".
iMod (Hwp) as (stateI num_laters_per_step fork_post stateI_mono) "[Hσ H]". iMod (Hwp) as (stateI num_laters_per_step fork_post stateI_mono) "[Hσ H]".
set (iG := IrisG _ _ Hinv stateI fork_post num_laters_per_step stateI_mono). set (iG := IrisG _ _ Hinv stateI fork_post num_laters_per_step stateI_mono).
iApply (@twptp_total _ _ iG _ _ n with "Hσ"). iApply (@twptp_total _ _ iG _ _ n with "Hσ").
......
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