Skip to content
Snippets Groups Projects

Compare revisions

Changes are shown as if the source revision was being merged into the target revision. Learn more about comparing revisions.

Source

Select target project
No results found

Target

Select target project
  • Villetaneuse/lambda-rust
  • iris/lambda-rust
  • maximedenes/LambdaRust-coq
  • msammler/lambda-rust
  • daniel.louwrink/lambda-rust
  • simonspies/lambda-rust
  • xldenis/lambda-rust
  • lgaeher/lambda-rust
  • JasonHuZS/lambda-rust
  • snyke7/lambda-rust
  • ivanbakel/lambda-rust
11 results
Show changes
Showing
with 3935 additions and 366 deletions
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import list.
From iris.bi Require Import fractional.
From lrust.typing Require Import lft_contexts.
From lrust.typing Require Export type.
From iris.prelude Require Import options.
Section sum.
Context `{!typeGS Σ}.
(* We define the actual empty type as being the empty sum, so that it is
convertible to it---and in particular, we can pattern-match on it
(as in, pattern-match in the language of lambda-rust, not in Coq). *)
Program Definition emp0 : type :=
{| ty_size := 1%nat;
ty_own tid vl := False%I;
ty_shr κ tid l := False%I |}.
Next Obligation. iIntros (tid vl) "[]". Qed.
Next Obligation.
iIntros (E κ l tid ??) "#LFT Hown Htok".
iMod (bor_acc with "LFT Hown Htok") as "[>H _]"; first done.
iDestruct "H" as (?) "[_ []]".
Qed.
Next Obligation. iIntros (κ κ' tid l) "#Hord []". Qed.
Definition is_pad i tyl (vl : list val) : iProp Σ :=
((nth i tyl emp0).(ty_size) + length vl)%nat = (max_list_with ty_size tyl)⌝%I.
Lemma split_sum_mt l tid q tyl :
(l ↦∗{q}: λ vl,
(i : nat) vl' vl'', vl = #i :: vl' ++ vl''
length vl = S (max_list_with ty_size tyl)
ty_own (nth i tyl emp0) tid vl')%I
⊣⊢ (i : nat), (l {q} #i
(l + (S $ (nth i tyl emp0).(ty_size))) ↦∗{q}: is_pad i tyl)
(l + 1) ↦∗{q}: (nth i tyl emp0).(ty_own) tid.
Proof.
iSplit; iIntros "H".
- iDestruct "H" as (vl) "[Hmt Hown]". iDestruct "Hown" as (i vl' vl'') "(% & % & Hown)".
subst. iExists i. iDestruct (heap_pointsto_vec_cons with "Hmt") as "[$ Hmt]".
iDestruct (ty_size_eq with "Hown") as "#EQ". iDestruct "EQ" as %Hvl'.
iDestruct (heap_pointsto_vec_app with "Hmt") as "[Hmt Htail]". iSplitL "Htail".
+ iExists vl''. rewrite (shift_loc_assoc_nat _ 1) Hvl'. iFrame. iPureIntro.
rewrite -Hvl'. simpl in *. rewrite -length_app. congruence.
+ iExists vl'. by iFrame.
- iDestruct "H" as (i) "[[Hmt1 Htail] Hown]".
iDestruct "Hown" as (vl') "[Hmt2 Hown]". iDestruct "Htail" as (vl'') "[Hmt3 %]".
iDestruct (ty_size_eq with "Hown") as "#EQ". iDestruct "EQ" as %Hvl'.
iExists (#i::vl'++vl'').
rewrite heap_pointsto_vec_cons heap_pointsto_vec_app (shift_loc_assoc_nat _ 1) Hvl'.
iFrame. iExists vl''. iSplit; first done. iPureIntro.
simpl. f_equal. by rewrite length_app Hvl'.
Qed.
Program Definition sum (tyl : list type) :=
{| ty_size := S (max_list_with ty_size tyl);
ty_own tid vl :=
( (i : nat) vl' vl'', vl = #i :: vl' ++ vl''
length vl = S (max_list_with ty_size tyl)
(nth i tyl emp0).(ty_own) tid vl')%I;
ty_shr κ tid l :=
( (i : nat),
&frac{κ} (λ q, l {q} #i
(l + (S $ (nth i tyl emp0).(ty_size))) ↦∗{q}: is_pad i tyl)
(nth i tyl emp0).(ty_shr) κ tid (l + 1))%I
|}.
Next Obligation.
iIntros (tyl tid vl) "Hown". iDestruct "Hown" as (i vl' vl'') "(%&%&_)".
subst. done.
Qed.
Next Obligation.
intros tyl E κ l tid. iIntros (??) "#LFT Hown Htok". rewrite split_sum_mt.
iMod (bor_exists with "LFT Hown") as (i) "Hown"; first solve_ndisj.
iMod (bor_sep with "LFT Hown") as "[Hmt Hown]"; first solve_ndisj.
iMod ((nth i tyl emp0).(ty_share) with "LFT Hown Htok") as "[#Hshr $]"; try done.
iMod (bor_fracture with "LFT [Hmt]") as "H'"; first solve_ndisj; last eauto.
by iFrame.
Qed.
Next Obligation.
iIntros (tyl κ κ' tid l) "#Hord H".
iDestruct "H" as (i) "[Hown0 Hown]". iExists i.
iSplitL "Hown0".
- by iApply (frac_bor_shorten with "Hord").
- iApply ((nth i tyl emp0).(ty_shr_mono) with "Hord"); done.
Qed.
Global Instance sum_wf tyl `{!ListTyWf tyl} : TyWf (sum tyl) :=
{ ty_lfts := tyl_lfts tyl; ty_wf_E := tyl_wf_E tyl }.
Global Instance sum_type_ne n : Proper (Forall2 (type_dist2 n) ==> type_dist2 n) sum.
Proof.
intros x y EQ.
assert (EQmax : max_list_with ty_size x = max_list_with ty_size y).
{ induction EQ as [|???? EQ _ IH]=>//=.
rewrite IH. f_equiv. apply EQ. }
(* TODO: If we had the right lemma relating nth, (Forall2 R) and R, we should
be able to make this more automatic. *)
assert (EQnth : i, type_dist2 n (nth i x emp0) (nth i y emp0)).
{ clear EQmax. induction EQ as [|???? EQ _ IH]=>-[|i] //=. }
constructor; simpl.
- by rewrite EQmax.
- intros tid vl. dist_later_fin_intro. rewrite /ty_own /= EQmax.
solve_proper_core ltac:(fun _ => exact:EQnth || f_type_equiv || f_equiv).
- intros κ tid l. unfold is_pad. rewrite EQmax.
solve_proper_core ltac:(fun _ => exact:EQnth || f_type_equiv || f_equiv).
Qed.
Global Instance sum_ne : NonExpansive sum.
Proof.
intros n x y EQ.
assert (EQmax : max_list_with ty_size x = max_list_with ty_size y).
{ induction EQ as [|???? EQ _ IH]=>//=.
rewrite IH. f_equiv. apply EQ. }
(* TODO: If we had the right lemma relating nth, (Forall2 R) and R, we should
be able to make this more automatic. *)
assert (EQnth : i, type_dist n (nth i x emp0) (nth i y emp0)).
{ clear EQmax. induction EQ as [|???? EQ _ IH]=>-[|i] //=. }
constructor; simpl.
- by rewrite EQmax.
- intros tid vl. rewrite EQmax.
solve_proper_core ltac:(fun _ => exact:EQnth || f_equiv).
- intros κ tid l. unfold is_pad. rewrite EQmax.
solve_proper_core ltac:(fun _ => exact:EQnth || (eapply ty_size_ne; try reflexivity) || f_equiv).
Qed.
Global Instance sum_mono E L :
Proper (Forall2 (subtype E L) ==> subtype E L) sum.
Proof.
iIntros (tyl1 tyl2 Htyl qmax qL) "HL".
iAssert ( (elctx_interp E -∗ max_list_with ty_size tyl1 = max_list_with ty_size tyl2))%I with "[#]" as "#Hleq".
{ iInduction Htyl as [|???? Hsub] "IH".
{ iClear "∗". iIntros "!> _". done. }
iDestruct (Hsub with "HL") as "#Hsub". iDestruct ("IH" with "HL") as "{IH} #IH".
iModIntro. iIntros "#HE". iDestruct ("Hsub" with "HE") as "(% & _ & _)".
iDestruct ("IH" with "HE") as %IH. iPureIntro. simpl. inversion_clear IH. by f_equal. }
iDestruct (subtype_Forall2_llctx_noend with "HL") as "#Htyl"; first done.
iClear "HL". iIntros "!> #HE".
iDestruct ("Hleq" with "HE") as %Hleq. iSpecialize ("Htyl" with "HE"). iClear "Hleq HE".
iAssert ( i, type_incl (nth i tyl1 emp0) (nth i tyl2 emp0))%I with "[#]" as "#Hty".
{ iIntros (i). edestruct (nth_lookup_or_length tyl1 i) as [Hl1|Hl]; last first.
{ rewrite nth_overflow // nth_overflow; first by iApply type_incl_refl.
by erewrite <-Forall2_length. }
edestruct @Forall2_lookup_l as (ty2 & Hl2 & _); [done..|].
iDestruct (big_sepL_lookup with "Htyl") as "Hty".
{ rewrite lookup_zip_with. erewrite Hl1. simpl.
rewrite Hl2 /=. done. }
rewrite (nth_lookup_Some tyl2 _ _ ty2) //. }
clear -Hleq. iClear "∗". iSplit; last iSplit.
- simpl. by rewrite Hleq.
- iModIntro. iIntros (tid vl) "H". iDestruct "H" as (i vl' vl'') "(% & % & Hown)".
iExists i, vl', vl''. iSplit; first done.
iSplit; first by rewrite -Hleq.
iDestruct ("Hty" $! i) as "(_ & #Htyi & _)". by iApply "Htyi".
- iModIntro. iIntros (κ tid l) "H". iDestruct "H" as (i) "(Hmt & Hshr)".
iExists i. iSplitR "Hshr".
+ rewrite /is_pad -Hleq. iDestruct ("Hty" $! i) as "(Hlen & _)".
iDestruct "Hlen" as %<-. done.
+ iDestruct ("Hty" $! i) as "(_ & _ & #Htyi)". by iApply "Htyi".
Qed.
Lemma sum_mono' E L tyl1 tyl2 :
Forall2 (subtype E L) tyl1 tyl2 subtype E L (sum tyl1) (sum tyl2).
Proof. apply sum_mono. Qed.
Global Instance sum_proper E L :
Proper (Forall2 (eqtype E L) ==> eqtype E L) sum.
Proof.
intros tyl1 tyl2 Heq; split; eapply sum_mono; [|rewrite -Forall2_flip];
(eapply Forall2_impl; [done|by intros ?? []]).
Qed.
Lemma sum_proper' E L tyl1 tyl2 :
Forall2 (eqtype E L) tyl1 tyl2 eqtype E L (sum tyl1) (sum tyl2).
Proof. apply sum_proper. Qed.
Lemma nth_empty {A : Type} i (d : A) :
nth i [] d = d.
Proof. by destruct i. Qed.
Global Instance sum_copy tyl : ListCopy tyl Copy (sum tyl).
Proof.
intros HFA. split.
- intros tid vl.
cut ( i vl', Persistent (ty_own (nth i tyl emp0) tid vl')); first by apply _.
intros. apply @copy_persistent.
edestruct nth_in_or_default as [| ->]; [by eapply List.Forall_forall| ].
split; first by apply _. iIntros (????????) "? []".
- intros κ tid E F l q ? HF.
iIntros "#LFT #H Htl [Htok1 Htok2]".
setoid_rewrite split_sum_mt. iDestruct "H" as (i) "[Hshr0 Hshr]".
iMod (frac_bor_acc with "LFT Hshr0 Htok1") as (q'1) "[>Hown Hclose]"; first solve_ndisj.
iAssert (( vl, is_pad i tyl vl)%I) with "[#]" as %[vl Hpad].
{ iDestruct "Hown" as "[_ Hpad]". iDestruct "Hpad" as (vl) "[_ %]".
eauto. }
iMod (@copy_shr_acc _ _ (nth i tyl emp0) with "LFT Hshr Htl Htok2")
as (q'2) "(Htl & HownC & Hclose')"; try done.
{ edestruct nth_in_or_default as [| ->]; first by eapply List.Forall_forall.
split; first by apply _. iIntros (????????) "? []". }
{ rewrite <-HF. simpl. rewrite <-union_subseteq_r.
apply shr_locsE_subseteq. lia. }
iDestruct (na_own_acc with "Htl") as "[$ Htlclose]".
{ apply difference_mono_l.
trans (shr_locsE (l + 1) (max_list_with ty_size tyl)).
- apply shr_locsE_subseteq. lia.
- set_solver+. }
destruct (Qp.lower_bound q'1 q'2) as (q' & q'01 & q'02 & -> & ->).
rewrite -(heap_pointsto_pred_op _ q' q'02); last (by intros; iApply ty_size_eq).
rewrite (fractional (Φ := λ q, _ {q} _ _ ↦∗{q}: _)%I).
iDestruct "HownC" as "[HownC1 HownC2]". iDestruct "Hown" as "[Hown1 Hown2]".
iExists q'. iModIntro. iSplitL "Hown1 HownC1".
+ iNext. iExists i. iFrame.
+ iIntros "Htl H". iDestruct "H" as (i') "[>Hown1 HownC1]".
iDestruct ("Htlclose" with "Htl") as "Htl".
iDestruct (heap_pointsto_agree with "[Hown1 Hown2]") as "#Heq".
{ iDestruct "Hown1" as "[$ _]". iDestruct "Hown2" as "[$ _]". }
iDestruct "Heq" as %[= ->%Nat2Z.inj].
iMod ("Hclose'" with "Htl [$HownC1 $HownC2]") as "[$ ?]".
iMod ("Hclose" with "[$Hown1 $Hown2]") as "$". by iFrame.
Qed.
Global Instance sum_send tyl : ListSend tyl Send (sum tyl).
Proof.
iIntros (Hsend tid1 tid2 vl) "H". iDestruct "H" as (i vl' vl'') "(% & % & Hown)".
iExists _, _, _. iSplit; first done. iSplit; first done. iApply @send_change_tid; last done.
edestruct nth_in_or_default as [| ->]; first by eapply List.Forall_forall.
iIntros (???) "[]".
Qed.
Global Instance sum_sync tyl : ListSync tyl Sync (sum tyl).
Proof.
iIntros (Hsync κ tid1 tid2 l) "H". iDestruct "H" as (i) "[Hframe Hown]".
iExists _. iFrame "Hframe". iApply @sync_change_tid; last done.
edestruct nth_in_or_default as [| ->]; first by eapply List.Forall_forall.
iIntros (????) "[]".
Qed.
Definition emp_type := sum [].
Global Instance emp_type_empty : Empty type := emp_type.
End sum.
(* Σ is commonly used for the current functor. So it cannot be defined
as Π for products. We stick to the following form. *)
Notation "Σ[ ty1 ; .. ; tyn ]" :=
(sum (cons ty1%T (..(cons tyn%T nil)..))) : lrust_type_scope.
Global Hint Opaque sum : lrust_typing lrust_typing_merge.
Global Hint Resolve sum_mono' sum_proper' : lrust_typing.
From iris.algebra Require Import numbers.
From iris.base_logic.lib Require Export na_invariants.
From lrust.lang Require Export proofmode notation.
From lrust.lifetime Require Export frac_borrow.
From lrust.typing Require Export base.
From lrust.typing Require Import lft_contexts.
From iris.prelude Require Import options.
Class typeGS Σ := TypeG {
type_lrustGS :: lrustGS Σ;
type_lftGS :: lftGS Σ lft_userE;
type_na_invG :: na_invG Σ;
type_frac_borrowG :: frac_borG Σ
}.
Definition lrustN := nroot .@ "lrust".
Definition shrN := lrustN .@ "shr".
Definition thread_id := na_inv_pool_name.
Record type `{!typeGS Σ} :=
{ ty_size : nat;
ty_own : thread_id list val iProp Σ;
ty_shr : lft thread_id loc iProp Σ;
ty_shr_persistent κ tid l : Persistent (ty_shr κ tid l);
ty_size_eq tid vl : ty_own tid vl length vl = ty_size;
(* The mask for starting the sharing does /not/ include the
namespace N, for allowing more flexibility for the user of
this type (typically for the [own] type). AFAIK, there is no
fundamental reason for this.
This should not be harmful, since sharing typically creates
invariants, which does not need the mask. Moreover, it is
more consistent with thread-local tokens, which we do not
give any.
The lifetime token is needed (a) to make the definition of simple types
nicer (they would otherwise require a "∨ □|=>[†κ]"), and (b) so that
we can have emp == sum [].
*)
ty_share E κ l tid q : lftN E
lft_ctx -∗ &{κ} (l ↦∗: ty_own tid) -∗ q.[κ] ={E}=∗
ty_shr κ tid l q.[κ];
ty_shr_mono κ κ' tid l :
κ' κ -∗ ty_shr κ tid l -∗ ty_shr κ' tid l
}.
Global Existing Instance ty_shr_persistent.
Global Instance: Params (@ty_size) 2 := {}.
Global Instance: Params (@ty_own) 2 := {}.
Global Instance: Params (@ty_shr) 2 := {}.
Global Arguments ty_own {_ _} !_ _ _ / : simpl nomatch.
Class TyWf `{!typeGS Σ} (ty : type) := { ty_lfts : list lft; ty_wf_E : elctx }.
Global Arguments ty_lfts {_ _} _ {_}.
Global Arguments ty_wf_E {_ _} _ {_}.
Definition ty_outlives_E `{!typeGS Σ} ty `{!TyWf ty} (κ : lft) : elctx :=
(λ α, κ α) <$> (ty_lfts ty).
Lemma ty_outlives_E_elctx_sat `{!typeGS Σ} E L ty `{!TyWf ty} α β :
ty_outlives_E ty β ⊆+ E
lctx_lft_incl E L α β
elctx_sat E L (ty_outlives_E ty α).
Proof.
unfold ty_outlives_E. induction (ty_lfts ty) as [|κ l IH]=>/= Hsub Hαβ.
- solve_typing.
- apply elctx_sat_lft_incl.
+ etrans; first done. eapply lctx_lft_incl_external, elem_of_submseteq, Hsub.
set_solver.
+ apply IH, Hαβ. etrans; last done. by apply submseteq_cons.
Qed.
(* Lift TyWf to lists. We cannot use `Forall` because that one is restricted to Prop. *)
Inductive ListTyWf `{!typeGS Σ} : list type Type :=
| list_ty_wf_nil : ListTyWf []
| list_ty_wf_cons ty tyl `{!TyWf ty, !ListTyWf tyl} : ListTyWf (ty::tyl).
Existing Class ListTyWf.
Global Existing Instances list_ty_wf_nil list_ty_wf_cons.
Fixpoint tyl_lfts `{!typeGS Σ} tyl {WF : ListTyWf tyl} : list lft :=
match WF with
| list_ty_wf_nil => []
| list_ty_wf_cons ty [] => ty_lfts ty
| list_ty_wf_cons ty tyl => ty_lfts ty ++ tyl_lfts tyl
end.
Fixpoint tyl_wf_E `{!typeGS Σ} tyl {WF : ListTyWf tyl} : elctx :=
match WF with
| list_ty_wf_nil => []
| list_ty_wf_cons ty [] => ty_wf_E ty
| list_ty_wf_cons ty tyl => ty_wf_E ty ++ tyl_wf_E tyl
end.
Fixpoint tyl_outlives_E `{!typeGS Σ} tyl {WF : ListTyWf tyl} (κ : lft) : elctx :=
match WF with
| list_ty_wf_nil => []
| list_ty_wf_cons ty [] => ty_outlives_E ty κ
| list_ty_wf_cons ty tyl => ty_outlives_E ty κ ++ tyl.(tyl_outlives_E) κ
end.
Lemma tyl_outlives_E_elctx_sat `{!typeGS Σ} E L tyl {WF : ListTyWf tyl} α β :
tyl_outlives_E tyl β ⊆+ E
lctx_lft_incl E L α β
elctx_sat E L (tyl_outlives_E tyl α).
Proof.
induction WF as [|? [] ?? IH]=>/=.
- solve_typing.
- intros. by eapply ty_outlives_E_elctx_sat.
- intros. apply elctx_sat_app, IH; [eapply ty_outlives_E_elctx_sat| |]=>//;
(etrans; [|done]); solve_typing.
Qed.
Record simple_type `{!typeGS Σ} :=
{ st_own : thread_id list val iProp Σ;
st_size_eq tid vl : st_own tid vl length vl = 1%nat;
st_own_persistent tid vl : Persistent (st_own tid vl) }.
Global Existing Instance st_own_persistent.
Global Instance: Params (@st_own) 2 := {}.
Program Definition ty_of_st `{!typeGS Σ} (st : simple_type) : type :=
{| ty_size := 1; ty_own := st.(st_own);
(* [st.(st_own) tid vl] needs to be outside of the fractured
borrow, otherwise I do not know how to prove the shr part of
[subtype_shr_mono]. *)
ty_shr := λ κ tid l,
( vl, &frac{κ} (λ q, l ↦∗{q} vl) st.(st_own) tid vl)%I
|}.
Next Obligation. intros. apply st_size_eq. Qed.
Next Obligation.
iIntros (?? st E κ l tid ??) "#LFT Hmt Hκ".
iMod (bor_exists with "LFT Hmt") as (vl) "Hmt"; first solve_ndisj.
iMod (bor_sep with "LFT Hmt") as "[Hmt Hown]"; first solve_ndisj.
iMod (bor_persistent with "LFT Hown Hκ") as "[Hown $]"; first solve_ndisj.
iMod (bor_fracture with "LFT [Hmt]") as "Hfrac"; by eauto with iFrame.
Qed.
Next Obligation.
iIntros (?? st κ κ' tid l) "#Hord H".
iDestruct "H" as (vl) "[#Hf #Hown]".
iExists vl. iFrame "Hown". by iApply (frac_bor_shorten with "Hord").
Qed.
Coercion ty_of_st : simple_type >-> type.
Declare Scope lrust_type_scope.
Delimit Scope lrust_type_scope with T.
Bind Scope lrust_type_scope with type.
(* OFE and COFE structures on types and simple types. *)
Section ofe.
Context `{!typeGS Σ}.
Inductive type_equiv' (ty1 ty2 : type) : Prop :=
Type_equiv :
ty1.(ty_size) = ty2.(ty_size)
( tid vs, ty1.(ty_own) tid vs ty2.(ty_own) tid vs)
( κ tid l, ty1.(ty_shr) κ tid l ty2.(ty_shr) κ tid l)
type_equiv' ty1 ty2.
Local Instance type_equiv : Equiv type := type_equiv'.
Inductive type_dist' (n : nat) (ty1 ty2 : type) : Prop :=
Type_dist :
ty1.(ty_size) = ty2.(ty_size)
( tid vs, ty1.(ty_own) tid vs {n} ty2.(ty_own) tid vs)
( κ tid l, ty1.(ty_shr) κ tid l {n} ty2.(ty_shr) κ tid l)
type_dist' n ty1 ty2.
Local Instance type_dist : Dist type := type_dist'.
Let T := prodO
(prodO natO (thread_id -d> list val -d> iPropO Σ))
(lft -d> thread_id -d> loc -d> iPropO Σ).
Let P (x : T) : Prop :=
( κ tid l, Persistent (x.2 κ tid l))
( tid vl, x.1.2 tid vl length vl = x.1.1)
( E κ l tid q, lftN E
lft_ctx -∗ &{κ} (l ↦∗: λ vs, x.1.2 tid vs) -∗
q.[κ] ={E}=∗ x.2 κ tid l q.[κ])
( κ κ' tid l, κ' κ -∗ x.2 κ tid l -∗ x.2 κ' tid l).
Definition type_unpack (ty : type) : T :=
(ty.(ty_size), λ tid vl, (ty.(ty_own) tid vl), ty.(ty_shr)).
Program Definition type_pack (x : T) (H : P x) : type :=
{| ty_size := x.1.1; ty_own tid vl := x.1.2 tid vl; ty_shr := x.2 |}.
Solve Obligations with by intros [[??] ?] (?&?&?&?).
Definition type_ofe_mixin : OfeMixin type.
Proof.
apply (iso_ofe_mixin type_unpack).
- split; [by destruct 1|by intros [[??] ?]; constructor].
- split; [by destruct 1|by intros [[??] ?]; constructor].
Qed.
Canonical Structure typeO : ofe := Ofe type type_ofe_mixin.
Global Instance ty_size_ne n : Proper (dist n ==> eq) ty_size.
Proof. intros ?? EQ. apply EQ. Qed.
Global Instance ty_size_proper : Proper (() ==> eq) ty_size.
Proof. intros ?? EQ. apply EQ. Qed.
Global Instance ty_own_ne n:
Proper (dist n ==> eq ==> eq ==> dist n) ty_own.
Proof. intros ?? EQ ??-> ??->. apply EQ. Qed.
Global Instance ty_own_proper : Proper (() ==> eq ==> eq ==> ()) ty_own.
Proof. intros ?? EQ ??-> ??->. apply EQ. Qed.
Global Instance ty_shr_ne n :
Proper (dist n ==> eq ==> eq ==> eq ==> dist n) ty_shr.
Proof. intros ?? EQ ??-> ??-> ??->. apply EQ. Qed.
Global Instance ty_shr_proper :
Proper (() ==> eq ==> eq ==> eq ==> ()) ty_shr.
Proof. intros ?? EQ ??-> ??-> ??->. apply EQ. Qed.
Global Instance type_cofe : Cofe typeO.
Proof.
apply (iso_cofe_subtype' P type_pack type_unpack).
- by intros [].
- split; [by destruct 1|by intros [[??] ?]; constructor].
- by intros [].
- (* TODO: automate this *)
repeat apply limit_preserving_and; repeat (apply limit_preserving_forall; intros ?).
+ apply bi.limit_preserving_Persistent=> n ty1 ty2 Hty; apply Hty.
+ apply bi.limit_preserving_entails=> n ty1 ty2 Hty; last by rewrite Hty. apply Hty.
+ apply bi.limit_preserving_emp_valid=> n ty1 ty2 Hty. repeat f_equiv; apply Hty.
+ apply bi.limit_preserving_emp_valid=> n ty1 ty2 Hty. repeat f_equiv; apply Hty.
Qed.
Inductive st_equiv' (ty1 ty2 : simple_type) : Prop :=
St_equiv :
( tid vs, ty1.(ty_own) tid vs ty2.(ty_own) tid vs)
st_equiv' ty1 ty2.
Local Instance st_equiv : Equiv simple_type := st_equiv'.
Inductive st_dist' (n : nat) (ty1 ty2 : simple_type) : Prop :=
St_dist :
( tid vs, ty1.(ty_own) tid vs {n} (ty2.(ty_own) tid vs))
st_dist' n ty1 ty2.
Local Instance st_dist : Dist simple_type := st_dist'.
Definition st_unpack (ty : simple_type) : thread_id -d> list val -d> iPropO Σ :=
λ tid vl, ty.(ty_own) tid vl.
Definition st_ofe_mixin : OfeMixin simple_type.
Proof.
apply (iso_ofe_mixin st_unpack).
- split; [by destruct 1|by constructor].
- split; [by destruct 1|by constructor].
Qed.
Canonical Structure stO : ofe := Ofe simple_type st_ofe_mixin.
Global Instance st_own_ne n :
Proper (dist n ==> eq ==> eq ==> dist n) st_own.
Proof. intros ?? EQ ??-> ??->. apply EQ. Qed.
Global Instance st_own_proper : Proper (() ==> eq ==> eq ==> ()) st_own.
Proof. intros ?? EQ ??-> ??->. apply EQ. Qed.
Global Instance ty_of_st_ne : NonExpansive ty_of_st.
Proof.
intros n ?? EQ. constructor; try apply EQ; first done.
- simpl. intros; repeat f_equiv. apply EQ.
Qed.
Global Instance ty_of_st_proper : Proper (() ==> ()) ty_of_st.
Proof. apply (ne_proper _). Qed.
End ofe.
(** Special metric for type-nonexpansive and Type-contractive functions. *)
Section type_dist2.
Context `{!typeGS Σ}.
(* Size and shr are n-equal, but own is only n-1-equal.
We need this to express what shr has to satisfy on a Type-NE-function:
It may only depend contractively on own. *)
(* TODO: Find a better name for this metric. *)
Inductive type_dist2 (n : nat) (ty1 ty2 : type) : Prop :=
Type_dist2 :
ty1.(ty_size) = ty2.(ty_size)
( tid vs, dist_later n (ty1.(ty_own) tid vs) (ty2.(ty_own) tid vs))
( κ tid l, ty1.(ty_shr) κ tid l {n} ty2.(ty_shr) κ tid l)
type_dist2 n ty1 ty2.
Global Instance type_dist2_equivalence n : Equivalence (type_dist2 n).
Proof.
constructor.
- by constructor.
- intros ?? Heq; constructor; symmetry; eapply Heq.
- intros ??? Heq1 Heq2; constructor; etrans; (eapply Heq1 || eapply Heq2).
Qed.
Definition type_dist2_later (n : nat) ty1 ty2 : Prop :=
match n with O => True | S n => type_dist2 n ty1 ty2 end.
Global Arguments type_dist2_later !_ _ _ /.
Global Instance type_dist2_later_equivalence n :
Equivalence (type_dist2_later n).
Proof. destruct n as [|n]; first by split. apply type_dist2_equivalence. Qed.
(* The hierarchy of metrics:
dist n → type_dist2 n → dist_later n → type_dist2_later n. *)
Lemma type_dist_dist2 n ty1 ty2 : dist n ty1 ty2 type_dist2 n ty1 ty2.
Proof. intros EQ. split; intros; try apply dist_dist_later; apply EQ. Qed.
Lemma type_dist2_dist_later n ty1 ty2 : type_dist2 n ty1 ty2 dist_later n ty1 ty2.
Proof.
intros EQ. dist_later_fin_intro.
split; intros; try apply EQ; eauto using SIdx.lt_succ_diag_r.
apply dist_S, EQ.
Qed.
Lemma type_later_dist2_later n ty1 ty2 : dist_later n ty1 ty2 type_dist2_later n ty1 ty2.
Proof. destruct n; first done. rewrite dist_later_fin_iff. exact: type_dist_dist2. Qed.
Lemma type_dist2_dist n ty1 ty2 : type_dist2 (S n) ty1 ty2 dist n ty1 ty2.
Proof. move=>/type_dist2_dist_later. rewrite dist_later_fin_iff. done. Qed.
Lemma type_dist2_S n ty1 ty2 : type_dist2 (S n) ty1 ty2 type_dist2 n ty1 ty2.
Proof. intros. apply type_dist_dist2, type_dist2_dist. done. Qed.
Lemma ty_size_type_dist n : Proper (type_dist2 n ==> eq) ty_size.
Proof. intros ?? EQ. apply EQ. Qed.
Lemma ty_own_type_dist n:
Proper (type_dist2 (S n) ==> eq ==> eq ==> dist n) ty_own.
Proof. intros ?? EQ ??-> ??->. apply EQ, SIdx.lt_succ_diag_r. Qed.
Lemma ty_shr_type_dist n :
Proper (type_dist2 n ==> eq ==> eq ==> eq ==> dist n) ty_shr.
Proof. intros ?? EQ ??-> ??-> ??->. apply EQ. Qed.
End type_dist2.
(** Type-nonexpansive and Type-contractive functions. *)
(* Note that TypeContractive is neither weaker nor stronger than Contractive, because
(a) it allows the dependency of own on shr to be non-expansive, and
(b) it forces the dependency of shr on own to be doubly-contractive.
It would be possible to weaken this so that no double-contractivity is required.
However, then it is no longer possible to write TypeContractive as just a
Proper, which makes it significantly more annoying to use.
For similar reasons, TypeNonExpansive is incomparable to NonExpansive.
*)
Notation TypeNonExpansive T := ( n, Proper (type_dist2 n ==> type_dist2 n) T).
Notation TypeContractive T := ( n, Proper (type_dist2_later n ==> type_dist2 n) T).
Section type_contractive.
Context `{!typeGS Σ}.
Lemma type_ne_dist_later T :
TypeNonExpansive T n, Proper (type_dist2_later n ==> type_dist2_later n) T.
Proof. intros Hf [|n]; last exact: Hf. hnf. by intros. Qed.
(* From the above, it easily follows that TypeNonExpansive functions compose with
TypeNonExpansive and with TypeContractive functions. *)
Lemma type_ne_ne_compose T1 T2 :
TypeNonExpansive T1 TypeNonExpansive T2 TypeNonExpansive (T1 T2).
Proof. intros NE1 NE2 ? ???; simpl. apply: NE1. exact: NE2. Qed.
Lemma type_contractive_compose_right T1 T2 :
TypeContractive T1 TypeNonExpansive T2 TypeContractive (T1 T2).
Proof. intros HT1 HT2 ? ???. apply: HT1. exact: type_ne_dist_later. Qed.
Lemma type_contractive_compose_left T1 T2 :
TypeNonExpansive T1 TypeContractive T2 TypeContractive (T1 T2).
Proof. intros HT1 HT2 ? ???; simpl. apply: HT1. exact: HT2. Qed.
(* Show some more relationships between properties. *)
Lemma type_contractive_type_ne T :
TypeContractive T TypeNonExpansive T.
Proof.
intros HT ? ???. eapply type_dist_dist2, dist_later_S, type_dist2_dist_later, HT. done.
Qed.
Lemma type_contractive_ne T :
TypeContractive T NonExpansive T.
Proof.
intros HT ? ???. apply dist_later_S, type_dist2_dist_later, HT, type_dist_dist2. done.
Qed.
(* Simple types *)
Global Instance ty_of_st_type_ne n :
Proper (dist_later n ==> type_dist2 n) ty_of_st.
Proof.
intros ?? Hdst. constructor.
- done.
- intros. dist_later_intro. eapply Hdst.
- intros. solve_contractive.
Qed.
End type_contractive.
(* Tactic automation. *)
Ltac f_type_equiv :=
first [ ((eapply ty_size_type_dist || eapply ty_shr_type_dist || eapply ty_own_type_dist); try reflexivity) |
dist_later_fin_intro ].
Ltac solve_type_proper :=
constructor;
solve_proper_core ltac:(fun _ => f_type_equiv || f_contractive_fin || f_equiv).
Fixpoint shr_locsE (l : loc) (n : nat) : coPset :=
match n with
| 0%nat =>
| S n => shrN.@l shr_locsE (l + 1%nat) n
end.
Class Copy `{!typeGS Σ} (t : type) := {
copy_persistent tid vl : Persistent (t.(ty_own) tid vl);
copy_shr_acc κ tid E F l q :
lftN shrN E shr_locsE l (t.(ty_size) + 1) F
lft_ctx -∗ t.(ty_shr) κ tid l -∗ na_own tid F -∗ q.[κ] ={E}=∗
q', na_own tid (F shr_locsE l t.(ty_size))
(l ↦∗{q'}: t.(ty_own) tid)
(na_own tid (F shr_locsE l t.(ty_size)) -∗ l ↦∗{q'}: t.(ty_own) tid
={E}=∗ na_own tid F q.[κ])
}.
Global Existing Instances copy_persistent.
Global Instance: Params (@Copy) 2 := {}.
Class ListCopy `{!typeGS Σ} (tys : list type) := lst_copy : Forall Copy tys.
Global Instance: Params (@ListCopy) 2 := {}.
Global Instance lst_copy_nil `{!typeGS Σ} : ListCopy [] := List.Forall_nil _.
Global Instance lst_copy_cons `{!typeGS Σ} ty tys :
Copy ty ListCopy tys ListCopy (ty :: tys) := List.Forall_cons _ _ _.
Class Send `{!typeGS Σ} (t : type) :=
send_change_tid tid1 tid2 vl : t.(ty_own) tid1 vl t.(ty_own) tid2 vl.
Global Instance: Params (@Send) 2 := {}.
Class ListSend `{!typeGS Σ} (tys : list type) := lst_send : Forall Send tys.
Global Instance: Params (@ListSend) 2 := {}.
Global Instance lst_send_nil `{!typeGS Σ} : ListSend [] := List.Forall_nil _.
Global Instance lst_send_cons `{!typeGS Σ} ty tys :
Send ty ListSend tys ListSend (ty :: tys) := List.Forall_cons _ _ _.
Class Sync `{!typeGS Σ} (t : type) :=
sync_change_tid κ tid1 tid2 l : t.(ty_shr) κ tid1 l t.(ty_shr) κ tid2 l.
Global Instance: Params (@Sync) 2 := {}.
Class ListSync `{!typeGS Σ} (tys : list type) := lst_sync : Forall Sync tys.
Global Instance: Params (@ListSync) 2 := {}.
Global Instance lst_sync_nil `{!typeGS Σ} : ListSync [] := List.Forall_nil _.
Global Instance lst_sync_cons `{!typeGS Σ} ty tys :
Sync ty ListSync tys ListSync (ty :: tys) := List.Forall_cons _ _ _.
Section type.
Context `{!typeGS Σ}.
(** Copy types *)
Lemma shr_locsE_shift l n m :
shr_locsE l (n + m) = shr_locsE l n shr_locsE (l + n) m.
Proof.
revert l; induction n as [|n IHn]; intros l.
- rewrite shift_loc_0. set_solver+.
- rewrite -Nat.add_1_l Nat2Z.inj_add /= IHn shift_loc_assoc.
set_solver+.
Qed.
Lemma shr_locsE_disj l n m :
shr_locsE l n ## shr_locsE (l + n) m.
Proof.
revert l; induction n as [|n IHn]; intros l.
- simpl. set_solver+.
- rewrite -Nat.add_1_l Nat2Z.inj_add /=.
apply disjoint_union_l. split; last (rewrite -shift_loc_assoc; exact: IHn).
clear IHn. revert n; induction m as [|m IHm]; intros n; simpl; first set_solver+.
rewrite shift_loc_assoc. apply disjoint_union_r. split.
+ apply ndot_ne_disjoint. destruct l. intros [=]. lia.
+ rewrite -Z.add_assoc. move:(IHm (n + 1)%nat). rewrite Nat2Z.inj_add //.
Qed.
Lemma shr_locsE_shrN l n :
shr_locsE l n shrN.
Proof.
revert l; induction n=>l /=; first by set_solver+.
apply union_least; last by auto. solve_ndisj.
Qed.
Lemma shr_locsE_subseteq l n m :
(n m)%nat shr_locsE l n shr_locsE l m.
Proof.
induction 1; first done. etrans; first done.
rewrite -Nat.add_1_l [(_ + _)%nat]comm_L shr_locsE_shift. set_solver+.
Qed.
Lemma shr_locsE_split_tok l n m tid :
na_own tid (shr_locsE l (n + m)) ⊣⊢
na_own tid (shr_locsE l n) na_own tid (shr_locsE (l + n) m).
Proof.
rewrite shr_locsE_shift na_own_union //. apply shr_locsE_disj.
Qed.
Global Instance copy_equiv : Proper (equiv ==> impl) Copy.
Proof.
intros ty1 ty2 [EQsz EQown EQshr] Hty1. split.
- intros. rewrite -EQown. apply _.
- intros *. rewrite -EQsz -EQshr. setoid_rewrite <-EQown.
apply copy_shr_acc.
Qed.
Global Program Instance ty_of_st_copy st : Copy (ty_of_st st).
Next Obligation.
iIntros (st κ tid E ? l q ? HF) "#LFT #Hshr Htok Hlft".
iDestruct (na_own_acc with "Htok") as "[$ Htok]"; first solve_ndisj.
iDestruct "Hshr" as (vl) "[Hf Hown]".
iMod (frac_bor_acc with "LFT Hf Hlft") as (q') "[Hmt Hclose]"; first solve_ndisj.
iModIntro. iExists _. iDestruct "Hmt" as "[Hmt1 Hmt2]".
iSplitL "Hmt1"; first by auto with iFrame.
iIntros "Htok2 Hmt1". iDestruct "Hmt1" as (vl') "[Hmt1 #Hown']".
iDestruct ("Htok" with "Htok2") as "$".
iAssert ( length vl = length vl')%I as ">%".
{ iNext. iDestruct (st_size_eq with "Hown") as %->.
iDestruct (st_size_eq with "Hown'") as %->. done. }
iCombine "Hmt1" "Hmt2" as "Hmt". rewrite heap_pointsto_vec_op // Qp.div_2.
iDestruct "Hmt" as "[>% Hmt]". subst. by iApply "Hclose".
Qed.
(** Send and Sync types *)
Global Instance send_equiv : Proper (equiv ==> impl) Send.
Proof.
intros ty1 ty2 [EQsz EQown EQshr] Hty1.
rewrite /Send=>???. rewrite -!EQown. auto.
Qed.
Global Instance sync_equiv : Proper (equiv ==> impl) Sync.
Proof.
intros ty1 ty2 [EQsz EQown EQshr] Hty1.
rewrite /Send=>????. rewrite -!EQshr. auto.
Qed.
Global Instance ty_of_st_sync st : Send (ty_of_st st) Sync (ty_of_st st).
Proof.
iIntros (Hsend κ tid1 tid2 l). iDestruct 1 as (vl) "[Hm Hown]".
iExists vl. iFrame "Hm". iNext. by iApply Hsend.
Qed.
Lemma send_change_tid' t tid1 tid2 vl :
Send t t.(ty_own) tid1 vl t.(ty_own) tid2 vl.
Proof.
intros ?. apply: anti_symm; apply send_change_tid.
Qed.
Lemma sync_change_tid' t κ tid1 tid2 l :
Sync t t.(ty_shr) κ tid1 l t.(ty_shr) κ tid2 l.
Proof.
intros ?. apply: anti_symm; apply sync_change_tid.
Qed.
End type.
(** iProp-level type inclusion / equality. *)
Definition type_incl `{!typeGS Σ} (ty1 ty2 : type) : iProp Σ :=
(ty1.(ty_size) = ty2.(ty_size)
( tid vl, ty1.(ty_own) tid vl -∗ ty2.(ty_own) tid vl)
( κ tid l, ty1.(ty_shr) κ tid l -∗ ty2.(ty_shr) κ tid l))%I.
Global Instance: Params (@type_incl) 2 := {}.
(* Typeclasses Opaque type_incl. *)
Definition type_equal `{!typeGS Σ} (ty1 ty2 : type) : iProp Σ :=
(ty1.(ty_size) = ty2.(ty_size)
( tid vl, ty1.(ty_own) tid vl ∗-∗ ty2.(ty_own) tid vl)
( κ tid l, ty1.(ty_shr) κ tid l ∗-∗ ty2.(ty_shr) κ tid l))%I.
Global Instance: Params (@type_equal) 2 := {}.
Section iprop_subtyping.
Context `{!typeGS Σ}.
Global Instance type_incl_ne : NonExpansive2 type_incl.
Proof.
intros n ?? [EQsz1 EQown1 EQshr1] ?? [EQsz2 EQown2 EQshr2].
rewrite /type_incl. repeat ((by auto) || f_equiv).
Qed.
Global Instance type_incl_proper :
Proper (() ==> () ==> (⊣⊢)) type_incl.
Proof. apply ne_proper_2, _. Qed.
Global Instance type_incl_persistent ty1 ty2 : Persistent (type_incl ty1 ty2) := _.
Lemma type_incl_refl ty : type_incl ty ty.
Proof. iSplit; first done. iSplit; iModIntro; iIntros; done. Qed.
Lemma type_incl_trans ty1 ty2 ty3 :
type_incl ty1 ty2 -∗ type_incl ty2 ty3 -∗ type_incl ty1 ty3.
Proof.
iIntros "(% & #Ho12 & #Hs12) (% & #Ho23 & #Hs23)".
iSplit; first (iPureIntro; etrans; done).
iSplit; iModIntro; iIntros.
- iApply "Ho23". iApply "Ho12". done.
- iApply "Hs23". iApply "Hs12". done.
Qed.
Global Instance type_equal_ne : NonExpansive2 type_equal.
Proof.
intros n ?? [EQsz1 EQown1 EQshr1] ?? [EQsz2 EQown2 EQshr2].
rewrite /type_equal. repeat ((by auto) || f_equiv).
Qed.
Global Instance type_equal_proper :
Proper (() ==> () ==> (⊣⊢)) type_equal.
Proof. apply ne_proper_2, _. Qed.
Global Instance type_equal_persistent ty1 ty2 : Persistent (type_equal ty1 ty2) := _.
Lemma type_equal_incl ty1 ty2 :
type_equal ty1 ty2 ⊣⊢ type_incl ty1 ty2 type_incl ty2 ty1.
Proof.
iSplit.
- iIntros "#(% & Ho & Hs)".
iSplit; (iSplit; first done; iSplit; iModIntro).
+ iIntros (??) "?". by iApply "Ho".
+ iIntros (???) "?". by iApply "Hs".
+ iIntros (??) "?". by iApply "Ho".
+ iIntros (???) "?". by iApply "Hs".
- iIntros "#[(% & Ho1 & Hs1) (% & Ho2 & Hs2)]".
iSplit; first done. iSplit; iModIntro.
+ iIntros (??). iSplit; [iApply "Ho1"|iApply "Ho2"].
+ iIntros (???). iSplit; [iApply "Hs1"|iApply "Hs2"].
Qed.
Lemma type_equal_refl ty :
type_equal ty ty.
Proof.
iApply type_equal_incl. iSplit; iApply type_incl_refl.
Qed.
Lemma type_equal_trans ty1 ty2 ty3 :
type_equal ty1 ty2 -∗ type_equal ty2 ty3 -∗ type_equal ty1 ty3.
Proof.
rewrite !type_equal_incl. iIntros "#[??] #[??]". iSplit.
- iApply (type_incl_trans _ ty2); done.
- iApply (type_incl_trans _ ty2); done.
Qed.
Lemma type_incl_simple_type (st1 st2 : simple_type) :
( tid vl, st1.(st_own) tid vl -∗ st2.(st_own) tid vl) -∗
type_incl st1 st2.
Proof.
iIntros "#Hst". iSplit; first done. iSplit; iModIntro.
- iIntros. iApply "Hst"; done.
- iIntros (???). iDestruct 1 as (vl) "[Hf Hown]". iExists vl. iFrame "Hf".
by iApply "Hst".
Qed.
Lemma type_equal_equiv ty1 ty2 :
( type_equal ty1 ty2) ty1 ty2.
Proof.
split.
- intros Heq. split.
+ eapply uPred.pure_soundness. iDestruct Heq as "[%Hsz _]". done.
+ iDestruct Heq as "[_ [Hown _]]". done.
+ iDestruct Heq as "[_ [_ Hshr]]". done.
- intros ->. apply type_equal_refl.
Qed.
End iprop_subtyping.
(** Prop-level conditional type inclusion/equality
(may depend on assumptions in [E, L].) *)
Definition subtype `{!typeGS Σ} (E : elctx) (L : llctx) (ty1 ty2 : type) : Prop :=
qmax qL, llctx_interp_noend qmax L qL (elctx_interp E -∗ type_incl ty1 ty2).
Global Instance: Params (@subtype) 4 := {}.
(* TODO: The prelude should have a symmetric closure. *)
Definition eqtype `{!typeGS Σ} (E : elctx) (L : llctx) (ty1 ty2 : type) : Prop :=
subtype E L ty1 ty2 subtype E L ty2 ty1.
Global Instance: Params (@eqtype) 4 := {}.
Section subtyping.
Context `{!typeGS Σ}.
Lemma type_incl_subtype ty1 ty2 E L :
( type_incl ty1 ty2) subtype E L ty1 ty2.
Proof.
intros Heq. rewrite /subtype.
iIntros (??) "_ !> _". done.
Qed.
Lemma subtype_refl E L ty : subtype E L ty ty.
Proof. iIntros (??) "_ !> _". iApply type_incl_refl. Qed.
Global Instance subtype_preorder E L : PreOrder (subtype E L).
Proof.
split; first by intros ?; apply subtype_refl.
iIntros (ty1 ty2 ty3 H12 H23 ??) "HL".
iDestruct (H12 with "HL") as "#H12".
iDestruct (H23 with "HL") as "#H23".
iClear "∗". iIntros "!> #HE".
iApply (type_incl_trans with "[#]"); first by iApply "H12". by iApply "H23".
Qed.
Lemma subtype_Forall2_llctx_noend E L tys1 tys2 qmax qL :
Forall2 (subtype E L) tys1 tys2
llctx_interp_noend qmax L qL -∗ (elctx_interp E -∗
[ list] tys (zip tys1 tys2), type_incl (tys.1) (tys.2)).
Proof.
iIntros (Htys) "HL".
iAssert ([ list] tys zip tys1 tys2,
(elctx_interp _ -∗ type_incl (tys.1) (tys.2)))%I as "#Htys".
{ iApply big_sepL_forall. iIntros (k [ty1 ty2] Hlookup).
move:Htys => /Forall2_Forall /Forall_forall=>Htys.
iApply (Htys (ty1, ty2)); first by exact: elem_of_list_lookup_2. done. }
iClear "∗". iIntros "!> #HE". iApply (big_sepL_impl with "[$Htys]").
iIntros "!> * % #Hincl". by iApply "Hincl".
Qed.
Lemma subtype_Forall2_llctx E L tys1 tys2 qmax :
Forall2 (subtype E L) tys1 tys2
llctx_interp qmax L -∗ (elctx_interp E -∗
[ list] tys (zip tys1 tys2), type_incl (tys.1) (tys.2)).
Proof.
iIntros (?) "HL". iApply subtype_Forall2_llctx_noend; first done.
iDestruct (llctx_interp_acc_noend with "HL") as "[$ _]".
Qed.
Lemma lft_invariant_subtype E L T :
Proper (lctx_lft_eq E L ==> subtype E L) T.
Proof.
iIntros (x y [Hxy Hyx] qmax qL) "L".
iPoseProof (Hxy with "L") as "#Hxy".
iPoseProof (Hyx with "L") as "#Hyx".
iIntros "!> #E". clear Hxy Hyx.
iDestruct ("Hxy" with "E") as %Hxy.
iDestruct ("Hyx" with "E") as %Hyx.
iClear "Hyx Hxy".
rewrite (anti_symm _ _ _ Hxy Hyx).
iApply type_incl_refl.
Qed.
Lemma lft_invariant_eqtype E L T :
Proper (lctx_lft_eq E L ==> eqtype E L) T.
Proof. split; by apply lft_invariant_subtype. Qed.
Lemma equiv_subtype E L ty1 ty2 : ty1 ty2 subtype E L ty1 ty2.
Proof. unfold subtype, type_incl=>EQ. setoid_rewrite EQ. apply subtype_refl. Qed.
Lemma eqtype_unfold E L ty1 ty2 :
eqtype E L ty1 ty2
( qmax qL, llctx_interp_noend qmax L qL -∗ (elctx_interp E -∗ type_equal ty1 ty2)).
Proof.
split.
- iIntros ([EQ1 EQ2] qmax qL) "HL".
iDestruct (EQ1 with "HL") as "#EQ1".
iDestruct (EQ2 with "HL") as "#EQ2".
iClear "∗". iIntros "!> #HE".
iDestruct ("EQ1" with "HE") as "[#Hsz [#H1own #H1shr]]".
iDestruct ("EQ2" with "HE") as "[_ [#H2own #H2shr]]".
iSplit; last iSplit.
+ done.
+ by iIntros "!>*"; iSplit; iIntros "H"; [iApply "H1own"|iApply "H2own"].
+ by iIntros "!>*"; iSplit; iIntros "H"; [iApply "H1shr"|iApply "H2shr"].
- intros EQ. split; (iIntros (qmax qL) "HL";
iDestruct (EQ with "HL") as "#EQ";
iClear "∗"; iIntros "!> #HE";
iDestruct ("EQ" with "HE") as "[% [#Hown #Hshr]]";
(iSplit; last iSplit)).
+ done.
+ iIntros "!>* H". by iApply "Hown".
+ iIntros "!>* H". by iApply "Hshr".
+ done.
+ iIntros "!>* H". by iApply "Hown".
+ iIntros "!>* H". by iApply "Hshr".
Qed.
Lemma type_equal_eqtype ty1 ty2 E L :
( type_equal ty1 ty2) eqtype E L ty1 ty2.
Proof.
intros Heq. apply eqtype_unfold.
iIntros (??) "_ !> _". done.
Qed.
Lemma eqtype_refl E L ty : eqtype E L ty ty.
Proof. by split. Qed.
Lemma equiv_eqtype E L ty1 ty2 : ty1 ty2 eqtype E L ty1 ty2.
Proof. by split; apply equiv_subtype. Qed.
Global Instance subtype_proper E L :
Proper (eqtype E L ==> eqtype E L ==> iff) (subtype E L).
Proof. intros ??[] ??[]. split; intros; by etrans; [|etrans]. Qed.
Global Instance eqtype_equivalence E L : Equivalence (eqtype E L).
Proof.
split.
- by split.
- intros ?? Heq; split; apply Heq.
- intros ??? H1 H2. split; etrans; (apply H1 || apply H2).
Qed.
Lemma subtype_simple_type E L (st1 st2 : simple_type) :
( qmax qL, llctx_interp_noend qmax L qL -∗ (elctx_interp E -∗
tid vl, st1.(st_own) tid vl -∗ st2.(st_own) tid vl))
subtype E L st1 st2.
Proof.
intros Hst. iIntros (qmax qL) "HL". iDestruct (Hst with "HL") as "#Hst".
iClear "∗". iIntros "!> #HE". iApply type_incl_simple_type.
iIntros "!>" (??) "?". by iApply "Hst".
Qed.
Lemma subtype_weaken E1 E2 L1 L2 ty1 ty2 :
E1 ⊆+ E2 L1 ⊆+ L2
subtype E1 L1 ty1 ty2 subtype E2 L2 ty1 ty2.
Proof.
iIntros (HE12 ? Hsub qmax qL) "HL". iDestruct (Hsub with "[HL]") as "#Hsub".
{ rewrite /llctx_interp. by iApply big_sepL_submseteq. }
iClear "∗". iIntros "!> #HE". iApply "Hsub".
rewrite /elctx_interp. by iApply big_sepL_submseteq.
Qed.
End subtyping.
Section type_util.
Context `{!typeGS Σ}.
Lemma heap_pointsto_ty_own l ty tid :
l ↦∗: ty_own ty tid ⊣⊢ (vl : vec val ty.(ty_size)), l ↦∗ vl ty_own ty tid vl.
Proof.
iSplit.
- iIntros "H". iDestruct "H" as (vl) "[Hl Hown]".
iDestruct (ty_size_eq with "Hown") as %<-.
iExists (list_to_vec vl). rewrite vec_to_list_to_vec. iFrame.
- iIntros "H". iDestruct "H" as (vl) "[Hl Hown]". eauto with iFrame.
Qed.
End type_util.
Global Hint Resolve ty_outlives_E_elctx_sat tyl_outlives_E_elctx_sat : lrust_typing.
Global Hint Resolve subtype_refl eqtype_refl : lrust_typing.
Global Hint Opaque subtype eqtype : lrust_typing.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import type lft_contexts.
From iris.prelude Require Import options.
Definition path := expr.
Bind Scope expr_scope with path.
(* TODO: Consider making this a pair of a path and the rest. We could
then e.g. formulate tctx_elt_hasty_path more generally. *)
Inductive tctx_elt `{!typeGS Σ} : Type :=
| TCtx_hasty (p : path) (ty : type)
| TCtx_blocked (p : path) (κ : lft) (ty : type).
Notation tctx := (list tctx_elt).
Notation "p ◁ ty" := (TCtx_hasty p ty%list%T) (at level 70).
Notation "p ◁{ κ } ty" := (TCtx_blocked p κ ty)
(at level 70, format "p ◁{ κ } ty").
Section type_context.
Context `{!typeGS Σ}.
Implicit Types T : tctx.
Fixpoint eval_path (p : path) : option val :=
match p with
| BinOp OffsetOp e (Lit (LitInt n)) =>
match eval_path e with
| Some (#(LitLoc l)) => Some (#(l + n))
| _ => None
end
| e => to_val e
end.
Lemma eval_path_of_val (v : val) :
eval_path v = Some v.
Proof. destruct v; first done. simpl. rewrite (decide_True_pi _). done. Qed.
Lemma wp_eval_path E p v :
eval_path p = Some v WP p @ E {{ v', v' = v }}.
Proof.
revert v; induction p; intros v; try done.
{ intros [=]. by iApply wp_value. }
{ move=> /of_to_val=> ?. by iApply wp_value. }
simpl.
case_match; try discriminate; [].
case_match; try (intros ?; by iApply wp_value); [].
case_match; try (intros ?; by iApply wp_value); [].
case_match; try done.
case_match; try discriminate; [].
case_match; try discriminate; [].
intros [=<-].
match goal with |- context[(?l + _)%E] => rename l into p1 end.
wp_bind p1. iApply (wp_wand with "[]").
{ match goal with H: context[(WP p1 @ _ {{ _, _ }})%I] |- _ => iApply H end. done. }
iIntros (v) "%". subst v. wp_op. done.
Qed.
Lemma eval_path_closed p v :
eval_path p = Some v Closed [] p.
Proof.
intros Hpv. revert v Hpv.
induction p as [| | |[] p IH [|[]| | | | | | | | | |] _| | | | | | | |]=>//.
- unfold eval_path=>? /of_to_val <-. apply is_closed_of_val.
- simpl. destruct (eval_path p) as [[[]|]|]; intros ? [= <-].
specialize (IH _ eq_refl). apply _.
Qed.
(** Type context element *)
Definition tctx_elt_interp (tid : thread_id) (x : tctx_elt) : iProp Σ :=
match x with
| p ty => v, eval_path p = Some v ty.(ty_own) tid [v]
| p {κ} ty => v, eval_path p = Some v
([κ] ={}=∗ ty.(ty_own) tid [v])
end%I.
(* Block tctx_elt_interp from reducing with simpl when x is a constructor. *)
Global Arguments tctx_elt_interp : simpl never.
Lemma tctx_hasty_val tid (v : val) ty :
tctx_elt_interp tid (v ty) ⊣⊢ ty.(ty_own) tid [v].
Proof.
rewrite /tctx_elt_interp eval_path_of_val. iSplit.
- iIntros "H". iDestruct "H" as (?) "[EQ ?]".
iDestruct "EQ" as %[=->]. done.
- iIntros "?". iExists _. auto.
Qed.
Lemma tctx_elt_interp_hasty_path p1 p2 ty tid :
eval_path p1 = eval_path p2
tctx_elt_interp tid (p1 ty) tctx_elt_interp tid (p2 ty).
Proof. intros Hp. rewrite /tctx_elt_interp /=. setoid_rewrite Hp. done. Qed.
Lemma tctx_hasty_val' tid p (v : val) ty :
eval_path p = Some v
tctx_elt_interp tid (p ty) ⊣⊢ ty.(ty_own) tid [v].
Proof.
intros ?. rewrite -tctx_hasty_val. apply tctx_elt_interp_hasty_path.
rewrite eval_path_of_val. done.
Qed.
Lemma wp_hasty E tid p ty Φ :
tctx_elt_interp tid (p ty) -∗
( v, eval_path p = Some v -∗ ty.(ty_own) tid [v] -∗ Φ v) -∗
WP p @ E {{ Φ }}.
Proof.
iIntros "Hty HΦ". iDestruct "Hty" as (v) "[% Hown]".
iApply (wp_wand with "[]"). { iApply wp_eval_path. done. }
iIntros (v') "%". subst v'. iApply ("HΦ" with "[] Hown"). by auto.
Qed.
Lemma closed_hasty tid p ty : tctx_elt_interp tid (p ty) -∗ Closed [] p⌝.
Proof.
iIntros "H". iDestruct "H" as (?) "[% _]". eauto using eval_path_closed.
Qed.
(** Type context *)
Definition tctx_interp (tid : thread_id) (T : tctx) : iProp Σ :=
([ list] x T, tctx_elt_interp tid x)%I.
Global Instance tctx_interp_permut tid:
Proper (() ==> (⊣⊢)) (tctx_interp tid).
Proof. intros ???. by apply big_opL_permutation. Qed.
Lemma tctx_interp_cons tid x T :
tctx_interp tid (x :: T) ⊣⊢ (tctx_elt_interp tid x tctx_interp tid T).
Proof. done. Qed.
Lemma tctx_interp_app tid T1 T2 :
tctx_interp tid (T1 ++ T2) ⊣⊢ (tctx_interp tid T1 tctx_interp tid T2).
Proof. rewrite /tctx_interp big_sepL_app //. Qed.
Definition tctx_interp_nil tid :
tctx_interp tid [] = True%I := eq_refl _.
Lemma tctx_interp_singleton tid x :
tctx_interp tid [x] ⊣⊢ tctx_elt_interp tid x.
Proof. rewrite /tctx_interp /= right_id //. Qed.
(** Copy typing contexts *)
Class CopyC (T : tctx) :=
copyc_persistent tid : Persistent (tctx_interp tid T).
Global Existing Instances copyc_persistent.
Global Instance tctx_nil_copy : CopyC [].
Proof. rewrite /CopyC. apply _. Qed.
Global Instance tctx_ty_copy T p ty :
CopyC T Copy ty CopyC ((p ty) :: T).
Proof. intros ???. rewrite tctx_interp_cons. apply _. Qed.
(** Send typing contexts *)
Class SendC (T : tctx) :=
sendc_change_tid tid1 tid2 : tctx_interp tid1 T tctx_interp tid2 T.
Global Instance tctx_nil_send : SendC [].
Proof. done. Qed.
Global Instance tctx_ty_send T p ty :
SendC T Send ty SendC ((p ty) :: T).
Proof.
iIntros (HT Hty ??). rewrite !tctx_interp_cons.
iIntros "[Hty HT]". iSplitR "HT".
- iDestruct "Hty" as (?) "[% Hty]". iExists _. iSplit; first done.
by iApply Hty.
- by iApply HT.
Qed.
(** Type context inclusion *)
Definition tctx_incl (E : elctx) (L : llctx) (T1 T2 : tctx): Prop :=
tid qmax q2, lft_ctx -∗ elctx_interp E -∗ llctx_interp_noend qmax L q2 -∗
tctx_interp tid T1 ={}=∗ llctx_interp_noend qmax L q2 tctx_interp tid T2.
Global Instance : E L, RewriteRelation (tctx_incl E L) := {}.
Global Instance tctx_incl_preorder E L : PreOrder (tctx_incl E L).
Proof.
split.
- by iIntros (????) "_ _ $ $".
- iIntros (??? H1 H2 ???) "#LFT #HE HL H".
iMod (H1 with "LFT HE HL H") as "(HL & H)".
by iMod (H2 with "LFT HE HL H") as "($ & $)".
Qed.
Lemma contains_tctx_incl E L T1 T2 : T1 ⊆+ T2 tctx_incl E L T2 T1.
Proof.
rewrite /tctx_incl. iIntros (Hc ???) "_ _ $ H". by iApply big_sepL_submseteq.
Qed.
Global Instance tctx_incl_app E L :
Proper (tctx_incl E L ==> tctx_incl E L ==> tctx_incl E L) app.
Proof.
intros ?? Hincl1 ?? Hincl2 ???. rewrite /tctx_interp !big_sepL_app.
iIntros "#LFT #HE HL [H1 H2]".
iMod (Hincl1 with "LFT HE HL H1") as "(HL & $)".
iApply (Hincl2 with "LFT HE HL H2").
Qed.
Global Instance tctx_incl_cons E L x :
Proper (tctx_incl E L ==> tctx_incl E L) (cons x).
Proof. by apply (tctx_incl_app E L [_] [_]). Qed.
Lemma tctx_incl_frame_l {E L} T T1 T2 :
tctx_incl E L T1 T2 tctx_incl E L (T++T1) (T++T2).
Proof. by apply tctx_incl_app. Qed.
Lemma tctx_incl_frame_r {E L} T T1 T2 :
tctx_incl E L T1 T2 tctx_incl E L (T1++T) (T2++T).
Proof. by intros; apply tctx_incl_app. Qed.
Lemma copy_tctx_incl E L p `{!Copy ty} :
tctx_incl E L [p ty] [p ty; p ty].
Proof. iIntros (???) "_ _ $ * [#$ $] //". Qed.
Lemma copy_elem_of_tctx_incl E L T p `{!Copy ty} :
p ty T tctx_incl E L T ((p ty) :: T).
Proof.
remember (p ty). induction 1 as [|???? IH]; subst.
- apply (tctx_incl_frame_r _ [_] [_;_]), copy_tctx_incl, _.
- etrans; first by apply (tctx_incl_frame_l [_]), IH, reflexivity.
apply contains_tctx_incl, submseteq_swap.
Qed.
Lemma type_incl_tctx_incl tid p ty1 ty2 :
type_incl ty1 ty2 -∗ tctx_interp tid [p ty1] -∗ tctx_interp tid [p ty2].
Proof.
iIntros "Hincl HT". rewrite !tctx_interp_singleton.
iDestruct "HT" as (v) "[% HT]". iExists _. iFrame "%".
iDestruct "Hincl" as "(_ & Hincl & _)". iApply "Hincl". done.
Qed.
Lemma subtype_tctx_incl E L p ty1 ty2 :
subtype E L ty1 ty2 tctx_incl E L [p ty1] [p ty2].
Proof.
iIntros (Hst ???) "#LFT #HE HL HT".
iDestruct (Hst with "HL") as "#Hst". iFrame "HL".
iApply type_incl_tctx_incl; last done.
by iApply "Hst".
Qed.
(* Extracting from a type context. *)
Definition tctx_extract_hasty E L p ty T T' : Prop :=
tctx_incl E L T ((p ty)::T').
Lemma tctx_extract_hasty_further E L p ty T T' x :
tctx_extract_hasty E L p ty T T'
tctx_extract_hasty E L p ty (x::T) (x::T').
Proof. unfold tctx_extract_hasty=>->. apply contains_tctx_incl, submseteq_swap. Qed.
Lemma tctx_extract_hasty_here_copy E L p1 p2 ty ty' T :
p1 = p2 Copy ty subtype E L ty ty'
tctx_extract_hasty E L p1 ty' ((p2 ty)::T) ((p2 ty)::T).
Proof.
intros -> ??. apply (tctx_incl_frame_r _ [_] [_;_]).
etrans; first by apply copy_tctx_incl, _.
by apply (tctx_incl_frame_r _ [_] [_]), subtype_tctx_incl.
Qed.
Lemma tctx_extract_hasty_here E L p1 p2 ty ty' T :
p1 = p2 subtype E L ty ty' tctx_extract_hasty E L p1 ty' ((p2 ty)::T) T.
Proof.
intros -> ?. by apply (tctx_incl_frame_r _ [_] [_]), subtype_tctx_incl.
Qed.
Lemma tctx_extract_hasty_here_eq E L p1 p2 ty T :
p1 = p2 tctx_extract_hasty E L p1 ty ((p2 ty)::T) T.
Proof. intros ->. by apply tctx_extract_hasty_here. Qed.
Definition tctx_extract_blocked E L p κ ty T T' : Prop :=
tctx_incl E L T ((p {κ} ty)::T').
Lemma tctx_extract_blocked_cons E L p κ ty T T' x :
tctx_extract_blocked E L p κ ty T T'
tctx_extract_blocked E L p κ ty (x::T) (x::T').
Proof.
move=> /(tctx_incl_frame_l [x]) /= Hincl. rewrite /tctx_extract_blocked.
etrans; first done. apply contains_tctx_incl, submseteq_swap.
Qed.
Lemma tctx_extract_blocked_here E L p κ ty T :
tctx_extract_blocked E L p κ ty ((p {κ} ty)::T) T.
Proof. intros. by apply (tctx_incl_frame_r _ [_] [_]). Qed.
Definition tctx_extract_ctx E L T T1 T2 : Prop :=
tctx_incl E L T1 (T++T2).
Lemma tctx_extract_ctx_nil E L T:
tctx_extract_ctx E L [] T T.
Proof. by unfold tctx_extract_ctx. Qed.
Lemma tctx_extract_ctx_hasty E L T T1 T2 T3 p ty:
tctx_extract_hasty E L p ty T1 T2 tctx_extract_ctx E L T T2 T3
tctx_extract_ctx E L ((p ty)::T) T1 T3.
Proof. unfold tctx_extract_ctx, tctx_extract_hasty=>->->//. Qed.
Lemma tctx_extract_ctx_blocked E L T T1 T2 T3 p κ ty:
tctx_extract_blocked E L p κ ty T1 T2 tctx_extract_ctx E L T T2 T3
tctx_extract_ctx E L ((p {κ} ty)::T) T1 T3.
Proof. unfold tctx_extract_ctx, tctx_extract_blocked=>->->//. Qed.
Lemma tctx_extract_ctx_incl E L T T' T0:
tctx_extract_ctx E L T' T T0
tctx_incl E L T T'.
Proof.
unfold tctx_extract_ctx=>->.
by apply contains_tctx_incl, submseteq_inserts_r.
Qed.
(* Unblocking a type context. *)
(* TODO : That would be great if this could also remove all the
instances mentionning the lifetime in question.
E.g., if [p ◁ &uniq{κ} ty] should be removed, because this is now
useless. *)
Class UnblockTctx (κ : lft) (T1 T2 : tctx) : Prop :=
unblock_tctx : tid, [κ] -∗ tctx_interp tid T1 ={}=∗ tctx_interp tid T2.
Global Instance unblock_tctx_nil κ : UnblockTctx κ [] [].
Proof. by iIntros (tid) "_ $". Qed.
Global Instance unblock_tctx_cons_unblock T1 T2 p κ ty :
UnblockTctx κ T1 T2
UnblockTctx κ ((p {κ} ty)::T1) ((p ty)::T2).
Proof.
iIntros (H12 tid) "#H†". rewrite !tctx_interp_cons. iIntros "[H HT1]".
iMod (H12 with "H† HT1") as "$". iDestruct "H" as (v) "[% H]".
iExists (v). by iMod ("H" with "H†") as "$".
Qed.
Global Instance unblock_tctx_cons κ T1 T2 x :
UnblockTctx κ T1 T2 UnblockTctx κ (x::T1) (x::T2) | 100.
Proof.
iIntros (H12 tid) "#H†". rewrite !tctx_interp_cons. iIntros "[$ HT1]".
by iMod (H12 with "H† HT1") as "$".
Qed.
End type_context.
Global Hint Resolve tctx_extract_hasty_here_copy | 1 : lrust_typing.
Global Hint Resolve tctx_extract_hasty_here | 20 : lrust_typing.
Global Hint Resolve tctx_extract_hasty_further | 50 : lrust_typing.
Global Hint Resolve tctx_extract_blocked_here tctx_extract_blocked_cons
tctx_extract_ctx_nil tctx_extract_ctx_hasty
tctx_extract_ctx_blocked tctx_extract_ctx_incl : lrust_typing.
Global Hint Opaque tctx_extract_ctx tctx_extract_hasty tctx_extract_blocked
tctx_incl : lrust_typing.
(* In general, we want reborrowing to be tried before subtyping, so
that we get the extraction. However, in the case the types match
exactly, we want to NOT use reborrowing. Therefore, we add
[tctx_extract_hasty_here_eq] as a hint with a very low cost. *)
Global Hint Resolve tctx_extract_hasty_here_eq | 2 : lrust_typing.
From iris.proofmode Require Import proofmode.
From lrust.lang Require Import memcpy.
From lrust.typing Require Import uninit uniq_bor shr_bor own sum.
From lrust.typing Require Import lft_contexts type_context programs product.
From iris.prelude Require Import options.
Section case.
Context `{!typeGS Σ}.
(* FIXME : have an Iris version of Forall2. *)
Lemma type_case_own' E L C T p n tyl el :
Forall2 (λ ty e,
( typed_body E L C ((p + #0 own_ptr n (uninit 1)) :: (p + #1 own_ptr n ty) ::
(p + #(S (ty.(ty_size)))
own_ptr n (uninit (max_list_with ty_size tyl - ty_size ty))) :: T) e)
( typed_body E L C ((p own_ptr n (sum tyl)) :: T) e))
tyl el
typed_body E L C ((p own_ptr n (sum tyl)) :: T) (case: !p of el).
Proof.
iIntros (Hel tid qmax) "#LFT #HE Hna HL HC HT". wp_bind p.
rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]".
iApply (wp_hasty with "Hp"). iIntros ([[]|] Hv) "Hp"; try done.
iDestruct "Hp" as "[H↦ Hf]". iDestruct "H↦" as (vl) "[H↦ Hown]".
iDestruct "Hown" as (i vl' vl'') "(>% & >EQlen & Hown)". subst.
simpl ty_size. iDestruct "EQlen" as %[=EQlen]. rewrite -EQlen. simpl length.
rewrite -Nat.add_1_l length_app -!freeable_sz_split
heap_pointsto_vec_cons heap_pointsto_vec_app.
iDestruct "H↦" as "(H↦i & H↦vl' & H↦vl'')".
iDestruct "Hf" as "(Hfi & Hfvl' & Hfvl'')".
rewrite nth_lookup.
destruct (tyl !! i) as [ty|] eqn:EQty; last iDestruct "Hown" as ">[]".
edestruct @Forall2_lookup_l as (e & He & Hety); eauto.
wp_read. wp_case; first (split; [lia|by rewrite Nat2Z.id]).
destruct Hety as [Hety|Hety]; iApply (Hety with "LFT HE Hna HL HC");
rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //=; iFrame "HT".
- rewrite /own_ptr /=. iDestruct (ty.(ty_size_eq) with "Hown") as %<-.
iSplitL "H↦i Hfi"; last iSplitR "H↦vl'' Hfvl''".
+ rewrite shift_loc_0. iFrame. iExists [ #i]. rewrite heap_pointsto_vec_singleton.
iFrame. auto.
+ eauto with iFrame.
+ rewrite -EQlen length_app Nat.add_comm Nat.add_sub shift_loc_assoc_nat.
by iFrame.
- rewrite /= -EQlen length_app -(Nat.add_1_l (_+_)) -!freeable_sz_split. iFrame.
iExists (#i :: vl' ++ vl''). iNext. rewrite heap_pointsto_vec_cons heap_pointsto_vec_app.
iFrame. iExists i, vl', vl''. rewrite /= length_app nth_lookup EQty /=. auto.
Qed.
Lemma type_case_own E L C T T' p n tyl el :
tctx_extract_hasty E L p (own_ptr n (sum tyl)) T T'
Forall2 (λ ty e,
( typed_body E L C ((p + #0 own_ptr n (uninit 1)) :: (p + #1 own_ptr n ty) ::
(p + #(S (ty.(ty_size)))
own_ptr n (uninit (max_list_with ty_size tyl - ty_size ty))) :: T') e)
( typed_body E L C ((p own_ptr n (sum tyl)) :: T') e))
tyl el
typed_body E L C T (case: !p of el).
Proof. unfold tctx_extract_hasty=>->. apply type_case_own'. Qed.
Lemma type_case_uniq' E L C T p κ tyl el :
lctx_lft_alive E L κ
Forall2 (λ ty e,
( typed_body E L C ((p + #1 &uniq{κ}ty) :: T) e)
( typed_body E L C ((p &uniq{κ}(sum tyl)) :: T) e)) tyl el
typed_body E L C ((p &uniq{κ}(sum tyl)) :: T) (case: !p of el).
Proof.
iIntros (Halive Hel tid qmax) "#LFT #HE Hna HL HC HT". wp_bind p.
rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]".
iApply (wp_hasty with "Hp"). iIntros ([[|l|]|] Hv) "Hp"; try iDestruct "Hp" as "[]".
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (Halive with "HE HL") as (q) "[Htok Hclose]"; first done.
iMod (bor_acc_cons with "LFT Hp Htok") as "[H↦ Hclose']"; first done.
iDestruct "H↦" as (vl) "[H↦ Hown]".
iDestruct "Hown" as (i vl' vl'') "(>% & >EQlen & Hown)". subst.
iDestruct "EQlen" as %[=EQlen].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_app nth_lookup.
iDestruct "H↦" as "(H↦i & H↦vl' & H↦vl'')".
destruct (tyl !! i) as [ty|] eqn:EQty; last iDestruct "Hown" as ">[]".
edestruct @Forall2_lookup_l as (e & He & Hety); eauto.
wp_read. wp_case; first (split; [lia|by rewrite Nat2Z.id]).
iDestruct (ty.(ty_size_eq) with "Hown") as %EQlenvl'.
destruct Hety as [Hety|Hety].
- iMod ("Hclose'" $! ((l + 1) ↦∗: ty.(ty_own) tid)%I
with "[H↦i H↦vl''] [H↦vl' Hown]") as "[Hb Htok]".
{ iIntros "!>Hown". iDestruct "Hown" as (vl'2) "[H↦ Hown]".
iExists (#i::vl'2++vl''). iIntros "!>". iNext.
iDestruct (ty.(ty_size_eq) with "Hown") as %EQlenvl'2.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_app EQlenvl' EQlenvl'2.
iFrame. iExists _, _, _. iSplit; first by auto.
rewrite /= -EQlen !length_app EQlenvl' EQlenvl'2 nth_lookup EQty /=. auto. }
{ iExists vl'. iFrame. }
iMod ("Hclose" with "Htok") as "HL".
iDestruct ("HLclose" with "HL") as "HL".
iApply (Hety with "LFT HE Hna HL HC").
rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //. iFrame.
- iMod ("Hclose'" with "[] [H↦i H↦vl' H↦vl'' Hown]") as "[Hb Htok]";
[by iIntros "!>$"| |].
{ iExists (#i::vl'++vl'').
rewrite heap_pointsto_vec_cons heap_pointsto_vec_app /= -EQlen. iFrame. iNext.
iExists i, vl', vl''. rewrite nth_lookup EQty. auto. }
iMod ("Hclose" with "Htok") as "HL".
iDestruct ("HLclose" with "HL") as "HL".
iApply (Hety with "LFT HE Hna HL HC").
rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //. iFrame.
Qed.
Lemma type_case_uniq E L C T T' p κ tyl el :
tctx_extract_hasty E L p (&uniq{κ}(sum tyl)) T T'
lctx_lft_alive E L κ
Forall2 (λ ty e,
( typed_body E L C ((p + #1 &uniq{κ}ty) :: T') e)
( typed_body E L C ((p &uniq{κ}(sum tyl)) :: T') e)) tyl el
typed_body E L C T (case: !p of el).
Proof. unfold tctx_extract_hasty=>->. apply type_case_uniq'. Qed.
Lemma type_case_shr' E L C T p κ tyl el:
lctx_lft_alive E L κ
Forall2 (λ ty e,
( typed_body E L C ((p + #1 &shr{κ}ty) :: T) e)
( typed_body E L C ((p &shr{κ}(sum tyl)) :: T) e)) tyl el
typed_body E L C ((p &shr{κ}(sum tyl)) :: T) (case: !p of el).
Proof.
iIntros (Halive Hel tid qmax) "#LFT #HE Hna HL HC HT". wp_bind p.
rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]".
iApply (wp_hasty with "Hp"). iIntros ([[]|] Hv) "Hp"; try done.
iDestruct "Hp" as (i) "[#Hb Hshr]".
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (Halive with "HE HL") as (q) "[Htok Hclose]"; first done.
iMod (frac_bor_acc with "LFT Hb Htok") as (q') "[[H↦i H↦vl''] Hclose']"; first done.
rewrite nth_lookup.
destruct (tyl !! i) as [ty|] eqn:EQty; last done.
edestruct @Forall2_lookup_l as (e & He & Hety); eauto.
wp_read. wp_case; first (split; [lia|by rewrite Nat2Z.id]).
iMod ("Hclose'" with "[$H↦i $H↦vl'']") as "Htok".
iMod ("Hclose" with "Htok") as "HL".
iDestruct ("HLclose" with "HL") as "HL".
destruct Hety as [Hety|Hety]; iApply (Hety with "LFT HE Hna HL HC");
rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //; iFrame.
iExists _. rewrite ->nth_lookup, EQty. auto.
Qed.
Lemma type_case_shr E L C T p κ tyl el :
p &shr{κ}(sum tyl) T
lctx_lft_alive E L κ
Forall2 (λ ty e, typed_body E L C ((p + #1 &shr{κ}ty) :: T) e) tyl el
typed_body E L C T (case: !p of el).
Proof.
intros. rewrite ->copy_elem_of_tctx_incl; last done; last apply _.
apply type_case_shr'; first done. eapply Forall2_impl; first done. auto.
Qed.
Lemma type_sum_assign_instr {E L} (i : nat) ty1 tyl ty ty2 p1 p2 :
tyl !! i = Some ty
( typed_write E L ty1 (sum tyl) ty2)
typed_instruction E L [p1 ty1; p2 ty] (p1 <-{Σ i} p2) (λ _, [p1 ty2]).
Proof.
iIntros (Hty Hw tid qmax) "#LFT #HE $ HL Hp".
rewrite tctx_interp_cons tctx_interp_singleton.
iDestruct "Hp" as "[Hp1 Hp2]". iDestruct (closed_hasty with "Hp1") as "%".
iDestruct (closed_hasty with "Hp2") as "%". wp_bind p1.
iApply (wp_hasty with "Hp1"). iIntros (v1 Hv1) "Hty1".
rewrite typed_write_eq in Hw.
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (Hw with "[] LFT HE HL Hty1") as (l vl) "(H & H↦ & Hw)"=>//=.
destruct vl as [|? vl]; iDestruct "H" as %[[= Hlen] ->].
rewrite heap_pointsto_vec_cons. iDestruct "H↦" as "[H↦0 H↦vl]".
wp_write. wp_bind p1. iApply (wp_wand with "[]"); first by iApply (wp_eval_path).
iIntros (? ->). wp_op. wp_bind p2. iApply (wp_hasty with "Hp2").
iIntros (v2 Hv2) "Hty2". iDestruct (ty_size_eq with "Hty2") as %Hlenty.
destruct vl as [|? vl].
{ exfalso. revert i Hty. clear - Hlen Hlenty.
induction tyl as [|ty' tyl IH]=>//= -[|i] /=.
- intros [= ->]. simpl in *. lia.
- apply IH. simpl in *. lia. }
rewrite heap_pointsto_vec_cons -wp_fupd. iDestruct "H↦vl" as "[H↦ H↦vl]". wp_write.
rewrite tctx_interp_singleton tctx_hasty_val' //.
iMod ("Hw" with "[-HLclose]") as "[HL $]"; last first.
{ iApply "HLclose". done. }
iNext.
iExists (_::_::_). rewrite !heap_pointsto_vec_cons. iFrame.
iExists i, [_], _. rewrite -Hlen nth_lookup Hty. auto.
Qed.
Lemma type_sum_assign {E L} sty tyl i ty1 ty ty1' C T T' p1 p2 e:
Closed [] e
0 i
sty = sum tyl
tctx_extract_ctx E L [p1 ty1; p2 ty] T T'
tyl !! (Z.to_nat i) = Some ty
( typed_write E L ty1 sty ty1')
typed_body E L C ((p1 ty1') :: T') e -∗
typed_body E L C T (p1 <-{Σ i} p2 ;; e).
Proof.
iIntros (??->) "* **". rewrite -(Z2Nat.id i) //.
iApply type_seq; [by eapply type_sum_assign_instr|done|done].
Qed.
Lemma type_sum_unit_instr {E L} (i : nat) tyl ty1 ty2 p :
tyl !! i = Some unit
( typed_write E L ty1 (sum tyl) ty2)
typed_instruction E L [p ty1] (p <-{Σ i} ()) (λ _, [p ty2]).
Proof.
iIntros (Hty Hw tid qmax) "#LFT #HE $ HL Hp". rewrite tctx_interp_singleton.
wp_bind p. iApply (wp_hasty with "Hp"). iIntros (v Hv) "Hty".
rewrite typed_write_eq in Hw.
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (Hw with "[] LFT HE HL Hty") as (l vl) "(H & H↦ & Hw)"; first done.
simpl. destruct vl as [|? vl]; iDestruct "H" as %[[= Hlen] ->].
rewrite heap_pointsto_vec_cons -wp_fupd. iDestruct "H↦" as "[H↦0 H↦vl]".
wp_write. rewrite tctx_interp_singleton tctx_hasty_val' //.
iMod ("Hw" with "[-HLclose]") as "[HL $]"; last first.
{ iApply "HLclose". done. }
iModIntro. iExists (_::_). rewrite heap_pointsto_vec_cons. iFrame.
iExists i, [], _. rewrite -Hlen nth_lookup Hty. auto.
Qed.
Lemma type_sum_unit {E L} sty tyl i ty1 ty1' C T T' p e:
Closed [] e
0 i
sty = sum tyl
tctx_extract_hasty E L p ty1 T T'
tyl !! (Z.to_nat i) = Some unit
( typed_write E L ty1 sty ty1')
typed_body E L C ((p ty1') :: T') e -∗
typed_body E L C T (p <-{Σ i} () ;; e).
Proof.
iIntros (??->) "* **". rewrite -(Z2Nat.id i) //.
iApply type_seq; [by eapply type_sum_unit_instr|solve_typing|done].
Qed.
Lemma type_sum_memcpy_instr {E L} (i : nat) tyl ty1 ty1' ty2 ty2' ty p1 p2 :
tyl !! i = Some ty
( typed_write E L ty1 (sum tyl) ty1')
( typed_read E L ty2 ty ty2')
typed_instruction E L [p1 ty1; p2 ty2]
(p1 <-{ty.(ty_size),Σ i} !p2) (λ _, [p1 ty1'; p2 ty2']).
Proof.
iIntros (Hty Hw Hr tid qmax) "#LFT #HE Htl HL Hp".
iDestruct (llctx_interp_acc_noend with "HL") as "[[HL1 HL2] HLclose]".
rewrite tctx_interp_cons tctx_interp_singleton.
iDestruct "Hp" as "[Hp1 Hp2]". iDestruct (closed_hasty with "Hp1") as "%".
iDestruct (closed_hasty with "Hp2") as "%". wp_bind p1.
iApply (wp_hasty with "Hp1"). iIntros (v1 Hv1) "Hty1".
rewrite typed_write_eq in Hw.
iMod (Hw with "[] LFT HE HL1 Hty1") as (l1 vl1) "(H & H↦ & Hw)"=>//=.
clear Hw. destruct vl1 as [|? vl1]; iDestruct "H" as %[[= Hlen] ->].
rewrite heap_pointsto_vec_cons -wp_fupd. iDestruct "H↦" as "[H↦0 H↦vl1]". wp_write.
wp_bind p1. iApply (wp_wand with "[]"); first by iApply (wp_eval_path). iIntros (? ->).
wp_op. wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2 Hv2) "Hty2".
rewrite typed_read_eq in Hr.
iMod (Hr with "[] LFT HE Htl HL2 Hty2") as (l2 vl2 q) "(% & H↦2 & Hty & Hr)"=>//=.
clear Hr. subst. assert (ty.(ty_size) length vl1).
{ revert i Hty. rewrite Hlen. clear.
induction tyl as [|ty' tyl IH]=>//= -[|i] /=.
- intros [= ->]. lia.
- specialize (IH i). intuition lia. }
rewrite -(take_drop (ty.(ty_size)) vl1) heap_pointsto_vec_app.
iDestruct "H↦vl1" as "[H↦vl1 H↦pad]".
iDestruct (ty_size_eq with "Hty") as "#>%".
iApply (wp_memcpy with "[$H↦vl1 $H↦2]"); [|lia|].
{ rewrite length_take. lia. }
iNext; iIntros "[H↦vl1 H↦2]".
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //.
iMod ("Hr" with "H↦2") as "($ & HL1 & $)".
iMod ("Hw" with "[-HLclose HL1]") as "[HL $]"; last first.
{ iApply "HLclose". by iFrame. }
iNext.
rewrite split_sum_mt /is_pad. iExists i. rewrite nth_lookup Hty. iFrame.
rewrite (shift_loc_assoc_nat _ 1) length_take Nat.min_l; last lia. iFrame.
rewrite /= length_drop. iPureIntro. lia.
Qed.
Lemma type_sum_memcpy {E L} sty tyl i ty1 ty2 ty n ty1' ty2' C T T' p1 p2 e:
Closed [] e
0 i
sty = sum tyl
tctx_extract_ctx E L [p1 ty1; p2 ty2] T T'
tyl !! (Z.to_nat i) = Some ty
( typed_write E L ty1 sty ty1')
( typed_read E L ty2 ty ty2')
Z.of_nat (ty.(ty_size)) = n
typed_body E L C ((p1 ty1') :: (p2 ty2') :: T') e -∗
typed_body E L C T (p1 <-{n,Σ i} !p2 ;; e).
Proof.
iIntros (?? -> ??? Hr ?) "?". subst. rewrite -(Z2Nat.id i) //.
by iApply type_seq; [eapply type_sum_memcpy_instr, Hr|done|done].
Qed.
Lemma ty_outlives_E_elctx_sat_sum E L tyl {Wf : ListTyWf tyl} α:
elctx_sat E L (tyl_outlives_E tyl α)
elctx_sat E L (ty_outlives_E (sum tyl) α).
Proof.
intro Hsat. eapply eq_ind; first done. clear Hsat. rewrite /ty_outlives_E /=.
induction Wf as [|ty [] ?? IH]=>//=. rewrite IH fmap_app //.
Qed.
End case.
Global Hint Resolve ty_outlives_E_elctx_sat_sum : lrust_typing.
From lrust.lang Require Export tactics memcpy notation.
From lrust.typing Require Export
lft_contexts type_context cont_context programs cont type
int bool own uniq_bor shr_bor uninit product sum fixpoint function
product_split borrow type_sum.
From iris.prelude Require Import options.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Export type.
From lrust.typing Require Import product.
From iris.prelude Require Import options.
Section uninit.
Context `{!typeGS Σ}.
Program Definition uninit_1 : type :=
{| st_own tid vl := length vl = 1%nat⌝%I |}.
Next Obligation. done. Qed.
Global Instance uninit_1_send : Send uninit_1.
Proof. iIntros (tid1 tid2 vl) "H". done. Qed.
Definition uninit0 (n : nat) : type :=
Π (replicate n uninit_1).
Lemma uninit0_sz n : ty_size (uninit0 n) = n.
Proof. induction n=>//=. by f_equal. Qed.
Lemma uninit0_own n tid vl :
ty_own (uninit0 n) tid vl ⊣⊢ length vl = n⌝.
Proof.
iSplit.
- iIntros "Hvl".
iInduction n as [|n] "IH" forall (vl); simpl.
+ iDestruct "Hvl" as "%". subst vl. done.
+ iDestruct "Hvl" as (vl1 vl2) "(% & % & Hprod)".
destruct vl1 as [|v [|]]; try done. subst vl. simpl.
iDestruct ("IH" with "Hprod") as "%". iPureIntro. by f_equal.
- iIntros "%". subst n. iInduction vl as [|v vl] "IH"; first done.
simpl. iExists [v], vl. auto.
Qed.
(* We redefine uninit as an alias of uninit0, so that ty_size and ty_own
compute directly. We re-use the sharing from the product as that saves a whole
lot of work. *)
Program Definition uninit (n : nat) : type :=
{| ty_size := n; ty_own tid vl := length vl = n⌝%I;
ty_shr := (uninit0 n).(ty_shr) |}.
Next Obligation. iIntros (???) "%". done. Qed.
Next Obligation.
iIntros (n ??????) "LFT Hvl". iApply (ty_share (uninit0 n) with "LFT"); first done.
iApply (bor_iff with "[] Hvl"). iIntros "!> !>". setoid_rewrite uninit0_own.
iSplit; iIntros; done.
Qed.
Next Obligation. intros. by apply ty_shr_mono. Qed.
Global Instance uninit_wf n : TyWf (uninit n) :=
{ ty_lfts := []; ty_wf_E := [] }.
Global Instance uninit_copy n : Copy (uninit n).
Proof.
assert (Copy (uninit0 n)) as [A B].
{ apply product_copy. by apply Forall_replicate, _. }
split; first by apply _.
rewrite uninit0_sz in B. setoid_rewrite uninit0_own in B. done.
Qed.
Global Instance uninit_send n : Send (uninit n).
Proof. iIntros (???) "?". done. Qed.
Global Instance uninit_sync n : Sync (uninit n).
Proof. apply product_sync, Forall_replicate, _. Qed.
(* Unfolding lemma, kep only for backwards compatibility. *)
Lemma uninit_own n tid vl :
(uninit n).(ty_own) tid vl ⊣⊢ length vl = n⌝.
Proof. done. Qed.
Lemma uninit_uninit0_eqtype E L n :
eqtype E L (uninit0 n) (uninit n).
Proof.
apply equiv_eqtype; constructor=>//=.
- apply uninit0_sz.
- iIntros (??). rewrite uninit0_own. done.
Qed.
Lemma uninit_product_subtype_cons_r {E L} (n : nat) ty tyl :
ty.(ty_size) n
subtype E L (uninit ty.(ty_size)) ty
subtype E L (uninit (n-ty.(ty_size))) (Π tyl)
subtype E L (uninit n) (Π(ty :: tyl)).
Proof.
intros ?%Nat2Z.inj_le. rewrite -!uninit_uninit0_eqtype /uninit0=>Hty Htyl.
rewrite -(Nat.sub_add ty.(ty_size) n) // Nat.add_comm // replicate_add
-(prod_flatten_r _ _ [_]) /= -prod_app. repeat (done || f_equiv).
Qed.
Lemma uninit_product_subtype_cons_l {E L} (n : nat) ty tyl :
ty.(ty_size) n
subtype E L ty (uninit ty.(ty_size))
subtype E L (Π tyl) (uninit (n-ty.(ty_size)))
subtype E L (Π(ty :: tyl)) (uninit n).
Proof.
intros ?%Nat2Z.inj_le. rewrite -!uninit_uninit0_eqtype /uninit0=>Hty Htyl.
rewrite -(Nat.sub_add ty.(ty_size) n) // Nat.add_comm replicate_add
-(prod_flatten_r _ _ [_]) /= -prod_app. repeat (done || f_equiv).
Qed.
Lemma uninit_product_eqtype_cons_r {E L} (n : nat) ty tyl :
ty.(ty_size) n
eqtype E L (uninit ty.(ty_size)) ty
eqtype E L (uninit (n-ty.(ty_size))) (Π tyl)
eqtype E L (uninit n) (Π(ty :: tyl)).
Proof.
intros ? [] []. split.
- by apply uninit_product_subtype_cons_r.
- by apply uninit_product_subtype_cons_l.
Qed.
Lemma uninit_product_eqtype_cons_l {E L} (n : nat) ty tyl :
ty.(ty_size) n
eqtype E L ty (uninit ty.(ty_size))
eqtype E L (Π tyl) (uninit (n-ty.(ty_size)))
eqtype E L (Π(ty :: tyl)) (uninit n).
Proof. symmetry. by apply uninit_product_eqtype_cons_r. Qed.
Lemma uninit_unit_eqtype E L n :
n = 0%nat
eqtype E L (uninit n) unit.
Proof. rewrite -uninit_uninit0_eqtype=>->//. Qed.
Lemma uninit_unit_eqtype' E L n :
n = 0%nat
eqtype E L unit (uninit n).
Proof. rewrite -uninit_uninit0_eqtype=>->//. Qed.
Lemma uninit_unit_subtype E L n :
n = 0%nat
subtype E L (uninit n) unit.
Proof. rewrite -uninit_uninit0_eqtype=>->//. Qed.
Lemma uninit_unit_subtype' E L n :
n = 0%nat
subtype E L unit (uninit n).
Proof. rewrite -uninit_uninit0_eqtype=>->//. Qed.
End uninit.
Global Hint Resolve uninit_product_eqtype_cons_l uninit_product_eqtype_cons_r
uninit_product_subtype_cons_l uninit_product_subtype_cons_r
uninit_unit_eqtype uninit_unit_eqtype'
uninit_unit_subtype uninit_unit_subtype' : lrust_typing.
From iris.proofmode Require Import proofmode.
From lrust.lang Require Import heap.
From lrust.typing Require Export type.
From lrust.typing Require Import util lft_contexts type_context programs.
From iris.prelude Require Import options.
Section uniq_bor.
Context `{!typeGS Σ}.
Program Definition uniq_bor (κ:lft) (ty:type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] => &{κ} (l ↦∗: ty.(ty_own) tid)
| _ => False
end;
ty_shr κ' tid l :=
l':loc, &frac{κ'}(λ q', l {q'} #l')
F q, ⌜↑shrN lftN F -∗ q.[κκ']
={F}[F∖↑shrN]▷=∗ ty.(ty_shr) (κκ') tid l' q.[κκ']
|}%I.
Next Obligation. by iIntros (q ty tid [|[[]|][]]) "H". Qed.
Next Obligation.
move=> κ ty N κ' l tid ??/=. iIntros "#LFT Hshr Htok".
iMod (bor_exists with "LFT Hshr") as ([|[[|l'|]|][]]) "Hb"; first solve_ndisj;
(iMod (bor_sep with "LFT Hb") as "[Hb1 Hb2]"; first solve_ndisj);
try (iMod (bor_persistent with "LFT Hb2 Htok") as "[>[] _]"; solve_ndisj).
iFrame. iExists l'. subst. rewrite heap_pointsto_vec_singleton.
iMod (bor_fracture (λ q, l {q} #l')%I with "LFT Hb1") as "$"; first solve_ndisj.
iApply delay_sharing_nested; try done. iApply lft_incl_refl.
Qed.
Next Obligation.
intros κ0 ty κ κ' tid l. iIntros "#Hκ #H".
iDestruct "H" as (l') "[Hfb Hvs]". iAssert (κ0κ' κ0κ)%I as "#Hκ0".
{ iApply lft_intersect_mono; last done. iApply lft_incl_refl. }
iExists l'. iSplit; first by iApply (frac_bor_shorten with "[]").
iIntros "!> %F %q % Htok". iApply (step_fupd_mask_mono F _ (F∖↑shrN)); try solve_ndisj.
iMod (lft_incl_acc with "Hκ0 Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hvs" with "[%] Htok") as "Hvs'"; first solve_ndisj. iModIntro. iNext.
iMod "Hvs'" as "[#Hshr Htok]". iMod ("Hclose" with "Htok") as "$".
by iApply (ty_shr_mono with "Hκ0").
Qed.
Global Instance uniq_bor_wf κ ty `{!TyWf ty} : TyWf (uniq_bor κ ty) :=
{ ty_lfts := [κ]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty κ }.
Lemma uniq_type_incl κ1 κ2 ty1 ty2 :
κ2 κ1 -∗
type_equal ty1 ty2 -∗
type_incl (uniq_bor κ1 ty1) (uniq_bor κ2 ty2).
Proof.
iIntros "#Hlft #Hty". iSplit; first done.
iSplit; iModIntro.
- iIntros (? [|[[]|][]]) "H"; try done.
iApply (bor_shorten with "Hlft"). iApply bor_iff; last done.
iNext. iModIntro.
iDestruct "Hty" as "(_ & Hty & _)".
iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]";
iExists vl; iFrame; by iApply "Hty".
- iIntros (κ ??) "H". iAssert (κ2 κ κ1 κ)%I as "#Hincl'".
{ iApply lft_intersect_mono; first done. iApply lft_incl_refl. }
iDestruct "H" as (l') "[Hbor #Hupd]". iExists l'. iIntros "{$Hbor}!> %%% Htok".
iMod (lft_incl_acc with "Hincl' Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hupd" with "[%] Htok") as "Hupd'"; try done. iModIntro. iNext.
iMod "Hupd'" as "[H Htok]". iMod ("Hclose" with "Htok") as "$".
iDestruct "Hty" as "(_ & _ & Hty)".
iApply ty_shr_mono; last by iApply "Hty".
done.
Qed.
Global Instance uniq_mono E L :
Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) uniq_bor.
Proof.
intros κ1 κ2 ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (??) "HL".
iDestruct (Hty with "HL") as "#Hty". iDestruct ( with "HL") as "#Hκ".
iIntros "!> #HE".
iApply uniq_type_incl.
- iDestruct ("Hκ" with "HE") as %H.
apply lft_incl_syn_sem in H. iApply H.
- iNext. iApply "Hty". done.
Qed.
Global Instance uniq_mono_flip E L :
Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) uniq_bor.
Proof. intros ??????. apply uniq_mono; first done. by symmetry. Qed.
Global Instance uniq_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) uniq_bor.
Proof. intros ??[]; split; by apply uniq_mono. Qed.
Global Instance uniq_type_contractive κ : TypeContractive (uniq_bor κ).
Proof. solve_type_proper. Qed.
Global Instance uniq_ne κ : NonExpansive (uniq_bor κ).
Proof. apply type_contractive_ne, _. Qed.
Global Instance uniq_send κ ty :
Send ty Send (uniq_bor κ ty).
Proof.
iIntros (Hsend tid1 tid2 [|[[]|][]]) "H"; try done.
iApply bor_iff; last done. iNext. iModIntro. iApply bi.equiv_iff.
do 3 f_equiv. iSplit; iIntros "."; by iApply Hsend.
Qed.
Global Instance uniq_sync κ ty :
Sync ty Sync (uniq_bor κ ty).
Proof.
iIntros (Hsync κ' tid1 tid2 l) "H". iDestruct "H" as (l') "[Hm #Hshr]".
iExists l'. iFrame "Hm". iModIntro. iIntros (F q) "% Htok".
iMod ("Hshr" with "[] Htok") as "Hfin"; first done. iClear "Hshr".
iModIntro. iNext. iMod "Hfin" as "[Hshr $]". iApply Hsync. done.
Qed.
End uniq_bor.
Notation "&uniq{ κ }" := (uniq_bor κ) (format "&uniq{ κ }") : lrust_type_scope.
Section typing.
Context `{!typeGS Σ}.
Lemma uniq_mono' E L κ1 κ2 ty1 ty2 :
lctx_lft_incl E L κ2 κ1 eqtype E L ty1 ty2
subtype E L (&uniq{κ1}ty1) (&uniq{κ2}ty2).
Proof. by intros; apply uniq_mono. Qed.
Lemma uniq_proper' E L κ1 κ2 ty1 ty2 :
lctx_lft_eq E L κ1 κ2 eqtype E L ty1 ty2 eqtype E L (&uniq{κ1}ty1) (&uniq{κ2}ty2).
Proof. by intros; apply uniq_proper. Qed.
Lemma tctx_reborrow_uniq E L p ty κ κ' :
lctx_lft_incl E L κ' κ
tctx_incl E L [p &uniq{κ}ty] [p &uniq{κ'}ty; p {κ'} &uniq{κ}ty].
Proof.
iIntros (Hκκ' tid ??) "#LFT HE HL H". iDestruct (Hκκ' with "HL HE") as %H.
iDestruct (lft_incl_syn_sem κ' κ H) as "Hκκ'".
iFrame. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as ([[]|]) "[% Hb]"; try done.
iMod (rebor with "LFT Hκκ' Hb") as "[Hb Hext]"; first done. iModIntro.
iSplitL "Hb"; iExists _; auto.
Qed.
Lemma tctx_extract_hasty_reborrow E L p ty ty' κ κ' T :
lctx_lft_incl E L κ' κ eqtype E L ty ty'
tctx_extract_hasty E L p (&uniq{κ'}ty) ((p &uniq{κ}ty')::T)
((p {κ'} &uniq{κ}ty')::T).
Proof.
intros. apply (tctx_incl_frame_r _ [_] [_;_]). rewrite tctx_reborrow_uniq //.
by apply (tctx_incl_frame_r _ [_] [_]), subtype_tctx_incl, uniq_mono'.
Qed.
Lemma read_uniq E L κ ty :
Copy ty lctx_lft_alive E L κ typed_read E L (&uniq{κ}ty) ty (&uniq{κ}ty).
Proof.
rewrite typed_read_eq. iIntros (Hcopy Halive) "!>".
iIntros ([[]|] tid F qmax qL ?) "#LFT #HE Htl HL Hown"; try done.
iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first solve_ndisj.
iMod (bor_acc with "LFT Hown Hκ") as "[H↦ Hclose']"; first solve_ndisj.
iDestruct "H↦" as (vl) "[>H↦ #Hown]".
iDestruct (ty_size_eq with "Hown") as "#>%". iIntros "!>".
iExists _, _, _. iSplit; first done. iFrame "∗#". iIntros "H↦".
iMod ("Hclose'" with "[H↦]") as "[$ Htok]"; first by iExists _; iFrame.
by iMod ("Hclose" with "Htok") as "($ & $ & $)".
Qed.
Lemma write_uniq E L κ ty :
lctx_lft_alive E L κ typed_write E L (&uniq{κ}ty) ty (&uniq{κ}ty).
Proof.
rewrite typed_write_eq. iIntros (Halive) "!>".
iIntros ([[]|] tid F qmax qL ?) "#LFT HE HL Hown"; try done.
iMod (Halive with "HE HL") as (q) "[Htok Hclose]"; first solve_ndisj.
iMod (bor_acc with "LFT Hown Htok") as "[H↦ Hclose']"; first solve_ndisj.
iDestruct "H↦" as (vl) "[>H↦ Hown]". rewrite ty.(ty_size_eq).
iDestruct "Hown" as ">%". iModIntro. iExists _, _. iSplit; first done.
iFrame. iIntros "Hown". iDestruct "Hown" as (vl') "[H↦ Hown]".
iMod ("Hclose'" with "[H↦ Hown]") as "[$ Htok]"; first by iExists _; iFrame.
by iMod ("Hclose" with "Htok") as "($ & $ & $)".
Qed.
End typing.
Global Hint Resolve uniq_mono' uniq_proper' write_uniq read_uniq : lrust_typing.
Global Hint Resolve tctx_extract_hasty_reborrow | 10 : lrust_typing.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Export type.
From iris.prelude Require Import options.
Section util.
Context `{!typeGS Σ}.
(* Delayed sharing is used by various types; in particular own and uniq.
It comes in two flavors: Borrows of "later something" and borrows of
"borrowed something".
TODO: Figure out a nice way to generalize that; the two proofs below are too
similar. *)
(* This is somewhat the general pattern here... but it doesn't seem
easy to make this usable in Coq, with the arbitrary quantifiers
and things. Also, it actually works not just for borrows but for
anything that you can split into a timeless and a persistent
part.
Lemma delay_borrow_step :
lfeE ⊆ N → (∀ x, Persistent (Post x)) →
lft_ctx -∗ &{κ} P -∗
□ (∀ x, &{κ} P -∗ Pre x -∗ Frame x ={F1 x}[F2 x]▷=∗ Post x ∗ Frame x) ={N}=∗
□ (∀ x, Pre x -∗ Frame x ={F1 x}[F2 x]▷=∗ Post x ∗ Frame x).
*)
Lemma delay_sharing_later N κ l ty tid :
lftN N
lft_ctx -∗ &{κ}( l ↦∗: ty_own ty tid) ={N}=∗
(F : coPset) (q : Qp),
⌜↑shrN lftN F -∗ (q).[κ] ={F}[F shrN]▷=∗ ty.(ty_shr) κ tid l (q).[κ].
Proof.
iIntros (?) "#LFT Hbor". rewrite bor_unfold_idx.
iDestruct "Hbor" as (i) "(#Hpb&Hpbown)".
iMod (inv_alloc shrN _ (idx_bor_own 1 i ty_shr ty κ tid l)%I
with "[Hpbown]") as "#Hinv"; first by eauto.
iIntros "!> !> * % Htok".
iMod (inv_acc with "Hinv") as "[INV Hclose]"; first solve_ndisj.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iMod (bor_later_tok with "LFT [Hbtok] Htok") as "Hdelay"; first solve_ndisj.
{ rewrite bor_unfold_idx. eauto. }
iModIntro. iNext. iMod "Hdelay" as "[Hb Htok]".
iMod (ty.(ty_share) with "LFT Hb Htok") as "[#$ $]"; first solve_ndisj.
iApply "Hclose". auto.
- iMod fupd_mask_subseteq as "Hclose'"; first solve_ndisj. iModIntro.
iNext. iMod "Hclose'" as "_". iMod ("Hclose" with "[]") as "_"; by eauto.
Qed.
Lemma delay_sharing_nested N κ κ' κ'' l ty tid :
lftN N
lft_ctx -∗ (κ'' κ κ') -∗ &{κ'}(&{κ}(l ↦∗: ty_own ty tid)) ={N}=∗
(F : coPset) (q : Qp),
⌜↑shrN lftN F -∗ (q).[κ''] ={F}[F shrN]▷=∗ ty.(ty_shr) κ'' tid l (q).[κ''].
Proof.
iIntros (?) "#LFT #Hincl Hbor". rewrite bor_unfold_idx.
iDestruct "Hbor" as (i) "(#Hpb&Hpbown)".
iMod (inv_alloc shrN _ (idx_bor_own 1 i ty_shr ty κ'' tid l)%I
with "[Hpbown]") as "#Hinv"; first by eauto.
iIntros "!> !> * % Htok".
iMod (inv_acc with "Hinv") as "[INV Hclose]"; first solve_ndisj.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iMod (bor_unnest with "LFT [Hbtok]") as "Hb"; first solve_ndisj.
{ iApply bor_unfold_idx. eauto. }
iModIntro. iNext. iMod "Hb".
iMod (ty.(ty_share) with "LFT [Hb] Htok") as "[#Hshr $]"; first solve_ndisj.
{ iApply bor_shorten; done. }
iMod ("Hclose" with "[]") as "_"; auto.
- iMod fupd_mask_subseteq as "Hclose'"; last iModIntro; first solve_ndisj.
iNext. iMod "Hclose'" as "_". iMod ("Hclose" with "[]") as "_"; by eauto.
Qed.
End util.
From iris.algebra Require Import gmap auth frac.
From iris.proofmode Require Import proofmode.
From lrust.lifetime Require Export lifetime.
From iris.prelude Require Import options.
(** Atomic persistent bors *)
(* TODO : update the TEX with the fact that we can choose the namespace. *)
Definition at_bor `{!invGS Σ, !lftGS Σ userE} κ N (P : iProp Σ) :=
( i, &{κ,i}P
(N ## lftN inv N (idx_bor_own 1 i)
N = lftN inv N ( q, idx_bor_own q i)))%I.
Notation "&at{ κ , N }" := (at_bor κ N) (format "&at{ κ , N }") : bi_scope.
Section atomic_bors.
Context `{!invGS Σ, !lftGS Σ userE} (P : iProp Σ) (N : namespace).
Global Instance at_bor_ne κ n : Proper (dist n ==> dist n) (at_bor κ N).
Proof. solve_proper. Qed.
Global Instance at_bor_contractive κ : Contractive (at_bor κ N).
Proof. solve_contractive. Qed.
Global Instance at_bor_proper κ : Proper ((⊣⊢) ==> (⊣⊢)) (at_bor κ N).
Proof. solve_proper. Qed.
Lemma at_bor_iff κ P' : (P P') -∗ &at{κ, N} P -∗ &at{κ, N} P'.
Proof.
iIntros "HPP' H". iDestruct "H" as (i) "[HP ?]". iExists i. iFrame.
iApply (idx_bor_iff with "HPP' HP").
Qed.
Global Instance at_bor_persistent κ N P : Persistent (&at{κ, N} P) := _.
Lemma bor_share E κ :
N ## lftN &{κ}P ={E}=∗ &at{κ, N}P.
Proof.
iIntros (HN) "HP". rewrite bor_unfold_idx. iDestruct "HP" as (i) "(#?&Hown)".
iExists i. iFrame "#".
iLeft. iSplitR; first done. by iMod (inv_alloc with "[Hown]") as "$"; auto.
Qed.
Lemma bor_share_lftN E κ :
lftN E &{κ}P ={E}=∗ &at{κ, lftN}P.
Proof.
iIntros (?) "HP". rewrite bor_unfold_idx. iDestruct "HP" as (i) "(#?&Hown)".
iExists i. iFrame "#". subst.
iRight. iSplitR; first done. by iMod (inv_alloc with "[Hown]") as "$"; auto.
Qed.
Lemma at_bor_acc E κ :
lftN E N E
lft_ctx -∗ &at{κ,N}P ={E,E∖↑N∖↑lftN}=∗ P (P ={E∖↑N∖↑lftN,E}=∗ True)
[κ] |={E∖↑N∖↑lftN,E}=> True.
Proof.
iIntros (??) "#LFT #HP". iDestruct "HP" as (i) "#[Hidx [[% Hinv]|[% Hinv]]]".
- iInv N as ">Hi" "Hclose". iMod (idx_bor_atomic_acc with "LFT Hidx Hi")
as "[[HP Hclose']|[H† Hclose']]"; first solve_ndisj.
+ iLeft. iFrame. iIntros "!>HP". iMod ("Hclose'" with "HP"). by iApply "Hclose".
+ iRight. iFrame. iIntros "!>". iMod "Hclose'". by iApply "Hclose".
- subst. rewrite difference_twice_L. iInv lftN as (q') ">[Hq'0 Hq'1]" "Hclose".
iMod ("Hclose" with "[Hq'1]") as "_"; first solve_ndisj.
iMod (idx_bor_atomic_acc with "LFT Hidx Hq'0") as "[[HP Hclose]|[H† Hclose]]"; first done.
+ iLeft. iFrame. iIntros "!>HP". by iMod ("Hclose" with "HP").
+ iRight. iFrame. iIntros "!>". by iMod "Hclose".
Qed.
Lemma at_bor_acc_tok E q κ :
lftN E N E
lft_ctx -∗ &at{κ,N}P -∗ q.[κ] ={E,E∖↑N}=∗ P (P ={E∖↑N,E}=∗ q.[κ]).
Proof.
iIntros (??) "#LFT #HP Hκ". iDestruct "HP" as (i) "#[Hidx [[% Hinv]|[% Hinv]]]".
- iInv N as ">Hi" "Hclose".
iMod (idx_bor_acc with "LFT Hidx Hi Hκ") as "[$ Hclose']"; first solve_ndisj.
iIntros "!> H". iMod ("Hclose'" with "H") as "[? $]". by iApply "Hclose".
- iMod (at_bor_acc with "LFT []") as "[[$ Hclose]|[H† _]]"; try done.
+ iExists i. auto.
+ subst. rewrite difference_twice_L. iIntros "!>HP". by iMod ("Hclose" with "HP").
+ iDestruct (lft_tok_dead with "Hκ H†") as "[]".
Qed.
Lemma at_bor_shorten κ κ': κ κ' -∗ &at{κ',N}P -∗ &at{κ,N}P.
Proof.
iIntros "#H⊑ H". iDestruct "H" as (i) "[??]". iExists i. iFrame.
by iApply (idx_bor_shorten with "H⊑").
Qed.
Lemma at_bor_fake E κ:
lftN E N ## lftN lft_ctx -∗ [κ] ={E}=∗ &at{κ,N}P.
Proof.
iIntros (??) "#LFT#H†". iApply (bor_share with "[>]"); try done.
by iApply (bor_fake with "LFT H†").
Qed.
Lemma at_bor_fake_lftN E κ:
lftN E lft_ctx -∗ [κ] ={E}=∗ &at{κ,lftN}P.
Proof.
iIntros (?) "#LFT#H†". iApply (bor_share_lftN with "[>]"); try done.
by iApply (bor_fake with "LFT H†").
Qed.
End atomic_bors.
Lemma at_bor_acc_lftN `{!invGS Σ, !lftGS Σ userE} (P : iProp Σ) E κ :
lftN E
lft_ctx -∗ &at{κ,lftN}P ={E,E∖↑lftN}=∗ P (P ={E∖↑lftN,E}=∗ True)
[κ] |={E∖↑lftN,E}=> True.
Proof.
iIntros (?) "H".
iDestruct (at_bor_acc _ lftN with "H") as "H"; [done..|].
by rewrite difference_twice_L.
Qed.
Global Typeclasses Opaque at_bor.
From iris.proofmode Require Import proofmode.
From iris.bi Require Import fractional.
From iris.algebra Require Import frac.
From lrust.lifetime Require Export at_borrow.
From iris.prelude Require Import options.
Class frac_borG Σ := frac_borG_inG :: inG Σ fracR.
Local Definition frac_bor_inv `{!invGS Σ, !lftGS Σ userE, !frac_borG Σ}
(φ : Qp iProp Σ) γ κ' :=
( q, φ q own γ q (q = 1%Qp q', (q + q' = 1)%Qp q'.[κ']))%I.
Definition frac_bor `{!invGS Σ, !lftGS Σ userE, !frac_borG Σ} κ (φ : Qp iProp Σ) :=
( γ κ', κ κ' &at{κ',lftN} (frac_bor_inv φ γ κ'))%I.
Notation "&frac{ κ }" := (frac_bor κ) (format "&frac{ κ }") : bi_scope.
Section frac_bor.
Context `{!invGS Σ, !lftGS Σ userE, !frac_borG Σ} (φ : Qp iProp Σ).
Implicit Types E : coPset.
Global Instance frac_bor_contractive κ n :
Proper (pointwise_relation _ (dist_later n) ==> dist n) (frac_bor κ).
Proof. rewrite /frac_bor /frac_bor_inv. solve_contractive. Qed.
Global Instance frac_bor_ne κ n :
Proper (pointwise_relation _ (dist n) ==> dist n) (frac_bor κ).
Proof. rewrite /frac_bor /frac_bor_inv. solve_proper. Qed.
Global Instance frac_bor_proper κ :
Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (frac_bor κ).
Proof. rewrite /frac_bor /frac_bor_inv. solve_proper. Qed.
Lemma frac_bor_iff κ φ' :
( q, φ q φ' q) -∗ &frac{κ} φ -∗ &frac{κ} φ'.
Proof.
iIntros "#Hφφ' H". iDestruct "H" as (γ κ') "[? Hφ]". iExists γ, κ'. iFrame.
iApply (at_bor_iff with "[Hφφ'] Hφ"). iNext. iModIntro.
iSplit; iIntros "H"; iDestruct "H" as (q) "[H ?]"; iExists q; iFrame; by iApply "Hφφ'".
Qed.
Global Instance frac_bor_persistent κ : Persistent (&frac{κ}φ) := _.
Lemma bor_fracture E κ :
lftN E lft_ctx -∗ &{κ}(φ 1%Qp) ={E}=∗ &frac{κ}φ.
Proof.
iIntros (?) "#LFT Hφ". iMod (own_alloc 1%Qp) as (γ) "?"; first done.
iMod (bor_acc_atomic_strong with "LFT Hφ") as "[H|[H† Hclose]]"; first done.
- iDestruct "H" as (κ') "(#Hκκ' & Hφ & Hclose)".
iMod ("Hclose" with "[] [-]") as "Hφ"; last first.
{ iExists γ, κ'. iFrame "#". iApply (bor_share_lftN with "Hφ"); auto. }
{ iExists 1%Qp. iFrame. eauto. }
iIntros "!>Hφ H†!>". iNext. iDestruct "Hφ" as (q') "(Hφ & _ & [%|Hκ])"; first by subst.
iDestruct "Hκ" as (q'') "[_ Hκ]".
iDestruct (lft_tok_dead with "Hκ H†") as "[]".
- iMod "Hclose" as "_"; last first.
iExists γ, κ. iSplitR; first by iApply lft_incl_refl.
by iApply at_bor_fake_lftN.
Qed.
Lemma frac_bor_atomic_acc E κ :
lftN E
lft_ctx -∗ &frac{κ}φ ={E,E∖↑lftN}=∗ ( q, φ q ( φ q ={E∖↑lftN,E}=∗ True))
[κ] |={E∖↑lftN,E}=> True.
Proof.
iIntros (?) "#LFT #Hφ". iDestruct "Hφ" as (γ κ') "[Hκκ' Hshr]".
iMod (at_bor_acc_lftN with "LFT Hshr") as "[[Hφ Hclose]|[H† Hclose]]"; try done.
- iLeft. iDestruct "Hφ" as (q) "(Hφ & Hγ & H)". iExists q. iFrame.
iIntros "!>Hφ". iApply "Hclose". iExists q. iFrame.
- iRight. iMod "Hclose" as "_". iMod (lft_incl_dead with "Hκκ' H†") as "$"; first done.
iApply fupd_mask_intro_subseteq; first set_solver. done.
Qed.
Local Lemma frac_bor_trade1 γ κ' q :
( q1 q2, φ (q1+q2)%Qp φ q1 φ q2) -∗
frac_bor_inv φ γ κ' own γ q φ q -∗
(frac_bor_inv φ γ κ' q.[κ']).
Proof.
iIntros "#Hφ (H & Hown & Hφ1)". iNext.
iDestruct "H" as () "(Hφqφ & Hown' & [%|Hq])".
{ subst. iCombine "Hown'" "Hown" as "Hown".
by iDestruct (own_valid with "Hown") as %Hval%Qp.not_add_le_l. }
rewrite /frac_bor_inv. iApply bi.sep_exist_r. iExists (q + )%Qp.
iDestruct "Hq" as (q' Hqφq') "Hq'κ". iCombine "Hown'" "Hown" as "Hown".
iDestruct (own_valid with "Hown") as %Hval. rewrite comm_L. iFrame "Hown".
iCombine "Hφ1 Hφqφ" as "Hφq". iDestruct ("Hφ" with "Hφq") as "$".
assert (q q')%Qp as [[r ->]%Qp.lt_sum|<-]%Qp.le_lteq.
{ apply (Qp.add_le_mono_l _ _ ). by rewrite Hqφq'. }
- iDestruct "Hq'κ" as "[$ Hr]".
iRight. iExists _. iIntros "{$Hr} !%".
by rewrite (comm_L Qp.add q) -assoc_L.
- iFrame "Hq'κ". iLeft. iPureIntro. rewrite comm_L. done.
Qed.
Local Lemma frac_bor_trade2 γ κ' qκ' :
( q1 q2, φ (q1+q2)%Qp φ q1 φ q2) -∗
frac_bor_inv φ γ κ' qκ'.[κ'] -∗
(frac_bor_inv φ γ κ' q0 q1, qκ' = (q0 + q1)%Qp q1.[κ'] own γ q0 φ q0).
Proof.
iIntros "#Hφ [H Hκ']". iNext.
iDestruct "H" as () "(Hφqφ & Hown & Hq)".
destruct (Qp.lower_bound qκ' ) as (qq & qκ'0 & qφ0 & Hqκ' & Hqφ).
iApply bi.sep_exist_l. iExists qq.
iApply bi.sep_exist_l. iExists qκ'0.
subst qκ' . rewrite /frac_bor_inv.
iApply bi.sep_exist_r. iExists qφ0.
iDestruct ("Hφ" with "Hφqφ") as "[$ $] {Hφ}".
iDestruct "Hown" as "[$ $]".
iDestruct "Hκ'" as "[Hκ' $]".
iSplit; last done.
iDestruct "Hq" as "[%|Hq]".
- iRight. iExists qq. iFrame. iPureIntro.
by rewrite (comm _ qφ0).
- iDestruct "Hq" as (q') "[% Hq'κ]". iRight. iExists (qq + q')%Qp.
iFrame. iPureIntro.
rewrite assoc (comm _ _ qq). done.
Qed.
Lemma frac_bor_acc' E q κ :
lftN E
lft_ctx -∗ ( q1 q2, φ (q1+q2)%Qp φ q1 φ q2) -∗
&frac{κ}φ -∗ q.[κ] ={E}=∗ q', φ q' ( φ q' ={E}=∗ q.[κ]).
Proof.
iIntros (?) "#LFT #Hφ Hfrac Hκ". unfold frac_bor.
iDestruct "Hfrac" as (γ κ') "#[#Hκκ' Hshr]".
iMod (lft_incl_acc with "Hκκ' Hκ") as (qκ') "[[Hκ1 Hκ2] Hclose]"; first done.
iMod (at_bor_acc_tok with "LFT Hshr Hκ1") as "[H Hclose']"; try done.
iDestruct (frac_bor_trade2 with "Hφ [$H $Hκ2]") as "[H Htrade]".
iDestruct "Htrade" as (q0 q1) "(>Hq & >Hκ2 & >Hown & Hqφ)".
iDestruct "Hq" as %Hq.
iMod ("Hclose'" with "H") as "Hκ1".
iExists q0. iFrame "Hqφ". iIntros "!>Hqφ".
iMod (at_bor_acc_tok with "LFT Hshr Hκ1") as "[H Hclose']"; try done.
iDestruct (frac_bor_trade1 with "Hφ [$H $Hown $Hqφ]") as "[H >Hκ3]".
iMod ("Hclose'" with "H") as "Hκ1".
iApply "Hclose". iFrame "Hκ1". rewrite Hq. iFrame.
Qed.
Lemma frac_bor_acc E q κ `{!Fractional φ} :
lftN E
lft_ctx -∗ &frac{κ}φ -∗ q.[κ] ={E}=∗ q', φ q' ( φ q' ={E}=∗ q.[κ]).
Proof.
iIntros (?) "LFT". iApply (frac_bor_acc' with "LFT"); first done.
iIntros "!>*". rewrite fractional. iSplit; auto.
Qed.
Lemma frac_bor_shorten κ κ' : κ κ' -∗ &frac{κ'}φ -∗ &frac{κ}φ.
Proof.
iIntros "Hκκ' H". iDestruct "H" as (γ κ0) "[#H⊑ ?]". iExists γ, κ0. iFrame.
iApply (lft_incl_trans with "Hκκ' []"). auto.
Qed.
Lemma frac_bor_fake E κ :
lftN E lft_ctx -∗ [κ] ={E}=∗ &frac{κ}φ.
Proof.
iIntros (?) "#LFT#H†". iApply (bor_fracture with "LFT [>]"); first done.
by iApply (bor_fake with "LFT").
Qed.
End frac_bor.
Lemma frac_bor_lft_incl `{!invGS Σ, !lftGS Σ userE, !frac_borG Σ} κ κ' q:
lft_ctx -∗ &frac{κ}(λ q', (q * q').[κ']) -∗ κ κ'.
Proof.
iIntros "#LFT#Hbor". iApply lft_incl_intro. iModIntro. iSplitR.
- iIntros (q') "Hκ'".
(* FIXME we should probably have a lemma for this in Iris *)
assert (Fractional (λ q' : Qp, (q * q').[κ']))%I.
{ intros ??. rewrite Qp.mul_add_distr_l lft_tok_fractional //. }
iMod (frac_bor_acc with "LFT Hbor Hκ'") as (q'') "[>? Hclose]"; first done.
iExists _. iFrame. iIntros "!>Hκ'". iApply "Hclose". auto.
- iIntros "H†'".
iMod (frac_bor_atomic_acc with "LFT Hbor") as "[H|[$ $]]"; first done.
iDestruct "H" as (q') "[>Hκ' _]".
iDestruct (lft_tok_dead with "Hκ' H†'") as "[]".
Qed.
Global Typeclasses Opaque frac_bor.
From lrust.lifetime Require Export lifetime_sig.
From lrust.lifetime.model Require definitions primitive accessors faking borrow
borrow_sep reborrow creation.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Module Export lifetime : lifetime_sig.
Definition lft := gmultiset positive.
Include definitions.
Include primitive.
Include borrow.
Include faking.
Include borrow_sep.
Include reborrow.
Include accessors.
Include creation.
End lifetime.
Notation lft_intersect_list l := (foldr () static l).
Global Instance lft_inhabited : Inhabited lft := populate static.
Canonical lftO := leibnizO lft.
Definition lft_incl_syntactic (κ κ' : lft) : Prop := κ'', κ'' κ' = κ.
Infix "⊑ˢʸⁿ" := lft_incl_syntactic (at level 40) : stdpp_scope.
Section derived.
Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft.
Lemma lft_kill_atomic (Λ : atomic_lft) :
lft_ctx -∗ (1.[@Λ] ={lftN userE}[userE]▷=∗ [@Λ]).
Proof.
iIntros "#LFT".
rewrite -(big_sepS_singleton (λ x, 1.[@x])%I)
-(big_sepS_singleton (λ x, [@x])%I).
by iApply (lft_kill_atomics with "LFT").
Qed.
Lemma lft_create_atomic E :
lftN E
lft_ctx ={E}=∗ Λ, 1.[@Λ].
Proof.
iIntros (?) "#LFT".
iMod (lft_create_strong (λ _, True) with "LFT") as (κ _) "$"=>//.
by apply pred_infinite_True.
Qed.
Lemma lft_create E :
lftN E
lft_ctx ={E}=∗ κ, 1.[κ] (1.[κ] ={lftN userE}[userE]▷=∗ [κ]).
Proof.
iIntros (?) "#LFT".
iMod (lft_create_atomic with "LFT") as "[% $]"=>//.
by iApply lft_kill_atomic.
Qed.
(* Deriving this just to prove that it can be derived.
(It is in the signature only for code sharing reasons.) *)
Lemma bor_shorten_ κ κ' P : κ κ' -∗ &{κ'}P -∗ &{κ}P.
Proof.
iIntros "#Hκκ'". rewrite !bor_unfold_idx. iDestruct 1 as (i) "[#? ?]".
iExists i. iFrame. by iApply idx_bor_shorten.
Qed.
Lemma bor_acc_atomic_cons E κ P :
lftN E
lft_ctx -∗ &{κ} P ={E,E∖↑lftN}=∗
( P Q, ( Q ={userE}=∗ P) -∗ Q ={E∖↑lftN,E}=∗ &{κ} Q)
([κ] |={E∖↑lftN,E}=> True).
Proof.
iIntros (?) "#LFT HP".
iMod (bor_acc_atomic_strong with "LFT HP") as "[H|[??]]"; first done.
- iLeft. iDestruct "H" as (κ') "(#Hκκ' & $ & Hclose)". iIntros "!>* HPQ HQ".
iMod ("Hclose" with "[HPQ] HQ") as "Hb".
+ iNext. iIntros "? _". by iApply "HPQ".
+ iApply (bor_shorten with "Hκκ' Hb").
- iRight. by iFrame.
Qed.
Lemma bor_acc_atomic E κ P :
lftN E
lft_ctx -∗ &{κ}P ={E,E∖↑lftN}=∗
( P ( P ={E∖↑lftN,E}=∗ &{κ}P)) ([κ] |={E∖↑lftN,E}=> True).
Proof.
iIntros (?) "#LFT HP".
iMod (bor_acc_atomic_cons with "LFT HP") as "[[HP Hclose]|[? ?]]"; first done.
- iLeft. iIntros "!> {$HP} HP". iMod ("Hclose" with "[] HP"); auto.
- iRight. by iFrame.
Qed.
Lemma bor_acc_cons E q κ P :
lftN E
lft_ctx -∗ &{κ} P -∗ q.[κ] ={E}=∗ P
Q, ( Q ={userE}=∗ P) -∗ Q ={E}=∗ &{κ} Q q.[κ].
Proof.
iIntros (?) "#LFT HP Htok".
iMod (bor_acc_strong with "LFT HP Htok") as (κ') "(#Hκκ' & $ & Hclose)"; first done.
iIntros "!>* HPQ HQ". iMod ("Hclose" with "[HPQ] HQ") as "[Hb $]".
- iNext. iIntros "? _". by iApply "HPQ".
- iApply (bor_shorten with "Hκκ' Hb").
Qed.
Lemma bor_acc E q κ P :
lftN E
lft_ctx -∗ &{κ}P -∗ q.[κ] ={E}=∗ P ( P ={E}=∗ &{κ}P q.[κ]).
Proof.
iIntros (?) "#LFT HP Htok".
iMod (bor_acc_cons with "LFT HP Htok") as "($ & Hclose)"; first done.
iIntros "!>HP". iMod ("Hclose" with "[] HP") as "[$ $]"; auto.
Qed.
Lemma bor_exists {A} (Φ : A iProp Σ) `{!Inhabited A} E κ :
lftN E
lft_ctx -∗ &{κ}( x, Φ x) ={E}=∗ x, &{κ}(Φ x).
Proof.
iIntros (?) "#LFT Hb".
iMod (bor_acc_atomic_cons with "LFT Hb") as "[H|[H† >_]]"; first done.
- iDestruct "H" as "[HΦ Hclose]". iDestruct "HΦ" as (x) "HΦ".
iExists x. iApply ("Hclose" with "[] HΦ"). iIntros "!> ?"; eauto.
- iExists inhabitant. by iApply (bor_fake with "LFT").
Qed.
Lemma bor_or E κ P Q :
lftN E
lft_ctx -∗ &{κ}(P Q) ={E}=∗ (&{κ}P &{κ}Q).
Proof.
iIntros (?) "#LFT H". rewrite bi.or_alt.
(* The (A:=...) is needed due to Coq bug #5458 *)
iMod (bor_exists (A:=bool) with "LFT H") as ([]) "H"; auto.
Qed.
Lemma bor_iff κ P P' : (P P') -∗ &{κ}P -∗ &{κ}P'.
Proof.
rewrite !bor_unfold_idx. iIntros "#HP Hbor".
iDestruct "Hbor" as (i) "[Hbor Htok]". iExists i. iFrame "Htok".
iApply idx_bor_iff; done.
Qed.
Lemma bor_iff_proper κ P P': (P P') -∗ (&{κ}P &{κ}P').
Proof.
iIntros "#HP !>". iSplit; iIntros "?"; iApply bor_iff; try done.
iNext. iModIntro. iSplit; iIntros "?"; iApply "HP"; done.
Qed.
Lemma bor_later E κ P :
lftN E
lft_ctx -∗ &{κ}( P) ={E}[E∖↑lftN]▷=∗ &{κ}P.
Proof.
iIntros (?) "#LFT Hb".
iMod (bor_acc_atomic_cons with "LFT Hb") as "[H|[H† Hclose]]"; first done.
- iDestruct "H" as "[HP Hclose]". iModIntro. iNext.
iApply ("Hclose" with "[] HP"). by iIntros "!> $".
- iIntros "!> !>". iMod "Hclose" as "_". by iApply (bor_fake with "LFT").
Qed.
Lemma bor_later_tok E q κ P :
lftN E
lft_ctx -∗ &{κ}( P) -∗ q.[κ] ={E}▷=∗ &{κ}P q.[κ].
Proof.
iIntros (?) "#LFT Hb Htok".
iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose]"; first done.
iModIntro. iNext. iApply ("Hclose" with "[] HP"). by iIntros "!> $".
Qed.
Lemma bor_persistent_notok P `{!Persistent P} E κ :
lftN E
lft_ctx -∗ &{κ}P ={E}=∗ P [κ].
Proof.
iIntros (?) "#LFT Hb".
iMod (bor_acc_atomic with "LFT Hb") as "[[#HP Hob]|[#H† Hclose]]"; first done.
- iMod ("Hob" with "HP") as "_". auto.
- iMod "Hclose" as "_". auto.
Qed.
Lemma bor_persistent P `{!Persistent P} E κ q :
lftN E
lft_ctx -∗ &{κ}P -∗ q.[κ] ={E}=∗ P q.[κ].
Proof.
iIntros (?) "#LFT Hb Htok".
iMod (bor_acc with "LFT Hb Htok") as "[#HP Hob]"; first done.
by iMod ("Hob" with "HP") as "[_ $]".
Qed.
Lemma later_bor_static E P :
lftN E
lft_ctx -∗ P ={E}=∗ &{static} P.
Proof.
iIntros (?) "#LFT HP". iMod (bor_create with "LFT HP") as "[$ _]"; done.
Qed.
Lemma bor_static_later E P :
lftN E
lft_ctx -∗ &{static} P ={E}=∗ P.
Proof.
iIntros (?) "LFT HP". iMod (bor_acc _ 1%Qp with "LFT HP []") as "[$ _]"; try done.
iApply lft_tok_static.
Qed.
Lemma rebor E κ κ' P :
lftN E
lft_ctx -∗ κ' κ -∗ &{κ}P ={E}=∗ &{κ'}P ([κ'] ={E}=∗ &{κ}P).
Proof.
iIntros (?) "#LFT #Hκ'κ Hbor". rewrite [(&{κ}P)%I]bor_unfold_idx.
iDestruct "Hbor" as (i) "[#Hbor Hidx]".
iMod (bor_create _ κ' with "LFT Hidx") as "[Hidx Hinh]"; first done.
iMod (idx_bor_unnest with "LFT Hbor Hidx") as "Hbor'"; first done.
iDestruct (bor_shorten with "[] Hbor'") as "$".
{ iApply lft_incl_glb; first done. iApply lft_incl_refl. }
iIntros "!> H†". iMod ("Hinh" with "H†") as ">Hidx". auto with iFrame.
Qed.
Lemma bor_unnest E κ κ' P :
lftN E
lft_ctx -∗ &{κ'} (&{κ} P) ={E}▷=∗ &{κ κ'} P.
Proof.
iIntros (?) "#LFT Hbor".
rewrite ->(bor_unfold_idx _ P).
iMod (bor_exists with "LFT Hbor") as (i) "Hbor"; [done|].
iMod (bor_sep with "LFT Hbor") as "[Hidx Hbor]"; [done|].
iMod (bor_persistent_notok with "LFT Hidx") as "#[Hidx|H†]"; [done| |].
- iIntros "!>". iNext. by iApply (idx_bor_unnest with "LFT Hidx Hbor").
- iApply (bor_fake with "LFT"); [done|]. rewrite -lft_dead_or. auto.
Qed.
Lemma lft_incl_static κ : κ static.
Proof.
iApply lft_incl_intro. iIntros "!>". iSplitR.
- iIntros (q) "?". iExists 1%Qp. iSplitR; last by auto. by iApply lft_tok_static.
- iIntros "Hst". by iDestruct (lft_dead_static with "Hst") as "[]".
Qed.
Lemma lft_intersect_list_elem_of_incl κs κ :
κ κs lft_intersect_list κs κ.
Proof.
induction 1 as [|???? IH].
- iApply lft_intersect_incl_l.
- iApply lft_incl_trans; last iApply IH. (* FIXME: Why does "done" not do this? Looks like "assumption" to me. *)
iApply lft_intersect_incl_r.
Qed.
Lemma lft_eternalize E q κ :
q.[κ] ={E}=∗ static κ.
Proof.
iIntros "Hκ". iMod (inv_alloc lftN _ ( q, q.[κ])%I with "[Hκ]") as "#Hnv".
{ iExists _. done. }
iApply lft_incl_intro. iIntros "!> !>". iSplitL.
- iIntros (q') "$". iInv lftN as ">Hκ" "Hclose". iDestruct "Hκ" as (q'') "[Hκ1 Hκ2]".
iMod ("Hclose" with "[Hκ2]") as "_". { iExists _. done. }
iModIntro. iExists _. iFrame. iIntros "_". done.
- iIntros "H†". iInv lftN as ">Hκ" "_". iDestruct "Hκ" as (q'') "Hκ".
iDestruct (lft_tok_dead with "Hκ H†") as "[]".
Qed.
(** Syntactic lifetime inclusion *)
Lemma lft_incl_syn_sem κ κ' :
κ ˢʸⁿ κ' κ κ'.
Proof. intros [z Hxy]. rewrite -Hxy. by apply lft_intersect_incl_r. Qed.
Lemma lft_intersect_incl_syn_l κ κ' : κ κ' ˢʸⁿ κ.
Proof. by exists κ'; rewrite (comm _ κ κ'). Qed.
Lemma lft_intersect_incl_syn_r κ κ' : κ κ' ˢʸⁿ κ'.
Proof. by exists κ. Qed.
Lemma lft_incl_syn_refl κ : κ ˢʸⁿ κ.
Proof. exists static; by rewrite left_id. Qed.
Lemma lft_incl_syn_trans κ κ' κ'' :
κ ˢʸⁿ κ' κ' ˢʸⁿ κ'' κ ˢʸⁿ κ''.
Proof.
intros [κ0 Hκ0] [κ'0 Hκ'0].
rewrite -Hκ0 -Hκ'0 assoc.
by apply lft_intersect_incl_syn_r.
Qed.
Lemma lft_intersect_mono_syn κ1 κ1' κ2 κ2' :
κ1 ˢʸⁿ κ1' κ2 ˢʸⁿ κ2' (κ1 κ2) ˢʸⁿ (κ1' κ2').
Proof.
intros [κ1'' Hκ1] [κ2'' Hκ2].
exists (κ1'' κ2'').
rewrite -!assoc (comm _ κ2'' _).
rewrite -assoc assoc (comm _ κ2' _).
by f_equal.
Qed.
Lemma lft_incl_syn_static κ : κ ˢʸⁿ static.
Proof.
exists κ; by apply lft_intersect_right_id.
Qed.
Lemma lft_intersect_list_elem_of_incl_syn κs κ :
κ κs lft_intersect_list κs ˢʸⁿ κ.
Proof.
induction 1 as [κ|κ κ' κs Hin IH].
- apply lft_intersect_incl_syn_l.
- eapply lft_incl_syn_trans; last done.
apply lft_intersect_incl_syn_r.
Qed.
Global Instance lft_incl_syn_anti_symm : AntiSymm eq (λ x y, x ˢʸⁿ y).
Proof.
intros κ1 κ2 [κ12 Hκ12] [κ21 Hκ21].
assert (κ21 = static) as Hstatic.
{ apply (lft_intersect_static_cancel_l _ κ12).
eapply (inj (κ1 ⊓.)).
rewrite assoc right_id.
eapply (inj (. κ2)).
rewrite (comm _ κ1 κ21) -assoc Hκ21 Hκ12 (comm _ κ2). done. }
by rewrite Hstatic left_id in Hκ21.
Qed.
End derived.
From iris.algebra Require Import frac.
From stdpp Require Export gmultiset strings.
From iris.base_logic.lib Require Export invariants.
From iris.base_logic.lib Require Import boxes.
From iris.bi Require Import fractional.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Definition lftN : namespace := nroot .@ "lft".
Module Type lifetime_sig.
(** CMRAs needed by the lifetime logic *)
(* We can't instantie the module parameters with inductive types, so we
have aliases here. *)
(** userE is a (disjoint) mask that is available in the "consequence" view
shift of borrow accessors. *)
Parameter lftGS' : (Σ : gFunctors) (userE : coPset), Set.
Global Notation lftGS := lftGS'.
Existing Class lftGS'.
Parameter lftGpreS' : gFunctors Set.
Global Notation lftGpreS := lftGpreS'.
Existing Class lftGpreS'.
(** * Definitions *)
Parameter lft : Type.
Parameter static : lft.
Global Declare Instance lft_intersect : Meet lft.
Parameter lft_ctx : `{!invGS Σ, !lftGS Σ userE}, iProp Σ.
Parameter lft_tok : `{!lftGS Σ userE} (q : Qp) (κ : lft), iProp Σ.
Parameter lft_dead : `{!lftGS Σ userE} (κ : lft), iProp Σ.
Parameter lft_incl : `{!invGS Σ, !lftGS Σ userE} (κ κ' : lft), iProp Σ.
Parameter bor : `{!invGS Σ, !lftGS Σ userE} (κ : lft) (P : iProp Σ), iProp Σ.
Parameter bor_idx : Type.
Parameter idx_bor_own : `{!lftGS Σ userE} (q : frac) (i : bor_idx), iProp Σ.
Parameter idx_bor : `{!invGS Σ, !lftGS Σ userE} (κ : lft) (i : bor_idx) (P : iProp Σ), iProp Σ.
Parameter atomic_lft : Type.
Parameter atomic_lft_to_lft : atomic_lft lft.
(** Our lifetime creation lemma offers allocating a lifetime that is defined
by a [positive] in some given infinite set. This operation converts the
[positive] to an atomic lifetime. *)
Parameter positive_to_atomic_lft : positive atomic_lft.
Notation positive_to_lft p := (atomic_lft_to_lft (positive_to_atomic_lft p)).
(** * Notation *)
Notation "q .[ κ ]" := (lft_tok q κ)
(format "q .[ κ ]", at level 2, left associativity) : bi_scope.
Notation "q .[@ Λ ]" := (lft_tok q (atomic_lft_to_lft Λ))
(format "q .[@ Λ ]", at level 2, left associativity) : bi_scope.
Notation "[† κ ]" := (lft_dead κ) (format "[† κ ]"): bi_scope.
Notation "[†@ Λ ]" := (lft_dead (atomic_lft_to_lft Λ)) (format "[†@ Λ ]"): bi_scope.
Notation "&{ κ }" := (bor κ) (format "&{ κ }") : bi_scope.
Notation "&{ κ , i }" := (idx_bor κ i) (format "&{ κ , i }") : bi_scope.
Infix "⊑" := lft_incl (at level 70) : bi_scope.
Section properties.
Context `{!invGS Σ, !lftGS Σ userE}.
(** * Instances *)
Global Declare Instance lft_inhabited : Inhabited lft.
Global Declare Instance lft_eq_dec : EqDecision lft.
Global Declare Instance lft_countable : Countable lft.
Global Declare Instance atomic_lft_inhabited : Inhabited atomic_lft.
Global Declare Instance atomic_lft_eq_dec : EqDecision atomic_lft.
Global Declare Instance atomic_lft_countable : Countable atomic_lft.
Global Declare Instance bor_idx_inhabited : Inhabited bor_idx.
Global Declare Instance bor_idx_eq_dec : EqDecision bor_idx.
Global Declare Instance bor_idx_countable : Countable bor_idx.
Global Declare Instance lft_intersect_comm : Comm (A:=lft) eq ().
Global Declare Instance lft_intersect_assoc : Assoc (A:=lft) eq ().
Global Declare Instance lft_intersect_inj_1 (κ : lft) : Inj eq eq (κ ⊓.).
Global Declare Instance lft_intersect_inj_2 (κ : lft) : Inj eq eq (. κ).
Global Declare Instance lft_intersect_left_id : LeftId eq static meet.
Global Declare Instance lft_intersect_right_id : RightId eq static meet.
Global Declare Instance lft_ctx_persistent : Persistent lft_ctx.
Global Declare Instance lft_dead_persistent κ : Persistent ([κ]).
Global Declare Instance lft_incl_persistent κ κ' : Persistent (κ κ').
Global Declare Instance idx_bor_persistent κ i P : Persistent (&{κ,i} P).
Global Declare Instance lft_tok_timeless q κ : Timeless (q.[κ]).
Global Declare Instance lft_dead_timeless κ : Timeless ([κ]).
Global Declare Instance idx_bor_own_timeless q i : Timeless (idx_bor_own q i).
Global Instance idx_bor_params : Params (@idx_bor) 5 := {}.
Global Instance bor_params : Params (@bor) 4 := {}.
Global Declare Instance bor_ne κ n : Proper (dist n ==> dist n) (bor κ).
Global Declare Instance bor_contractive κ : Contractive (bor κ).
Global Declare Instance bor_proper κ : Proper (() ==> ()) (bor κ).
Global Declare Instance idx_bor_ne κ i n : Proper (dist n ==> dist n) (idx_bor κ i).
Global Declare Instance idx_bor_contractive κ i : Contractive (idx_bor κ i).
Global Declare Instance idx_bor_proper κ i : Proper (() ==> ()) (idx_bor κ i).
Global Declare Instance lft_tok_fractional κ : Fractional (λ q, q.[κ])%I.
Global Declare Instance lft_tok_as_fractional κ q :
AsFractional q.[κ] (λ q, q.[κ])%I q.
Global Declare Instance frame_lft_tok p κ q1 q2 q :
FrameFractionalQp q1 q2 q
Frame p q1.[κ] q2.[κ] q.[κ] | 5.
Global Declare Instance idx_bor_own_fractional i : Fractional (λ q, idx_bor_own q i)%I.
Global Declare Instance idx_bor_own_as_fractional i q :
AsFractional (idx_bor_own q i) (λ q, idx_bor_own q i)%I q.
Global Declare Instance frame_idx_bor_own p i q1 q2 q :
FrameFractionalQp q1 q2 q
Frame p (idx_bor_own q1 i) (idx_bor_own q2 i) (idx_bor_own q i) | 5.
Global Declare Instance atomic_lft_to_lft_inj : Inj eq eq atomic_lft_to_lft.
Global Declare Instance positive_to_atomic_lft_inj : Inj eq eq positive_to_atomic_lft.
(** * Laws *)
Parameter lft_tok_sep : q κ1 κ2, q.[κ1] q.[κ2] ⊣⊢ q.[κ1 κ2].
Parameter lft_dead_or : κ1 κ2, [κ1] [κ2] ⊣⊢ [ κ1 κ2].
Parameter lft_tok_dead : q κ, q.[κ] -∗ [ κ] -∗ False.
Parameter lft_tok_static : q, q.[static].
Parameter lft_tok_valid : q κ, q.[κ] -∗ (q 1)%Qp κ = static⌝.
Parameter lft_dead_static : [ static] -∗ False.
Parameter lft_intersect_static_cancel_l : κ κ', κ κ' = static κ = static.
(** Create a lifetime in some given set of names [P]. This lemma statement
requires exposing [atomic_lft], because [P] restricted to the image of
[atomic_to_lft] might well not be infinite. *)
Parameter lft_create_strong : P E, pred_infinite P lftN E
lft_ctx ={E}=∗
p : positive, let Λ := positive_to_atomic_lft p in P p (1).[@Λ].
(** Given two lifetimes [κ] and [κ'], find a lower bound on atomic lifetimes [m]
such that [m ⊓ κ'] is syntactically different from [κ].
This is useful in conjunction with [lft_create_strong] to generate
fresh lifetimes when using lifetimes as keys to index into a map. *)
Parameter lft_fresh : κ κ', i : positive,
m : positive, (i < m)%positive (positive_to_lft m) κ' κ.
Parameter lft_kill_atomics : Λs,
lft_ctx -∗ (([ set] Λ Λs, (1).[@Λ]) ={lftN userE}[userE]▷=∗ [ set] ΛΛs, [@Λ]).
Parameter bor_create : E κ P,
lftN E lft_ctx -∗ P ={E}=∗ &{κ} P ([κ] ={E}=∗ P).
Parameter bor_fake : E κ P,
lftN E lft_ctx -∗ [κ] ={E}=∗ &{κ}P.
(* This is in the signature only to share the derived proof between the
model and the outside. *)
Parameter bor_shorten : κ κ' P, κ κ' -∗ &{κ'}P -∗ &{κ}P.
Parameter bor_sep : E κ P Q,
lftN E lft_ctx -∗ &{κ} (P Q) ={E}=∗ &{κ} P &{κ} Q.
Parameter bor_combine : E κ P Q,
lftN E lft_ctx -∗ &{κ} P -∗ &{κ} Q ={E}=∗ &{κ} (P Q).
Parameter bor_unfold_idx : κ P, &{κ}P ⊣⊢ i, &{κ,i}P idx_bor_own 1 i.
Parameter idx_bor_shorten : κ κ' i P, κ κ' -∗ &{κ',i} P -∗ &{κ,i} P.
Parameter idx_bor_iff : κ i P P', (P P') -∗ &{κ,i}P -∗ &{κ,i}P'.
Parameter idx_bor_unnest : E κ κ' i P,
lftN E lft_ctx -∗ &{κ,i} P -∗ &{κ'}(idx_bor_own 1 i) ={E}=∗ &{κ κ'} P.
Parameter idx_bor_acc : E q κ i P, lftN E
lft_ctx -∗ &{κ,i}P -∗ idx_bor_own 1 i -∗ q.[κ] ={E}=∗
P ( P ={E}=∗ idx_bor_own 1 i q.[κ]).
Parameter idx_bor_atomic_acc : E q κ i P, lftN E
lft_ctx -∗ &{κ,i}P -∗ idx_bor_own q i ={E,E∖↑lftN}=∗
( P ( P ={E∖↑lftN,E}=∗ idx_bor_own q i))
([κ] |={E∖↑lftN,E}=> idx_bor_own q i).
(* We have to expose κ' here as without [lftN] in the mask of the Q-to-P view
shift, we cannot turn [†κ'] into [†κ]. *)
Parameter bor_acc_strong : E q κ P, lftN E
lft_ctx -∗ &{κ} P -∗ q.[κ] ={E}=∗ κ', κ κ' P
Q, ( Q -∗ [κ'] ={userE}=∗ P) -∗ Q ={E}=∗ &{κ'} Q q.[κ].
Parameter bor_acc_atomic_strong : E κ P, lftN E
lft_ctx -∗ &{κ} P ={E,E∖↑lftN}=∗
( κ', κ κ' P
Q, ( Q -∗ [κ'] ={userE}=∗ P) -∗ Q ={E∖↑lftN,E}=∗ &{κ'} Q)
([κ] |={E∖↑lftN,E}=> True).
(* Because Coq's module system is horrible, we have to repeat properties of lft_incl here
unless we want to prove them twice (both here and in model.primitive) *)
Parameter lft_intersect_acc : κ κ' q q', q.[κ] -∗ q'.[κ'] -∗
q'', q''.[κ κ'] (q''.[κ κ'] -∗ q.[κ] q'.[κ']).
Parameter lft_intersect_incl_l : κ κ', κ κ' κ.
Parameter lft_intersect_incl_r : κ κ', κ κ' κ'.
Parameter lft_incl_refl : κ, κ κ.
Parameter lft_incl_trans : κ κ' κ'', κ κ' -∗ κ' κ'' -∗ κ κ''.
Parameter lft_incl_glb : κ κ' κ'', κ κ' -∗ κ κ'' -∗ κ κ' κ''.
Parameter lft_intersect_mono : κ1 κ1' κ2 κ2',
κ1 κ1' -∗ κ2 κ2' -∗ κ1 κ2 κ1' κ2'.
Parameter lft_incl_acc : E κ κ' q,
lftN E κ κ' -∗ q.[κ] ={E}=∗ q', q'.[κ'] (q'.[κ'] ={E}=∗ q.[κ]).
Parameter lft_incl_dead : E κ κ', lftN E κ κ' -∗ [κ'] ={E}=∗ [κ].
Parameter lft_incl_intro : κ κ',
(( q, q.[κ] ={lftN}=∗ q', q'.[κ'] (q'.[κ'] ={lftN}=∗ q.[κ]))
([κ'] ={lftN}=∗ [κ])) -∗ κ κ'.
End properties.
Parameter lftΣ : gFunctors.
Global Declare Instance subG_lftGpreS Σ : subG lftΣ Σ lftGpreS Σ.
Parameter lft_init : `{!invGS Σ, !lftGpreS Σ} E userE,
lftN E lftN ## userE
|={E}=> _ : lftGS Σ userE, lft_ctx.
End lifetime_sig.
From iris.algebra Require Import dyn_reservation_map agree.
From iris.proofmode Require Import proofmode.
From lrust.lifetime Require Export lifetime.
From iris.prelude Require Import options.
(** This module provides support for attaching metadata (specifically, a
[gname]) to a lifetime (as is required for types using branding). *)
Class lft_metaG Σ := LftMetaG {
lft_meta_inG :: inG Σ (dyn_reservation_mapR (agreeR gnameO));
}.
Definition lft_metaΣ : gFunctors :=
#[GFunctor (dyn_reservation_mapR (agreeR gnameO))].
Global Instance subG_lft_meta Σ :
subG (lft_metaΣ) Σ lft_metaG Σ.
Proof. solve_inG. Qed.
(** We need some global ghost state, but we do not actually care about the name:
we always use a frame-preserving update starting from ε to obtain the ownership
we need. In other words, we use [own_unit] instead of [own_alloc]. As a result
we can just hard-code an arbitrary name here. *)
Local Definition lft_meta_gname : gname := 42%positive.
Definition lft_meta `{!lftGS Σ userE, lft_metaG Σ} (κ : lft) (γ : gname) : iProp Σ :=
p : positive, κ = positive_to_lft p
own lft_meta_gname (dyn_reservation_map_data p (to_agree γ)).
Section lft_meta.
Context `{!invGS Σ, !lftGS Σ userE, lft_metaG Σ}.
Global Instance lft_meta_timeless κ γ : Timeless (lft_meta κ γ).
Proof. apply _. Qed.
Global Instance lft_meta_persistent κ γ : Persistent (lft_meta κ γ).
Proof. apply _. Qed.
Lemma lft_create_meta {E : coPset} (γ : gname) :
lftN E
lft_ctx ={E}=∗
κ, lft_meta κ γ (1).[κ] ((1).[κ] ={lftN userE}[userE]▷=∗ [κ]).
Proof.
iIntros (HE) "#LFT".
iMod (own_unit (dyn_reservation_mapUR (agreeR gnameO)) lft_meta_gname) as "Hown".
iMod (own_updateP _ _ _ dyn_reservation_map_reserve' with "Hown")
as (? [Etok [Hinf ->]]) "Hown".
iMod (lft_create_strong (. Etok) with "LFT") as (p HEtok) "Hκ"; [done..|].
iExists (positive_to_lft p). iFrame "Hκ".
iDestruct (lft_kill_atomic with "LFT") as "$".
iMod (own_update with "Hown") as "Hown".
{ eapply (dyn_reservation_map_alloc _ p (to_agree γ)); done. }
iModIntro. iExists p. eauto.
Qed.
Lemma lft_meta_agree (κ : lft) (γ1 γ2 : gname) :
lft_meta κ γ1 -∗ lft_meta κ γ2 -∗ γ1 = γ2⌝.
Proof.
iIntros "Hidx1 Hidx2".
iDestruct "Hidx1" as (p1) "(% & Hidx1)". subst κ.
iDestruct "Hidx2" as (p2) "(Hlft & Hidx2)".
iDestruct "Hlft" as %<-%(inj atomic_lft_to_lft)%(inj positive_to_atomic_lft).
iCombine "Hidx1 Hidx2" as "Hidx".
iDestruct (own_valid with "Hidx") as %Hval.
rewrite ->(dyn_reservation_map_data_valid (A:=agreeR gnameO)) in Hval.
apply to_agree_op_inv_L in Hval.
done.
Qed.
End lft_meta.
Global Typeclasses Opaque lft_meta.
From lrust.lifetime Require Export primitive. From lrust.lifetime Require Export primitive.
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.algebra Require Import csum auth frac gmap dec_agree gset. From iris.algebra Require Import csum auth frac gmap agree gset.
From iris.base_logic Require Import big_op.
From iris.base_logic.lib Require Import boxes. From iris.base_logic.lib Require Import boxes.
From iris.prelude Require Import options.
Section accessors. Section accessors.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft. Implicit Types κ : lft.
(* Helper internal lemmas *) (* Helper internal lemmas *)
...@@ -14,57 +14,58 @@ Lemma bor_open_internal E P i Pb q : ...@@ -14,57 +14,58 @@ Lemma bor_open_internal E P i Pb q :
slice borN (i.2) P -∗ lft_bor_alive (i.1) Pb -∗ slice borN (i.2) P -∗ lft_bor_alive (i.1) Pb -∗
idx_bor_own 1%Qp i -∗ (q).[i.1] ={E}=∗ idx_bor_own 1%Qp i -∗ (q).[i.1] ={E}=∗
lft_bor_alive (i.1) Pb lft_bor_alive (i.1) Pb
own_bor (i.1) ( {[i.2 := (1%Qp, DecAgree (Bor_open q))]}) P. own_bor (i.1) ( {[i.2 := (1%Qp, to_agree (Bor_open q))]}) P.
Proof. Proof.
iIntros (?) "Hslice Halive Hbor Htok". unfold lft_bor_alive, idx_bor_own. iIntros (?) "Hslice Halive Hbor Htok". unfold lft_bor_alive, idx_bor_own.
iDestruct "Halive" as (B) "(Hbox & >Hown & HB)". iDestruct "Halive" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor") iDestruct (own_bor_valid_2 with "Hown Hbor")
as %[EQB%to_borUR_included _]%auth_valid_discrete_2. as %[EQB%to_borUR_included _]%auth_both_valid_discrete.
iMod (slice_empty _ _ true with "Hslice Hbox") as "[$ Hbox]". iMod (slice_empty _ _ true with "Hslice Hbox") as "[$ Hbox]".
solve_ndisj. by rewrite lookup_fmap EQB. { solve_ndisj. } { by rewrite lookup_fmap EQB. }
rewrite -(fmap_insert bor_filled _ _ (Bor_open q)). rewrite -(fmap_insert bor_filled _ _ (Bor_open q)).
iMod (own_bor_update_2 with "Hown Hbor") as "Hbor". apply auth_update. iMod (own_bor_update_2 with "Hown Hbor") as "Hbor".
{ eapply singleton_local_update. by rewrite lookup_fmap EQB. { apply auth_update.
by apply (exclusive_local_update _ (1%Qp, DecAgree (Bor_open q))). } eapply singleton_local_update.
- by rewrite lookup_fmap EQB.
- by apply (exclusive_local_update _ (1%Qp, to_agree (Bor_open q))). }
rewrite own_bor_op -fmap_insert. iDestruct "Hbor" as "[Hown $]". rewrite own_bor_op -fmap_insert. iDestruct "Hbor" as "[Hown $]".
iExists _. iFrame "∗". iExists _. iFrame "∗".
rewrite -insert_delete big_sepM_insert ?lookup_delete // big_sepM_delete /=; last done. rewrite -insert_delete_insert big_sepM_insert ?lookup_delete // big_sepM_delete /=; last done.
iDestruct "HB" as "[_ $]". auto. iDestruct "HB" as "[_ $]". auto.
Qed. Qed.
Lemma bor_close_internal E P i Pb q : Lemma bor_close_internal E P i Pb q :
borN E borN E
slice borN (i.2) P -∗ lft_bor_alive (i.1) Pb -∗ slice borN (i.2) P -∗ lft_bor_alive (i.1) Pb -∗
own_bor (i.1) ( {[i.2 := (1%Qp, DecAgree (Bor_open q))]}) -∗ P ={E}=∗ own_bor (i.1) ( {[i.2 := (1%Qp, to_agree (Bor_open q))]}) -∗ P ={E}=∗
lft_bor_alive (i.1) Pb idx_bor_own 1%Qp i (q).[i.1]. lft_bor_alive (i.1) Pb idx_bor_own 1%Qp i (q).[i.1].
Proof. Proof.
iIntros (?) "Hslice Halive Hbor HP". unfold lft_bor_alive, idx_bor_own. iIntros (?) "Hslice Halive Hbor HP". unfold lft_bor_alive, idx_bor_own.
iDestruct "Halive" as (B) "(Hbox & >Hown & HB)". iDestruct "Halive" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor") iDestruct (own_bor_valid_2 with "Hown Hbor")
as %[EQB%to_borUR_included _]%auth_valid_discrete_2. as %[EQB%to_borUR_included _]%auth_both_valid_discrete.
iMod (slice_fill _ _ true with "Hslice HP Hbox") as "Hbox". iMod (slice_fill _ _ true with "Hslice HP Hbox") as "Hbox".
solve_ndisj. by rewrite lookup_fmap EQB. { solve_ndisj. } { by rewrite lookup_fmap EQB. }
rewrite -(fmap_insert bor_filled _ _ Bor_in). rewrite -(fmap_insert bor_filled _ _ Bor_in).
iMod (own_bor_update_2 with "Hown Hbor") as "Hbor". apply auth_update. iMod (own_bor_update_2 with "Hown Hbor") as "Hbor".
{ eapply singleton_local_update. by rewrite lookup_fmap EQB. { apply auth_update.
by apply (exclusive_local_update _ (1%Qp, DecAgree Bor_in)). } eapply singleton_local_update.
- by rewrite lookup_fmap EQB.
- by apply (exclusive_local_update _ (1%Qp, to_agree Bor_in)). }
rewrite own_bor_op -fmap_insert. iDestruct "Hbor" as "[Hown $]". rewrite own_bor_op -fmap_insert. iDestruct "Hbor" as "[Hown $]".
rewrite big_sepM_delete //. simpl. iDestruct "HB" as "[>$ HB]". rewrite big_sepM_delete //. simpl. iDestruct "HB" as "[>$ HB]".
iExists _. iFrame "Hbox Hown". iExists _. iFrame "Hbox Hown".
rewrite -insert_delete big_sepM_insert ?lookup_delete //. simpl. by iFrame. rewrite -insert_delete_insert big_sepM_insert ?lookup_delete //. simpl. by iFrame.
Qed. Qed.
Lemma add_vs Pb Pb' P Q Pi κ : Lemma add_vs Pb Pb' P Q Pi κ :
(Pb (P Pb')) -∗ lft_vs κ Pb Pi -∗ ( Q -∗ [κ] ={⊤∖↑lftN}=∗ P) -∗ (Pb (P Pb')) -∗ lft_vs κ Pb Pi -∗ ( Q -∗ [κ] ={userE}=∗ P) -∗
lft_vs κ (Q Pb') Pi. lft_vs κ (Q Pb') Pi.
Proof. Proof.
iIntros "#HEQ Hvs HvsQ". iNext. rewrite !lft_vs_unfold. iIntros "#HEQ Hvs HvsQ". iApply (lft_vs_cons with "[-Hvs] Hvs").
iDestruct "Hvs" as (n) "[Hcnt Hvs]". iExists n. iIntros "[HQ HPb'] #H†".
iIntros "{$Hcnt}*Hinv[HQ HPb] #H†". iApply fupd_trans. iApply fupd_mask_mono; last iMod ("HvsQ" with "HQ H†") as "HP"; first set_solver.
iApply fupd_mask_mono; last iMod ("HvsQ" with "HQ H†") as "HP". solve_ndisj. iModIntro. iNext. iRewrite "HEQ". iFrame.
iModIntro. iAssert ( Pb)%I with "[HPb HP]" as "HPb".
{ iNext. iRewrite "HEQ". iFrame. }
iApply ("Hvs" with "* Hinv HPb H†").
Qed. Qed.
(** Indexed borrow *) (** Indexed borrow *)
...@@ -74,19 +75,22 @@ Lemma idx_bor_acc E q κ i P : ...@@ -74,19 +75,22 @@ Lemma idx_bor_acc E q κ i P :
P ( P ={E}=∗ idx_bor_own 1 i q.[κ]). P ( P ={E}=∗ idx_bor_own 1 i q.[κ]).
Proof. Proof.
unfold idx_bor. iIntros (HE) "#LFT [#Hle #Hs] Hbor Htok". unfold idx_bor. iIntros (HE) "#LFT [#Hle #Hs] Hbor Htok".
iMod (lft_incl_acc with "Hle Htok") as (q') "[Htok Hclose]". solve_ndisj. iDestruct "Hs" as (P') "[#HPP' Hs]".
iMod (lft_incl_acc with "Hle Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'".
iDestruct (own_bor_auth with "HI [Hbor]") as %Hκ'. by unfold idx_bor_own. iDestruct (own_bor_auth with "HI [Hbor]") as %Hκ'.
{ by unfold idx_bor_own. }
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
/lfts_inv /lft_inv /lft_inv_dead /lft_alive_in lft_inv_alive_unfold. /lfts_inv /lft_inv /lft_inv_dead /lft_alive_in lft_inv_alive_unfold.
iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >%]] Hclose'']". iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >%]] Hclose'']".
- iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)". - iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
iMod (bor_open_internal with "Hs Halive Hbor Htok") as "(Halive & Hf & $)". iMod (bor_open_internal with "Hs Halive Hbor Htok") as "(Halive & Hf & HP')".
solve_ndisj. { solve_ndisj. }
iDestruct ("HPP'" with "HP'") as "$".
iMod ("Hclose'" with "[-Hf Hclose]") as "_". iMod ("Hclose'" with "[-Hf Hclose]") as "_".
{ iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". { iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''".
iLeft. iFrame "%". iExists _, _. iFrame. } iLeft. iFrame "%". iExists _, _. iFrame. }
iIntros "!>HP". clear -HE. iIntros "!>HP'". iDestruct ("HPP'" with "HP'") as "HP". clear -HE.
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'".
iDestruct (own_bor_auth with "HI Hf") as %Hκ'. iDestruct (own_bor_auth with "HI Hf") as %Hκ'.
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
...@@ -94,17 +98,17 @@ Proof. ...@@ -94,17 +98,17 @@ Proof.
iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >Hdead_in]] Hclose'']". iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >Hdead_in]] Hclose'']".
+ rewrite lft_inv_alive_unfold. + rewrite lft_inv_alive_unfold.
iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)". iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
iMod (bor_close_internal with "Hs Halive Hf HP") as "(Halive & $ & Htok)". iMod (bor_close_internal with "Hs Halive Hf HP") as "(Halive & $ & Htok)";
solve_ndisj. first solve_ndisj.
iMod ("Hclose'" with "[-Htok Hclose]") as "_"; last by iApply "Hclose". iMod ("Hclose'" with "[-Htok Hclose]") as "_"; last by iApply "Hclose".
iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''".
iLeft. iFrame "%". iExists _, _. iFrame. iLeft. iFrame "%". iExists _, _. iFrame.
+ iDestruct "Hinv" as (?) "[Hinv _]". iDestruct "Hinv" as (B ?) "[>Hinv _]". + iDestruct "Hinv" as (?) "[Hinv _]". iDestruct "Hinv" as (B ?) "[>Hinv _]".
iDestruct (own_bor_valid_2 with "Hinv Hf") iDestruct (own_bor_valid_2 with "Hinv Hf")
as %[(_ & <- & INCL%option_included)%singleton_included _]%auth_valid_discrete_2. as %[(_ & <- & INCL%option_included)%singleton_included_l _]%auth_both_valid_discrete.
by destruct INCL as [[=]|(? & ? & [=<-] & by destruct INCL as [[=]|(? & ? & [=<-] &
[[_<-]%lookup_to_gmap_Some [[_[=]]|[]%(exclusive_included _)]])]. [[_<-]%lookup_gset_to_gmap_Some [[_ ?%(inj to_agree)]|[]%(exclusive_included _)]])].
- iMod (lft_dead_in_tok with "HA") as "[_ H†]". done. - iMod (lft_dead_in_tok with "HA") as "[_ H†]"; first done.
iDestruct (lft_tok_dead with "Htok H†") as "[]". iDestruct (lft_tok_dead with "Htok H†") as "[]".
Qed. Qed.
...@@ -115,6 +119,7 @@ Lemma idx_bor_atomic_acc E q κ i P : ...@@ -115,6 +119,7 @@ Lemma idx_bor_atomic_acc E q κ i P :
([κ] |={E∖↑lftN,E}=> idx_bor_own q i). ([κ] |={E∖↑lftN,E}=> idx_bor_own q i).
Proof. Proof.
unfold idx_bor, idx_bor_own. iIntros (HE) "#LFT [#Hle #Hs] Hbor". unfold idx_bor, idx_bor_own. iIntros (HE) "#LFT [#Hle #Hs] Hbor".
iDestruct "Hs" as (P') "[#HPP' Hs]".
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'".
iDestruct (own_bor_auth with "HI Hbor") as %Hκ'. iDestruct (own_bor_auth with "HI Hbor") as %Hκ'.
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
...@@ -124,21 +129,25 @@ Proof. ...@@ -124,21 +129,25 @@ Proof.
unfold lft_bor_alive. unfold lft_bor_alive.
iDestruct "Halive" as (B) "(Hbox & >Hown & HB)". iDestruct "Halive" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor") iDestruct (own_bor_valid_2 with "Hown Hbor")
as %[EQB%to_borUR_included _]%auth_valid_discrete_2. as %[EQB%to_borUR_included _]%auth_both_valid_discrete.
iMod (slice_empty _ _ true with "Hs Hbox") as "[$ Hbox]". iMod (slice_empty _ _ true with "Hs Hbox") as "[HP' Hbox]";
solve_ndisj. by rewrite lookup_fmap EQB. first solve_ndisj.
iMod fupd_intro_mask' as "Hclose"; last iIntros "!>HP". solve_ndisj. { by rewrite lookup_fmap EQB. }
iMod "Hclose" as "_". iMod (slice_fill _ _ true with "Hs HP Hbox") as "Hbox". iDestruct ("HPP'" with "HP'") as "$".
solve_ndisj. by rewrite lookup_insert. iFrame. iMod fupd_mask_subseteq as "Hclose"; last iIntros "!>HP'"; first solve_ndisj.
iApply "Hclose'". iExists _, _. iFrame. rewrite big_sepS_later. iDestruct ("HPP'" with "HP'") as "HP".
iApply "Hclose''". iLeft. iFrame "%". iExists _, _. iFrame. iMod "Hclose" as "_". iMod (slice_fill _ _ true with "Hs HP Hbox") as "Hbox";
iExists _. iFrame. first solve_ndisj.
{ by rewrite lookup_insert. }
iFrame.
iApply "Hclose'". iFrame. rewrite big_sepS_later.
iApply "Hclose''". iLeft. iFrame "% ∗".
rewrite insert_insert -(fmap_insert _ _ _ Bor_in) insert_id //. rewrite insert_insert -(fmap_insert _ _ _ Bor_in) insert_id //.
- iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]". done. - iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]"; first done.
iMod ("Hclose'" with "[-Hbor]") as "_". iMod ("Hclose'" with "[-Hbor]") as "_".
+ iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto. + iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto.
+ iMod (lft_incl_dead with "Hle H†"). done. iFrame. + iMod (lft_incl_dead with "Hle H†"); first done. iFrame.
iApply fupd_intro_mask'. solve_ndisj. iApply fupd_mask_subseteq. solve_ndisj.
Qed. Qed.
(** Basic borrows *) (** Basic borrows *)
...@@ -146,23 +155,26 @@ Qed. ...@@ -146,23 +155,26 @@ Qed.
Lemma bor_acc_strong E q κ P : Lemma bor_acc_strong E q κ P :
lftN E lftN E
lft_ctx -∗ &{κ} P -∗ q.[κ] ={E}=∗ κ', κ κ' P lft_ctx -∗ &{κ} P -∗ q.[κ] ={E}=∗ κ', κ κ' P
Q, Q -∗ ( Q -∗ [κ'] ={⊤∖↑lftN}=∗ P) ={E}=∗ &{κ'} Q q.[κ]. Q, ( Q -∗ [κ'] ={userE}=∗ P) -∗ Q ={E}=∗ &{κ'} Q q.[κ].
Proof. Proof.
unfold bor, raw_bor. iIntros (HE) "#LFT Hbor Htok". unfold bor, raw_bor. iIntros (HE) "#LFT Hbor Htok".
iDestruct "Hbor" as (κ'') "(#Hle & Htmp)". iDestruct "Htmp" as (s'') "(Hbor & #Hs)". iDestruct "Hbor" as (κ'') "(#Hle & Htmp)". iDestruct "Htmp" as (s'') "(Hbor & #Hs)".
iMod (lft_incl_acc with "Hle Htok") as (q') "[Htok Hclose]". solve_ndisj. iDestruct "Hs" as (P') "[#HPP' Hs]".
iMod (lft_incl_acc with "Hle Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'".
iDestruct (own_bor_auth with "HI [Hbor]") as %Hκ'. by unfold idx_bor_own. iDestruct (own_bor_auth with "HI [Hbor]") as %Hκ'.
{ by unfold idx_bor_own. }
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
/lfts_inv /lft_inv /lft_inv_dead /lft_alive_in lft_inv_alive_unfold. /lfts_inv /lft_inv /lft_inv_dead /lft_alive_in lft_inv_alive_unfold.
iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >%]] Hclose'']". iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >%]] Hclose'']".
- iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)". - iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
iMod (bor_open_internal _ _ (κ'', s'') with "Hs Halive Hbor Htok") as "(Halive & Hbor & $)". iMod (bor_open_internal _ _ (κ'', s'') with "Hs Halive Hbor Htok")
solve_ndisj. as "(Halive & Hbor & HP')"; first solve_ndisj.
iDestruct ("HPP'" with "HP'") as "$".
iMod ("Hclose'" with "[-Hbor Hclose]") as "_". iMod ("Hclose'" with "[-Hbor Hclose]") as "_".
{ iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". { iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''".
iLeft. iFrame "%". iExists _, _. iFrame. } iLeft. iFrame "%". iExists _, _. iFrame. }
iExists κ''. iFrame "#". iIntros "!>*HQ HvsQ". clear -HE. iExists κ''. iFrame "Hle". iIntros "!> %Q HvsQ HQ". clear -HE.
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'".
iDestruct (own_bor_auth with "HI Hbor") as %Hκ'. iDestruct (own_bor_auth with "HI Hbor") as %Hκ'.
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
...@@ -172,36 +184,40 @@ Proof. ...@@ -172,36 +184,40 @@ Proof.
iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)". iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
iDestruct "Halive" as (B) "(Hbox & >Hown & HB)". iDestruct "Halive" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor") iDestruct (own_bor_valid_2 with "Hown Hbor")
as %[EQB%to_borUR_included _]%auth_valid_discrete_2. as %[EQB%to_borUR_included _]%auth_both_valid_discrete.
iMod (slice_delete_empty _ _ true with "Hs Hbox") as (Pb') "[EQ Hbox]". iMod (slice_delete_empty _ _ true with "Hs Hbox") as (Pb') "[EQ Hbox]";
solve_ndisj. by rewrite lookup_fmap EQB. first solve_ndisj.
iDestruct (add_vs with "EQ Hvs HvsQ") as "Hvs". { by rewrite lookup_fmap EQB. }
iMod (slice_insert_empty _ _ true with "Hbox") as (j) "(% & #Hs' & Hbox)". iDestruct (add_vs with "EQ Hvs [HvsQ]") as "Hvs".
{ iNext. iIntros "HQ H†". iApply "HPP'". iApply ("HvsQ" with "HQ H†"). }
iMod (slice_insert_empty _ _ true _ Q with "Hbox") as (j) "(% & #Hs' & Hbox)".
iMod (own_bor_update_2 with "Hown Hbor") as "Hown". iMod (own_bor_update_2 with "Hown Hbor") as "Hown".
{ apply auth_update. etrans. { apply auth_update. etrans.
- apply delete_singleton_local_update, _. - apply delete_singleton_local_update, _.
- apply (alloc_singleton_local_update _ j - apply (alloc_singleton_local_update _ j
(1%Qp, DecAgree (Bor_open q'))); last done. (1%Qp, to_agree (Bor_open q'))); last done.
rewrite -fmap_delete lookup_fmap fmap_None rewrite -fmap_delete lookup_fmap fmap_None
-fmap_None -lookup_fmap fmap_delete //. } -fmap_None -lookup_fmap fmap_delete //. }
rewrite own_bor_op. iDestruct "Hown" as "[Hown Hbor]". rewrite own_bor_op. iDestruct "Hown" as "[Hown Hbor]".
iMod (bor_close_internal _ _ (_, _) with "Hs' [Hbox Hown HB] Hbor HQ") iMod (bor_close_internal _ _ (_, _) with "Hs' [Hbox Hown HB] Hbor HQ")
as "(Halive & Hbor & Htok)". solve_ndisj. as "(Halive & Hbor & Htok)"; first solve_ndisj.
{ rewrite -!fmap_delete -fmap_insert -(fmap_insert _ _ _ (Bor_open q')) { rewrite -!fmap_delete -fmap_insert -(fmap_insert _ _ _ (Bor_open q'))
/lft_bor_alive. iExists _. iFrame "Hbox Hown". /lft_bor_alive. iExists _. iFrame "Hbox Hown".
rewrite big_sepM_insert. by rewrite big_sepM_delete. rewrite big_sepM_insert.
by rewrite -fmap_None -lookup_fmap fmap_delete. } - by rewrite big_sepM_delete.
- by rewrite -fmap_None -lookup_fmap fmap_delete. }
iMod ("Hclose'" with "[-Hbor Htok Hclose]") as "_". iMod ("Hclose'" with "[-Hbor Htok Hclose]") as "_".
{ iExists _, _. iFrame. rewrite big_sepS_later /lft_bor_alive. iApply "Hclose''". { iFrame. rewrite big_sepS_later /lft_bor_alive. iApply "Hclose''".
iLeft. iFrame "%". iExists _, _. iFrame. } iLeft. iFrame "% ∗". }
iMod ("Hclose" with "Htok") as "$". iExists κ''. iModIntro. iMod ("Hclose" with "Htok") as "$". iExists κ''. iModIntro.
iSplit; first by iApply lft_incl_refl. iExists j. iFrame "∗#". iSplit; first by iApply lft_incl_refl. iFrame.
iExists Q. rewrite -(bi.iff_refl emp). eauto.
+ iDestruct "Hinv" as (?) "[Hinv _]". iDestruct "Hinv" as (B ?) "[>Hinv _]". + iDestruct "Hinv" as (?) "[Hinv _]". iDestruct "Hinv" as (B ?) "[>Hinv _]".
iDestruct (own_bor_valid_2 with "Hinv Hbor") iDestruct (own_bor_valid_2 with "Hinv Hbor")
as %[(_ & <- & INCL%option_included)%singleton_included _]%auth_valid_discrete_2. as %[(_ & <- & INCL%option_included)%singleton_included_l _]%auth_both_valid_discrete.
by destruct INCL as [[=]|(? & ? & [=<-] & by destruct INCL as [[=]|(? & ? & [=<-] &
[[_<-]%lookup_to_gmap_Some [[_[=]]|[]%(exclusive_included _)]])]. [[_<-]%lookup_gset_to_gmap_Some [[_?%(inj to_agree)]|[]%(exclusive_included _)]])].
- iMod (lft_dead_in_tok with "HA") as "[_ H†]". done. - iMod (lft_dead_in_tok with "HA") as "[_ H†]"; first done.
iDestruct (lft_tok_dead with "Htok H†") as "[]". iDestruct (lft_tok_dead with "Htok H†") as "[]".
Qed. Qed.
...@@ -209,95 +225,54 @@ Lemma bor_acc_atomic_strong E κ P : ...@@ -209,95 +225,54 @@ Lemma bor_acc_atomic_strong E κ P :
lftN E lftN E
lft_ctx -∗ &{κ} P ={E,E∖↑lftN}=∗ lft_ctx -∗ &{κ} P ={E,E∖↑lftN}=∗
( κ', κ κ' P ( κ', κ κ' P
Q, Q -∗ ( Q -∗ [κ'] ={⊤∖↑lftN}=∗ P) ={E∖↑lftN,E}=∗ &{κ'} Q) Q, ( Q -∗ [κ'] ={userE}=∗ P) -∗ Q ={E∖↑lftN,E}=∗ &{κ'} Q)
([κ] |={E∖↑lftN,E}=> True). ([κ] |={E∖↑lftN,E}=> True).
Proof. Proof.
iIntros (HE) "#LFT Hbor". rewrite bor_unfold_idx /idx_bor /bor /raw_bor. iIntros (HE) "#LFT Hbor". rewrite bor_unfold_idx /idx_bor /bor /raw_bor.
iDestruct "Hbor" as (i) "((#Hle & #Hs) & Hbor)". iDestruct "Hbor" as (i) "((#Hle & #Hs) & Hbor)".
iDestruct "Hs" as (P') "[#HPP' Hs]".
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose'".
iDestruct (own_bor_auth with "HI [Hbor]") as %Hκ'. by unfold idx_bor_own. iDestruct (own_bor_auth with "HI [Hbor]") as %Hκ'.
{ by unfold idx_bor_own. }
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
/lfts_inv /lft_inv /lft_inv_dead /lft_alive_in lft_inv_alive_unfold. /lfts_inv /lft_inv /lft_inv_dead /lft_alive_in lft_inv_alive_unfold.
iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >%]] Hclose'']". iDestruct "Hinv" as "[[[Hinv >%]|[Hinv >%]] Hclose'']".
- iLeft. iExists (i.1). iFrame "#". - iLeft. iFrame "Hle".
iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)". iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
unfold lft_bor_alive, idx_bor_own. unfold lft_bor_alive, idx_bor_own.
iDestruct "Halive" as (B) "(Hbox & >Hown & HB)". iDestruct "Halive" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor") iDestruct (own_bor_valid_2 with "Hown Hbor")
as %[EQB%to_borUR_included _]%auth_valid_discrete_2. as %[EQB%to_borUR_included _]%auth_both_valid_discrete.
iMod (slice_delete_full _ _ true with "Hs Hbox") as (Pb') "($ & EQ & Hbox)". iMod (slice_delete_full _ _ true with "Hs Hbox") as (Pb') "(HP' & EQ & Hbox)";
solve_ndisj. by rewrite lookup_fmap EQB. first solve_ndisj.
iMod fupd_intro_mask' as "Hclose"; last iIntros "!>*HQ HvsQ". solve_ndisj. { by rewrite lookup_fmap EQB. }
iMod "Hclose" as "_". iDestruct (add_vs with "EQ Hvs HvsQ") as "Hvs". iDestruct ("HPP'" with "HP'") as "$".
iMod (slice_insert_full _ _ true with "HQ Hbox") as (j) "(% & #Hs' & Hbox)". iMod fupd_mask_subseteq as "Hclose"; last iIntros "!> %Q HvsQ HQ"; first solve_ndisj.
solve_ndisj. iMod "Hclose" as "_". iDestruct (add_vs with "EQ Hvs [HvsQ]") as "Hvs".
{ iNext. iIntros "HQ H†". iApply "HPP'". iApply ("HvsQ" with "HQ H†"). }
iMod (slice_insert_full _ _ true with "HQ Hbox") as (j) "(% & #Hs' & Hbox)";
first solve_ndisj.
iMod (own_bor_update_2 with "Hown Hbor") as "Hown". iMod (own_bor_update_2 with "Hown Hbor") as "Hown".
{ apply auth_update. etrans. { apply auth_update. etrans.
- apply delete_singleton_local_update, _. - apply delete_singleton_local_update, _.
- apply (alloc_singleton_local_update _ j (1%Qp, DecAgree (Bor_in))); last done. - apply (alloc_singleton_local_update _ j (1%Qp, to_agree (Bor_in))); last done.
rewrite -fmap_delete lookup_fmap fmap_None rewrite -fmap_delete lookup_fmap fmap_None
-fmap_None -lookup_fmap fmap_delete //. } -fmap_None -lookup_fmap fmap_delete //. }
rewrite own_bor_op. iDestruct "Hown" as "[H● H◯]". rewrite own_bor_op. iDestruct "Hown" as "[H● H◯]".
iMod ("Hclose'" with "[- H◯]"); last first. iMod ("Hclose'" with "[- H◯]"); last first.
{ iModIntro. iExists (i.1). iSplit; first by iApply lft_incl_refl. { iModIntro. iExists (i.1). iSplit; first by iApply lft_incl_refl.
iExists j. iFrame. done. } iFrame "Hs' ∗". rewrite -(bi.iff_refl emp). auto. }
iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". iFrame. rewrite big_sepS_later. iApply "Hclose''".
iLeft. iFrame "%". iExists _, _. iFrame. iNext. iLeft. iFrame "%". iExists _, _.
rewrite -!fmap_delete -fmap_insert -(fmap_insert _ _ _ Bor_in). rewrite -!fmap_delete -fmap_insert -(fmap_insert _ _ _ Bor_in).
iExists _. iFrame "Hbox H●". iFrame.
rewrite big_sepM_insert. by rewrite big_sepM_delete. rewrite big_sepM_insert; last first.
by rewrite -fmap_None -lookup_fmap fmap_delete. { by rewrite -fmap_None -lookup_fmap fmap_delete. }
- iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]". done. by rewrite big_sepM_delete.
- iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]"; first done.
iMod ("Hclose'" with "[-Hbor]") as "_". iMod ("Hclose'" with "[-Hbor]") as "_".
+ iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto. + iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto.
+ iMod (lft_incl_dead with "Hle H†") as "$". done. + iMod (lft_incl_dead with "Hle H†") as "$"; first done.
iApply fupd_intro_mask'. solve_ndisj. iApply fupd_mask_subseteq; first solve_ndisj.
Qed.
Lemma bor_acc_cons E q κ P :
lftN E
lft_ctx -∗ &{κ} P -∗ q.[κ] ={E}=∗ P
Q, Q -∗ ( Q ={⊤∖↑lftN}=∗ P) ={E}=∗ &{κ} Q q.[κ].
Proof.
iIntros (?) "#LFT HP Htok".
iMod (bor_acc_strong with "LFT HP Htok") as (κ') "(#Hκκ' & $ & Hclose)"; first done.
iIntros "!>*HQ HPQ". iMod ("Hclose" with "*HQ [HPQ]") as "[Hb $]".
- iNext. iIntros "? _". by iApply "HPQ".
- iApply (bor_shorten with "Hκκ' Hb").
Qed.
Lemma bor_acc_atomic_cons E κ P :
lftN E
lft_ctx -∗ &{κ} P ={E,E∖↑lftN}=∗
( P Q, Q -∗ ( Q ={⊤∖↑lftN}=∗ P) ={E∖↑lftN,E}=∗ &{κ} Q)
([κ] |={E∖↑lftN,E}=> True).
Proof.
iIntros (?) "#LFT HP".
iMod (bor_acc_atomic_strong with "LFT HP") as "[H|[??]]"; first done.
- iLeft. iDestruct "H" as (κ') "(#Hκκ' & $ & Hclose)". iIntros "!>*HQ HPQ".
iMod ("Hclose" with "* HQ [HPQ]") as "Hb".
+ iNext. iIntros "? _". by iApply "HPQ".
+ iApply (bor_shorten with "Hκκ' Hb").
- iRight. by iFrame.
Qed.
Lemma bor_acc E q κ P :
lftN E
lft_ctx -∗ &{κ}P -∗ q.[κ] ={E}=∗ P ( P ={E}=∗ &{κ}P q.[κ]).
Proof.
iIntros (?) "#LFT HP Htok".
iMod (bor_acc_cons with "LFT HP Htok") as "($ & Hclose)"; first done.
iIntros "!>HP". iMod ("Hclose" with "*HP []") as "[$ $]"; auto.
Qed.
Lemma bor_acc_atomic E κ P :
lftN E
lft_ctx -∗ &{κ}P ={E,E∖↑lftN}=∗
( P ( P ={E∖↑lftN,E}=∗ &{κ}P)) ([κ] |={E∖↑lftN,E}=> True).
Proof.
iIntros (?) "#LFT HP".
iMod (bor_acc_atomic_cons with "LFT HP") as "[[HP Hclose]|[? ?]]"; first done.
- iLeft. iIntros "!> {$HP} HP". iMod ("Hclose" with "* HP []"); auto.
- iRight. by iFrame.
Qed. Qed.
End accessors. End accessors.
From lrust.lifetime Require Export primitive.
From lrust.lifetime Require Export faking.
From iris.algebra Require Import csum auth frac gmap agree gset.
From iris.base_logic.lib Require Import boxes.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Section borrow.
Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft.
Lemma raw_bor_create E κ P :
lftN E
lft_ctx -∗ P ={E}=∗ raw_bor κ P ([κ] ={E}=∗ P).
Proof.
iIntros (HE) "#LFT HP". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
iMod (ilft_create _ _ κ with "HA HI Hinv") as (A' I') "(Hκ & HA & HI & Hinv)".
iDestruct "Hκ" as %. iDestruct (@big_sepS_later with "Hinv") as "Hinv".
iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hinv Hclose']".
{ by apply elem_of_dom. }
rewrite {1}/lft_inv. iDestruct "Hinv" as "[[Hinv >%]|[Hinv >%]]".
- rewrite {1}lft_inv_alive_unfold;
iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
rewrite /lft_bor_alive; iDestruct "Halive" as (B) "(HboxB & >HownB & HB)".
iMod (lft_inh_extend _ _ P with "Hinh")
as "(Hinh & HIlookup & Hinh_close)"; first solve_ndisj.
iMod (slice_insert_full _ _ true with "HP HboxB")
as (γB) "(HBlookup & HsliceB & HboxB)"; first by solve_ndisj.
rewrite lookup_fmap. iDestruct "HBlookup" as %HBlookup.
rewrite -(fmap_insert bor_filled _ _ Bor_in).
iMod (own_bor_update with "HownB") as "[HB● HB◯]".
{ eapply auth_update_alloc,
(alloc_singleton_local_update _ γB (1%Qp, to_agree Bor_in)); last done.
rewrite lookup_fmap. case:(B !! γB) HBlookup; done. }
rewrite -fmap_insert.
iSpecialize ("Hclose'" with "[Hvs Hinh HboxB HB● HB]").
{ iNext. rewrite /lft_inv. iLeft. iFrame "%".
rewrite lft_inv_alive_unfold. iExists (P Pb)%I, (P Pi)%I.
iFrame "Hinh". iSplitL "HboxB HB● HB"; last by iApply lft_vs_frame.
rewrite /lft_bor_alive. iExists _. iFrame "HboxB HB●".
iApply @big_sepM_insert; first by destruct (B !! γB).
simpl. iFrame. }
iMod ("Hclose" with "[HA HI Hclose']") as "_"; [by iNext; iExists _, _; iFrame|].
iSplitL "HB◯ HsliceB".
+ rewrite /bor /raw_bor /idx_bor_own. iModIntro. iFrame.
rewrite -(bi.iff_refl emp). auto.
+ clear -HE. iIntros "!> H†".
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
iDestruct ("HIlookup" with "HI") as %.
iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hinv Hclose']".
{ by apply elem_of_dom. }
rewrite /lft_dead; iDestruct "H†" as (Λ) "[% #H†]".
iDestruct (own_alft_auth_agree A Λ false with "HA H†") as %EQAΛ.
rewrite {1}/lft_inv; iDestruct "Hinv" as "[[_ >%]|[Hinv >%]]".
{ unfold lft_alive_in in *. naive_solver. }
rewrite /lft_inv_dead; iDestruct "Hinv" as (Pinh) "(Hdead & >Hcnt & Hinh)".
iMod ("Hinh_close" $! Pinh with "Hinh") as (Pinh') "(? & $ & ?)".
iApply "Hclose". iExists A, I. iFrame. iNext. iApply "Hclose'".
rewrite /lft_inv. iRight. iFrame "%".
rewrite /lft_inv_dead. iExists Pinh'. iFrame.
- iFrame "HP". iApply fupd_frame_r. iSplitR ""; last by auto.
rewrite /lft_inv_dead. iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)" .
iMod (raw_bor_fake with "Hdead") as "[Hdead Hbor]"; first solve_ndisj.
unfold bor. iFrame. iApply "Hclose". iExists _, _. iFrame. rewrite big_sepS_later.
iApply "Hclose'". iNext. rewrite /lft_inv. iRight.
rewrite /lft_inv_dead. iFrame. eauto.
Qed.
Lemma bor_create E κ P :
lftN E
lft_ctx -∗ P ={E}=∗ &{κ}P ([κ] ={E}=∗ P).
Proof.
iIntros (?) "#LFT HP". iMod (raw_bor_create with "LFT HP") as "[HP $]"; [done|].
rewrite /bor. iExists _. iFrame. iApply lft_incl_refl.
Qed.
End borrow.
From lrust.lifetime Require Export primitive. From lrust.lifetime Require Export primitive.
From lrust.lifetime Require Export faking raw_reborrow. From lrust.lifetime Require Export faking reborrow.
From iris.algebra Require Import csum auth frac gmap dec_agree gset. From iris.algebra Require Import csum auth frac gmap agree.
From iris.base_logic Require Import big_op.
From iris.base_logic.lib Require Import boxes. From iris.base_logic.lib Require Import boxes.
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Section borrow. Section borrow.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft. Implicit Types κ : lft.
Lemma bor_create E κ P :
lftN E
lft_ctx -∗ P ={E}=∗ &{κ} P ([κ] ={E}=∗ P).
Proof.
iIntros (HE) "#LFT HP". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
iMod (ilft_create _ _ κ with "HA HI Hinv") as (A' I') "(Hκ & HA & HI & Hinv)".
iDestruct "Hκ" as %. iDestruct (@big_sepS_later with "Hinv") as "Hinv".
iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hinv Hclose']".
{ by apply elem_of_dom. }
rewrite {1}/lft_inv. iDestruct "Hinv" as "[[Hinv >%]|[Hinv >%]]".
- rewrite {1}lft_inv_alive_unfold;
iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)".
rewrite /lft_bor_alive; iDestruct "Halive" as (B) "(HboxB & >HownB & HB)".
iMod (lft_inh_extend _ _ P with "Hinh")
as "(Hinh & HIlookup & Hinh_close)"; first solve_ndisj.
iMod (slice_insert_full _ _ true with "HP HboxB")
as (γB) "(HBlookup & HsliceB & HboxB)"; first by solve_ndisj.
rewrite lookup_fmap. iDestruct "HBlookup" as %HBlookup.
rewrite -(fmap_insert bor_filled _ _ Bor_in).
iMod (own_bor_update with "HownB") as "[HB● HB◯]".
{ eapply auth_update_alloc,
(alloc_singleton_local_update _ γB (1%Qp, DecAgree Bor_in)); last done.
rewrite lookup_fmap. by destruct (B !! γB). }
rewrite -fmap_insert.
iSpecialize ("Hclose'" with "[Hvs Hinh HboxB HB● HB]").
{ iNext. rewrite /lft_inv. iLeft. iFrame "%".
rewrite lft_inv_alive_unfold. iExists (P Pb)%I, (P Pi)%I.
iFrame "Hinh". iSplitL "HboxB HB● HB"; last by iApply lft_vs_frame.
rewrite /lft_bor_alive. iExists _. iFrame "HboxB HB●".
iApply @big_sepM_insert; first by destruct (B !! γB).
simpl. iFrame. }
iMod ("Hclose" with "[HA HI Hclose']") as "_"; [by iExists _, _; iFrame|].
iSplitL "HB◯ HsliceB".
+ rewrite /bor /raw_bor /idx_bor_own. iExists κ; simpl.
iModIntro. iSplit; first by iApply lft_incl_refl. iExists γB. by iFrame.
+ clear -HE. iIntros "!> H†".
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
iDestruct ("HIlookup" with "* HI") as %.
iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hinv Hclose']".
{ by apply elem_of_dom. }
rewrite /lft_dead; iDestruct "H†" as (Λ) "[% #H†]".
iDestruct (own_alft_auth_agree A Λ false with "HA H†") as %EQAΛ.
rewrite {1}/lft_inv; iDestruct "Hinv" as "[[_ >%]|[Hinv >%]]".
{ unfold lft_alive_in in *. naive_solver. }
rewrite /lft_inv_dead; iDestruct "Hinv" as (Pinh) "(Hdead & >Hcnt & Hinh)".
iMod ("Hinh_close" $! Pinh with "Hinh") as (Pinh') "(? & $ & ?)".
iApply "Hclose". iExists A, I. iFrame. iNext. iApply "Hclose'".
rewrite /lft_inv. iRight. iFrame "%".
rewrite /lft_inv_dead. iExists Pinh'. iFrame.
- iFrame "HP". iApply fupd_frame_r. iSplitR ""; last by auto.
rewrite /lft_inv_dead. iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)" .
iMod (raw_bor_fake _ true with "Hdead") as "[Hdead Hbor]"; first solve_ndisj.
unfold bor. iExists κ. iFrame. rewrite -lft_incl_refl -big_sepS_later.
iApply "Hclose". iExists _, _. iFrame. iApply "Hclose'". iNext.
rewrite /lft_inv. iRight. rewrite /lft_inv_dead. iFrame. eauto.
Qed.
Lemma bor_sep E κ P Q : Lemma bor_sep E κ P Q :
lftN E lftN E
lft_ctx -∗ &{κ} (P Q) ={E}=∗ &{κ} P &{κ} Q. lft_ctx -∗ &{κ} (P Q) ={E}=∗ &{κ} P &{κ} Q.
...@@ -73,6 +16,7 @@ Proof. ...@@ -73,6 +16,7 @@ Proof.
iIntros (HE) "#LFT Hbor". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". iIntros (HE) "#LFT Hbor". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
rewrite {1}/bor. iDestruct "Hbor" as (κ') "[#Hκκ' Hbor]". rewrite {1}/bor. iDestruct "Hbor" as (κ') "[#Hκκ' Hbor]".
rewrite /raw_bor /idx_bor_own. iDestruct "Hbor" as (s) "[Hbor Hslice]". rewrite /raw_bor /idx_bor_own. iDestruct "Hbor" as (s) "[Hbor Hslice]".
iDestruct "Hslice" as (P') "[#HPP' Hslice]".
iDestruct (own_bor_auth with "HI Hbor") as %Hκ'. iDestruct (own_bor_auth with "HI Hbor") as %Hκ'.
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
/lfts_inv /lft_inv /lft_inv_dead /lft_alive_in. simpl. /lfts_inv /lft_inv /lft_inv_dead /lft_alive_in. simpl.
...@@ -81,41 +25,48 @@ Proof. ...@@ -81,41 +25,48 @@ Proof.
iDestruct "Hinv" as (Pb Pi) "(H & Hvs & Hinh)". iDestruct "Hinv" as (Pb Pi) "(H & Hvs & Hinh)".
iDestruct "H" as (B) "(Hbox & >Hown & HB)". iDestruct "H" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor") iDestruct (own_bor_valid_2 with "Hown Hbor")
as %[EQB%to_borUR_included _]%auth_valid_discrete_2. as %[EQB%to_borUR_included _]%auth_both_valid_discrete.
iMod (slice_split _ _ true with "Hslice Hbox") iMod (slice_iff _ _ true with "HPP' Hslice Hbox")
as (γ1 γ2) "(% & % & % & Hs1 & Hs2 & Hbox)"; first solve_ndisj. as (s' Pb') "(% & #HPbPb' & Hslice & Hbox)"; first solve_ndisj.
{ by rewrite lookup_fmap EQB. } { by rewrite lookup_fmap EQB. }
iAssert ( lft_vs κ' Pb' Pi)%I with "[Hvs]" as "Hvs".
{ iNext. iApply (lft_vs_cons with "[] Hvs"). iIntros "? _ !>".
by iApply "HPbPb'". }
iMod (slice_split _ _ true with "Hslice Hbox")
as (γ1 γ2) "(Hγ1 & Hγ2 & % & Hs1 & Hs2 & Hbox)"; first solve_ndisj.
{ by rewrite lookup_insert. }
rewrite delete_insert //. iDestruct "Hγ1" as %Hγ1. iDestruct "Hγ2" as %Hγ2.
iMod (own_bor_update_2 with "Hown Hbor") as "Hbor". iMod (own_bor_update_2 with "Hown Hbor") as "Hbor".
{ etrans; last etrans. { etrans; last etrans.
- apply auth_update_dealloc. by apply delete_singleton_local_update, _. - apply auth_update_dealloc. by apply delete_singleton_local_update, _.
- apply auth_update_alloc, - apply auth_update_alloc,
(alloc_singleton_local_update _ γ1 (1%Qp, DecAgree Bor_in)); last done. (alloc_singleton_local_update _ γ1 (1%Qp, to_agree Bor_in)); last done.
rewrite /to_borUR -fmap_delete lookup_fmap fmap_None rewrite /to_borUR -fmap_delete lookup_fmap fmap_None
-fmap_None -lookup_fmap fmap_delete //. -fmap_None -lookup_fmap fmap_delete //.
- apply cmra_update_op; last done. - apply cmra_update_op; last done.
apply auth_update_alloc, apply auth_update_alloc,
(alloc_singleton_local_update _ γ2 (1%Qp, DecAgree Bor_in)); last done. (alloc_singleton_local_update _ γ2 (1%Qp, to_agree Bor_in)); last done.
rewrite lookup_insert_ne // /to_borUR -fmap_delete lookup_fmap fmap_None rewrite lookup_insert_ne // /to_borUR -fmap_delete lookup_fmap fmap_None
-fmap_None -lookup_fmap fmap_delete //. } -fmap_None -lookup_fmap fmap_delete //. }
rewrite !own_bor_op. iDestruct "Hbor" as "[[H● H◯2] H◯1]". rewrite !own_bor_op. iDestruct "Hbor" as "[[H● H◯2] H◯1]".
iAssert (&{κ}P)%I with "[H◯1 Hs1]" as "$". iAssert (&{κ}P)%I with "[H◯1 Hs1]" as "$".
{ rewrite /bor /raw_bor /idx_bor_own. iExists κ'. iFrame "#". iExists γ1. iFrame. } { rewrite /bor /raw_bor /idx_bor_own. iFrame "#". by rewrite -(bi.iff_refl emp). }
iAssert (&{κ}Q)%I with "[H◯2 Hs2]" as "$". iAssert (&{κ}Q)%I with "[H◯2 Hs2]" as "$".
{ rewrite /bor /raw_bor /idx_bor_own. iExists κ'. iFrame "#". iExists γ2. iFrame. } { rewrite /bor /raw_bor /idx_bor_own. iFrame "#". by rewrite -(bi.iff_refl emp). }
iApply "Hclose". iExists A, I. iFrame. rewrite big_sepS_later. iApply "Hclose". iFrame. rewrite big_sepS_later.
iApply "Hclose'". iLeft. iFrame "%". iExists Pb, Pi. iFrame. iExists _. iApply "Hclose'". iLeft.
rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "Hbox H●". rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "%∗".
rewrite !big_sepM_insert /=. rewrite !big_sepM_insert /=.
+ by rewrite left_id -(big_sepM_delete _ _ _ _ EQB). + by rewrite left_id -(big_sepM_delete _ _ _ _ EQB).
+ by rewrite -fmap_None -lookup_fmap fmap_delete. + by rewrite -fmap_None -lookup_fmap fmap_delete.
+ by rewrite lookup_insert_ne // -fmap_None -lookup_fmap fmap_delete. + by rewrite lookup_insert_ne // -fmap_None -lookup_fmap fmap_delete.
- iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)". - iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)".
iMod (raw_bor_fake _ true with "Hdead") as "[Hdead Hbor1]"; first solve_ndisj. iMod (raw_bor_fake with "Hdead") as "[Hdead Hbor1]"; first solve_ndisj.
iMod (raw_bor_fake _ true with "Hdead") as "[Hdead Hbor2]"; first solve_ndisj. iMod (raw_bor_fake with "Hdead") as "[Hdead Hbor2]"; first solve_ndisj.
iMod ("Hclose" with "[-Hbor1 Hbor2]") as "_". iMod ("Hclose" with "[-Hbor1 Hbor2]") as "_".
{ iExists A, I. iFrame. rewrite big_sepS_later. iApply "Hclose'". { iFrame. rewrite big_sepS_later. iApply "Hclose'".
iRight. iSplit; last by auto. iExists _. iFrame. } iRight. by iFrame. }
unfold bor. iSplitL "Hbor1"; iExists _; eauto. unfold bor. by iFrame "#∗".
Qed. Qed.
Lemma bor_combine E κ P Q : Lemma bor_combine E κ P Q :
...@@ -124,12 +75,14 @@ Lemma bor_combine E κ P Q : ...@@ -124,12 +75,14 @@ Lemma bor_combine E κ P Q :
Proof. Proof.
iIntros (?) "#LFT HP HQ". rewrite {1 2}/bor. iIntros (?) "#LFT HP HQ". rewrite {1 2}/bor.
iDestruct "HP" as (κ1) "[#Hκ1 Hbor1]". iDestruct "HQ" as (κ2) "[#Hκ2 Hbor2]". iDestruct "HP" as (κ1) "[#Hκ1 Hbor1]". iDestruct "HQ" as (κ2) "[#Hκ2 Hbor2]".
iMod (raw_rebor _ _ (κ1 κ2) with "LFT Hbor1") as "[Hbor1 _]". iMod (raw_bor_shorten _ _ (κ1 κ2) with "LFT Hbor1") as "Hbor1"; first solve_ndisj.
done. by apply gmultiset_union_subseteq_l. { by apply gmultiset_disj_union_subseteq_l. }
iMod (raw_rebor _ _ (κ1 κ2) with "LFT Hbor2") as "[Hbor2 _]". iMod (raw_bor_shorten _ _ (κ1 κ2) with "LFT Hbor2") as "Hbor2"; first solve_ndisj.
done. by apply gmultiset_union_subseteq_r. { by apply gmultiset_disj_union_subseteq_r. }
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". unfold raw_bor, idx_bor_own. iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". unfold raw_bor, idx_bor_own.
iDestruct "Hbor1" as (j1) "[Hbor1 Hslice1]". iDestruct "Hbor2" as (j2) "[Hbor2 Hslice2]". iDestruct "Hbor1" as (j1) "[Hbor1 Hslice1]". iDestruct "Hbor2" as (j2) "[Hbor2 Hslice2]".
iDestruct "Hslice1" as (P') "[#HPP' Hslice1]".
iDestruct "Hslice2" as (Q') "[#HQQ' Hslice2]".
iDestruct (own_bor_auth with "HI Hbor1") as %Hκ'. iDestruct (own_bor_auth with "HI Hbor1") as %Hκ'.
rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom //
/lfts_inv /lft_inv /lft_inv_dead /lft_alive_in. simpl. /lfts_inv /lft_inv /lft_inv_dead /lft_alive_in. simpl.
...@@ -138,13 +91,13 @@ Proof. ...@@ -138,13 +91,13 @@ Proof.
iDestruct "Hinv" as (Pb Pi) "(H & Hvs & Hinh)". iDestruct "Hinv" as (Pb Pi) "(H & Hvs & Hinh)".
iDestruct "H" as (B) "(Hbox & >Hown & HB)". iDestruct "H" as (B) "(Hbox & >Hown & HB)".
iDestruct (own_bor_valid_2 with "Hown Hbor1") iDestruct (own_bor_valid_2 with "Hown Hbor1")
as %[EQB1%to_borUR_included _]%auth_valid_discrete_2. as %[EQB1%to_borUR_included _]%auth_both_valid_discrete.
iDestruct (own_bor_valid_2 with "Hown Hbor2") iDestruct (own_bor_valid_2 with "Hown Hbor2")
as %[EQB2%to_borUR_included _]%auth_valid_discrete_2. as %[EQB2%to_borUR_included _]%auth_both_valid_discrete.
iAssert j1 j2⌝%I with "[#]" as %Hj1j2. iAssert j1 j2⌝%I with "[#]" as %Hj1j2.
{ iIntros (->). { iIntros (->). iCombine "Hbor1 Hbor2" gives %Hj1j2.
iDestruct (own_bor_valid_2 with "Hbor1 Hbor2") as %Hj1j2%auth_valid_discrete. exfalso; revert Hj1j2.
exfalso. revert Hj1j2. rewrite /= op_singleton singleton_valid. rewrite -auth_frag_op auth_frag_valid singleton_op singleton_valid.
by intros [[]]. } by intros [[]]. }
iMod (slice_combine _ _ true with "Hslice1 Hslice2 Hbox") iMod (slice_combine _ _ true with "Hslice1 Hslice2 Hbox")
as (γ) "(% & Hslice & Hbox)"; first solve_ndisj. as (γ) "(% & Hslice & Hbox)"; first solve_ndisj.
...@@ -158,27 +111,28 @@ Proof. ...@@ -158,27 +111,28 @@ Proof.
apply auth_update_dealloc. by apply delete_singleton_local_update, _. apply auth_update_dealloc. by apply delete_singleton_local_update, _.
- apply auth_update_dealloc. by apply delete_singleton_local_update, _. - apply auth_update_dealloc. by apply delete_singleton_local_update, _.
- apply auth_update_alloc, - apply auth_update_alloc,
(alloc_singleton_local_update _ γ (1%Qp, DecAgree Bor_in)); last done. (alloc_singleton_local_update _ γ (1%Qp, to_agree Bor_in)); last done.
rewrite /to_borUR -!fmap_delete lookup_fmap fmap_None rewrite /to_borUR -!fmap_delete lookup_fmap fmap_None
-fmap_None -lookup_fmap !fmap_delete //. } -fmap_None -lookup_fmap !fmap_delete //. }
rewrite own_bor_op. iDestruct "Hbor" as "[H● H◯]". rewrite own_bor_op. iDestruct "Hbor" as "[H● H◯]".
iAssert (&{κ}(P Q))%I with "[H◯ Hslice]" as "$". iAssert (&{κ}(P Q))%I with "[H◯ Hslice]" as "$".
{ rewrite /bor /raw_bor /idx_bor_own. iExists (κ1 κ2). { rewrite /bor /raw_bor /idx_bor_own. iExists (κ1 κ2).
iSplit; first by iApply (lft_incl_glb with "Hκ1 Hκ2"). iSplit; first by iApply (lft_incl_glb with "Hκ1 Hκ2"). iFrame. iNext. iModIntro.
iExists γ. iFrame. } by iSplit; iIntros "[HP HQ]"; iSplitL "HP"; (iApply "HPP'" || iApply "HQQ'"). }
iApply "Hclose". iExists A, I. iFrame. rewrite big_sepS_later. iApply "Hclose". iFrame. rewrite big_sepS_later.
iApply "Hclose'". iLeft. iFrame "%". iExists Pb, Pi. iFrame. iExists _. iApply "Hclose'". iLeft.
rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "Hbox H●". rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "%∗".
rewrite !big_sepM_insert /=. rewrite !big_sepM_insert /=.
+ rewrite (big_sepM_delete _ _ _ _ EQB1) /=. + rewrite (big_sepM_delete _ _ _ _ EQB1) /=. iNext. simpl.
rewrite (big_sepM_delete _ _ j2 Bor_in) /=. by iDestruct "HB" as "[_ $]". rewrite [([ map] _ delete _ _, _)%I](big_sepM_delete _ _ j2 Bor_in) /=; last first.
rewrite lookup_delete_ne //. { rewrite lookup_delete_ne //. }
by iDestruct "HB" as "[_ $]".
+ rewrite -fmap_None -lookup_fmap !fmap_delete //. + rewrite -fmap_None -lookup_fmap !fmap_delete //.
- iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)". - iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)".
iMod (raw_bor_fake _ true with "Hdead") as "[Hdead Hbor]"; first solve_ndisj. iMod (raw_bor_fake with "Hdead") as "[Hdead Hbor]"; first solve_ndisj.
iMod ("Hclose" with "[-Hbor]") as "_". iMod ("Hclose" with "[-Hbor]") as "_".
{ iExists A, I. iFrame. rewrite big_sepS_later. iApply "Hclose'". { iFrame. rewrite big_sepS_later. iApply "Hclose'".
iRight. iSplit; last by auto. iExists _. iFrame. } iRight. by iFrame. }
unfold bor. iExists _. iFrame. iApply (lft_incl_glb with "Hκ1 Hκ2"). unfold bor. iFrame. iApply (lft_incl_glb with "Hκ1 Hκ2").
Qed. Qed.
End borrow. End borrow.
From lrust.lifetime Require Export primitive. From lrust.lifetime Require Export primitive.
From lrust.lifetime Require Import faking. From lrust.lifetime Require Import faking.
From iris.algebra Require Import csum auth frac gmap dec_agree gset. From iris.algebra Require Import csum auth frac gmap agree gset numbers.
From iris.base_logic Require Import big_op.
From iris.base_logic.lib Require Import boxes. From iris.base_logic.lib Require Import boxes.
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Section creation. Section creation.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft. Implicit Types κ : lft.
Lemma lft_kill (I : gmap lft lft_names) (K K' : gset lft) (κ : lft) : Lemma lft_kill (I : gmap lft lft_names) (K K' : gset lft) (κ : lft) :
let Iinv := ( let Iinv := (
own_ilft_auth I own_ilft_auth I
([ set] κ' K, lft_inv_alive κ') ([ set] κ' K, lft_inv_dead κ')
([ set] κ' K', lft_inv_dead κ'))%I in ([ set] κ' K', lft_inv_alive κ'))%I in
( κ', is_Some (I !! κ') κ' κ κ' K) ( κ', is_Some (I !! κ') κ κ' κ' K)
( κ', is_Some (I !! κ') κ κ' κ' K') ( κ', is_Some (I !! κ') κ' κ κ' K')
Iinv -∗ lft_inv_alive κ -∗ [κ] ={⊤∖↑mgmtN}=∗ Iinv lft_inv_dead κ. Iinv -∗ lft_inv_alive κ -∗ [κ] ={userE borN inhN}=∗ Iinv lft_inv_dead κ.
Proof. Proof.
iIntros (Iinv HK HK') "(HI & Halive & Hdead) Hlalive Hκ". iIntros (Iinv HK HK') "(HI & Hdead & Halive) Hlalive Hκ".
rewrite lft_inv_alive_unfold; rewrite lft_inv_alive_unfold;
iDestruct "Hlalive" as (P Q) "(Hbor & Hvs & Hinh)". iDestruct "Hlalive" as (P Q) "(Hbor & Hvs & Hinh)".
rewrite /lft_bor_alive; iDestruct "Hbor" as (B) "(Hbox & Hbor & HB)". rewrite /lft_bor_alive; iDestruct "Hbor" as (B) "(Hbox & Hbor & HB)".
...@@ -32,110 +32,180 @@ Proof. ...@@ -32,110 +32,180 @@ Proof.
iDestruct (@big_sepS_elem_of with "Hdead") as "Hdead"; first by eauto. iDestruct (@big_sepS_elem_of with "Hdead") as "Hdead"; first by eauto.
rewrite /lft_inv_dead; iDestruct "Hdead" as (R) "(_ & Hcnt' & _)". rewrite /lft_inv_dead; iDestruct "Hdead" as (R) "(_ & Hcnt' & _)".
iDestruct (own_cnt_valid_2 with "Hcnt' Hcnt") iDestruct (own_cnt_valid_2 with "Hcnt' Hcnt")
as %[?%nat_included _]%auth_valid_discrete_2; omega. } as %[?%nat_included _]%auth_both_valid_discrete; lia. }
iMod (box_empty with "Hbox") as "[HP Hbox]"=>//; first solve_ndisj. iMod (box_empty with "Hbox") as "[HP Hbox]"=>//; first by solve_ndisj.
{ intros i s. by rewrite lookup_fmap fmap_Some=> -[? [/HB -> ->]]. } { intros i s. by rewrite lookup_fmap fmap_Some=> -[? [/HB -> ->]]. }
rewrite lft_vs_unfold; iDestruct "Hvs" as (n) "[Hcnt Hvs]". rewrite lft_vs_unfold; iDestruct "Hvs" as (n) "[Hcnt Hvs]".
iDestruct (big_sepS_filter_acc ( κ) _ _ (dom _ I) with "Halive") iDestruct (big_sepS_filter_acc (. κ) _ _ (dom I) with "Halive")
as "[Halive Halive']". as "[Halive Halive']".
{ intros κ'. rewrite elem_of_dom. eauto. } { intros κ'. rewrite elem_of_dom. eauto. }
iMod ("Hvs" $! I with "[HI Halive Hbox Hbor] HP Hκ") as "(Hinv & HQ & Hcnt')". iApply fupd_trans. iApply fupd_mask_mono; last
{ rewrite lft_vs_inv_unfold. iFrame. rewrite /lft_bor_dead. iMod ("Hvs" $! I with "[HI Halive] HP Hκ") as "(Hinv & HQ & Hcnt')".
iExists (dom _ B), P. rewrite !to_gmap_dom -map_fmap_compose. { set_solver+. }
rewrite (map_fmap_ext _ ((1%Qp,) DecAgree) B); last naive_solver. { rewrite lft_vs_inv_unfold. iFrame. }
iFrame. } rewrite lft_vs_inv_unfold; iDestruct "Hinv" as "(HI&Halive)".
rewrite lft_vs_inv_unfold; iDestruct "Hinv" as "(?&HI&Halive)".
iSpecialize ("Halive'" with "Halive"). iSpecialize ("Halive'" with "Halive").
iMod (own_cnt_update_2 with "Hcnt Hcnt'") as "?". iMod (own_cnt_update_2 with "Hcnt Hcnt'") as "?".
{ apply auth_update_dealloc, (nat_local_update _ _ 0 0); omega. } { apply auth_update_dealloc, (nat_local_update _ _ 0 0); lia. }
rewrite /Iinv. iFrame "Hdead Halive' HI". rewrite /Iinv. iFrame "Hdead Halive' HI".
iMod (lft_inh_kill with "[$Hinh $HQ]"); first solve_ndisj. iModIntro. iMod (lft_inh_kill with "[$Hinh $HQ]"); first set_solver+.
iModIntro. rewrite /lft_inv_dead. iExists Q. by iFrame. iModIntro. rewrite /lft_inv_dead. iExists Q. iFrame.
rewrite /lft_bor_dead. iExists (dom B), P.
rewrite !gset_to_gmap_dom -map_fmap_compose.
rewrite (map_fmap_ext _ ((1%Qp,.) to_agree) B); last naive_solver.
iFrame.
Qed. Qed.
Lemma lfts_kill A (I : gmap lft lft_names) (K K' : gset lft) : Lemma lfts_kill (A : gmap atomic_lft _) (I : gmap lft lft_names) (K K' : gset lft) :
let Iinv K' := (own_ilft_auth I [ set] κ' K', lft_inv_alive κ')%I in let Iinv K' := (own_ilft_auth I [ set] κ' K', lft_inv_alive κ')%I in
K ## K'
( κ κ', κ K is_Some (I !! κ') κ κ' κ' K) ( κ κ', κ K is_Some (I !! κ') κ κ' κ' K)
( κ κ', κ K lft_alive_in A κ is_Some (I !! κ') ( κ, lft_alive_in A κ is_Some (I !! κ) κ K κ K')
κ' K κ' κ κ' K')
Iinv K' -∗ ([ set] κ K, lft_inv A κ [κ]) Iinv K' -∗ ([ set] κ K, lft_inv A κ [κ])
={⊤∖↑mgmtN}=∗ Iinv K' [ set] κ K, lft_inv_dead κ. ={userE borN inhN}=∗ Iinv K' [ set] κ K, lft_inv_dead κ.
Proof. Proof.
intros Iinv. revert K'. intros Iinv. revert K'.
induction (collection_wf K) as [K _ IH]=> K' HK HK'. induction (set_wf K) as [K _ IH]=> K' HKK' HK HK'.
iIntros "[HI Halive] HK". iIntros "[HI Halive] HK".
pose (Kalive := filter (lft_alive_in A) K). pose (Kalive := filter (lft_alive_in A) K).
destruct (decide (Kalive = )) as [HKalive|]. destruct (decide (Kalive = )) as [HKalive|].
{ iModIntro. rewrite /Iinv. iFrame. { iModIntro. rewrite /Iinv. iFrame.
iApply (@big_sepS_impl with "[$HK]"); iAlways. iApply (@big_sepS_impl with "[$HK]"); iModIntro.
rewrite /lft_inv. iIntros (κ ) "[[[_ %]|[$ _]] _]". set_solver. } rewrite /lft_inv. iIntros (κ ) "[[[_ %]|[$ _]] _]". set_solver. }
destruct (minimal_exists_L () Kalive) destruct (minimal_exists_elem_of_L () Kalive)
as (κ & [Hκalive HκK]%elem_of_filter & Hκmin); first done. as (κ & [Hκalive HκK]%elem_of_filter & Hκmin); first done.
iDestruct (@big_sepS_delete with "HK") as "[[Hκinv Hκ] HK]"; first done. iDestruct (@big_sepS_delete with "HK") as "[[Hκinv Hκ] HK]"; first done.
iDestruct (lft_inv_alive_in with "Hκinv") as "Hκalive"; first done. iDestruct (lft_inv_alive_in with "Hκinv") as "Hκalive"; first done.
iAssert κ K'⌝%I with "[#]" as %HκK'. assert (κ K') as HκK' by set_solver +HκK HKK'.
{ iIntros (). iApply (lft_inv_alive_twice κ with "Hκalive"). ospecialize (IH (K {[ κ ]}) _); [set_solver +HκK|].
by iApply (@big_sepS_elem_of with "Halive"). }
specialize (IH (K {[ κ ]})). feed specialize IH; [set_solver +HκK|].
iMod (IH ({[ κ ]} K') with "[HI Halive Hκalive] HK") as "[[HI Halive] Hdead]". iMod (IH ({[ κ ]} K') with "[HI Halive Hκalive] HK") as "[[HI Halive] Hdead]".
{ set_solver +HKK'. }
{ intros κ' κ''. { intros κ' κ''.
rewrite !elem_of_difference !elem_of_singleton=> -[? Hneq] ??. rewrite !elem_of_difference !elem_of_singleton=> -[? Hneq] ??.
split; [by eauto|]; intros ->. split; [by eauto|]; intros ->.
eapply (minimal_strict_1 _ _ κ' Hκmin), strict_spec_alt; eauto. eapply (minimal_strict_1 _ _ κ' Hκmin), strict_spec_alt; eauto.
apply elem_of_filter; eauto using lft_alive_in_subseteq. } apply elem_of_filter; eauto using lft_alive_in_subseteq. }
{ intros κ' κ'' Hκ' ? [γs'' ?]. { intros κ' Hκ'. destruct (decide (κ' = κ)) as [->|Hκκ']; [set_solver +|].
destruct (decide (κ'' = κ)) as [->|]; [set_solver +|]. specialize (HK' _ Hκ'). set_solver +Hκκ' HK'. }
move: Hκ'; rewrite not_elem_of_difference elem_of_difference
elem_of_union not_elem_of_singleton elem_of_singleton.
intros [??] [?|?]; eauto. }
{ rewrite /Iinv big_sepS_insert //. iFrame. } { rewrite /Iinv big_sepS_insert //. iFrame. }
iDestruct (@big_sepS_insert with "Halive") as "[Hκalive Halive]"; first done. iDestruct (@big_sepS_insert with "Halive") as "[Hκalive Halive]"; first done.
iMod (lft_kill with "[$HI $Halive $Hdead] Hκalive Hκ") iMod (lft_kill with "[$HI $Halive $Hdead] Hκalive Hκ")
as "[(HI&Halive&Hdead) Hκdead]". as "[(HI&Halive&Hdead) Hκdead]".
{ intros κ' ??. eapply (HK' κ); eauto.
intros ?. eapply (minimal_strict_1 _ _ _ Hκmin); eauto.
apply elem_of_filter; split; last done.
eapply lft_alive_in_subseteq, gmultiset_subset_subseteq; eauto. }
{ intros κ' ? [??]%strict_spec_alt. { intros κ' ? [??]%strict_spec_alt.
rewrite elem_of_difference elem_of_singleton; eauto. } rewrite elem_of_difference elem_of_singleton; eauto. }
{ intros κ' ??. eapply HK'; [|done|].
- by eapply lft_alive_in_subseteq, gmultiset_subset_subseteq.
- intros ?. eapply (minimal_strict_1 _ _ _ Hκmin); eauto.
apply elem_of_filter; split; last done.
eapply lft_alive_in_subseteq, gmultiset_subset_subseteq; eauto. }
iModIntro. rewrite /Iinv (big_sepS_delete _ K) //. iFrame. iModIntro. rewrite /Iinv (big_sepS_delete _ K) //. iFrame.
Qed. Qed.
Definition kill_set (I : gmap lft lft_names) (Λ : atomic_lft) : gset lft := Definition kill_set (I : gmap lft lft_names) (Λ : atomic_lft) : gset lft :=
filter (Λ ) (dom (gset lft) I). filter (Λ .) (dom I).
Lemma elem_of_kill_set I Λ κ : κ kill_set I Λ Λ κ is_Some (I !! κ). Lemma elem_of_kill_set I Λ κ : κ kill_set I Λ Λ κ is_Some (I !! κ).
Proof. by rewrite /kill_set elem_of_filter elem_of_dom. Qed. Proof. by rewrite /kill_set elem_of_filter elem_of_dom. Qed.
Lemma kill_set_subseteq I Λ : kill_set I Λ dom (gset lft) I.
Proof. intros κ [??]%elem_of_kill_set. by apply elem_of_dom. Qed.
Definition down_close (A : gmap atomic_lft bool) Lemma lupd_gmap_mass_insert (Λs : gset atomic_lft) (m : alftUR) (x y : lft_stateR) `{!Exclusive x} : y
(I : gmap lft lft_names) (K : gset lft) : gset lft := (m, gset_to_gmap x Λs) ~l~> (gset_to_gmap y Λs m, gset_to_gmap y Λs).
filter (λ κ, κ K Proof.
set_Exists (λ κ', κ κ' lft_alive_in A κ') K) (dom (gset lft) I). intros ?.
Lemma elem_of_down_close A I K κ : apply gmap_local_update.
κ down_close A I K intros Λ; destruct (m !! Λ) as [|] eqn:He1, (decide (Λ Λs)) as [|].
is_Some (I !! κ) κ K κ', κ' K κ κ' lft_alive_in A κ'. - rewrite lookup_union_l' !lookup_gset_to_gmap !option_guard_True // He1.
Proof. rewrite /down_close elem_of_filter elem_of_dom /set_Exists. tauto. Qed. by apply option_local_update, exclusive_local_update.
- rewrite lookup_union_r !lookup_gset_to_gmap !option_guard_False //= He1 //.
- rewrite lookup_union_l' !lookup_gset_to_gmap !option_guard_True // He1.
by apply alloc_option_local_update.
- rewrite lookup_union_r !lookup_gset_to_gmap !option_guard_False //= He1 //.
Qed.
Lemma down_close_lft_alive_in A I K κ : κ down_close A I K lft_alive_in A κ. Lemma lft_kill_atomics (Λs : gset atomic_lft) :
lft_ctx -∗
(([ set] ΛΛs, 1.[@Λ]) ={lftN userE}[userE]▷=∗
([ set] ΛΛs, [@Λ])).
Proof. Proof.
rewrite elem_of_down_close; intros (?&?&?&?&?&?). assert (userE_lftN_disj:=userE_lftN_disj). iIntros "#LFT !> HΛs".
eapply lft_alive_in_subseteq, gmultiset_subset_subseteq; eauto. destruct (decide (Λs = )) as [->|Hne].
{ rewrite !big_sepS_empty.
iApply fupd_mask_intro; [set_solver|iIntros "$"].
}
iApply (step_fupd_mask_mono (lftN userE) _ ((lftN userE)∖↑mgmtN)); [solve_ndisj..|].
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
rewrite /lft_tok.
iAssert (own alft_name ( gset_to_gmap (Cinl 1%Qp) Λs))%I with "[HΛs]" as "HΛs".
{ setoid_rewrite big_sepMS_singleton.
iDestruct (big_opS_own _ _ _ Hne with "HΛs") as "HΛs".
rewrite -(big_opS_gset_to_gmap Λs (Cinl 1%Qp)) big_opS_auth_frag //. }
iDestruct (own_valid_2 with "HA HΛs")
as %[[s Hs] _]%auth_both_valid_discrete.
iMod (own_update_2 with "HA HΛs") as "[HA HΛs]".
{ eapply auth_update.
by apply (lupd_gmap_mass_insert Λs _ _ (Cinr ())). }
rewrite -[in own alft_name ( _)]big_opS_gset_to_gmap big_opS_auth_frag big_opS_own //.
iDestruct "HΛs" as "#HΛs".
iModIntro; iNext.
pose (K := set_bind (λ Λ, kill_set I Λ) Λs).
pose (K' := filter (lft_alive_in A) (dom I) K).
destruct (proj1 (subseteq_disjoint_union_L (K K') (dom I))) as (K''&HI&HK'').
{ set_solver+. }
assert (K ## K') by set_solver+.
rewrite HI !big_sepS_union //. iDestruct "Hinv" as "[[HinvK HinvD] Hinv]".
iAssert ([ set] κ K', lft_inv_alive κ)%I with "[HinvD]" as "HinvD".
{ iApply (@big_sepS_impl with "[$HinvD]"); iIntros "!>".
iIntros (κ [[ _]%elem_of_filter _]%elem_of_difference) "?".
by iApply lft_inv_alive_in. }
iAssert ([ set] κ K, lft_inv A κ [ κ])%I with "[HinvK]" as "HinvK".
{ iApply (@big_sepS_impl with "[$HinvK]"); iIntros "!>".
iIntros (κ [Λ [ [? _]%elem_of_kill_set]]%elem_of_set_bind) "$".
rewrite /lft_dead.
by iDestruct (big_sepS_elem_of with "HΛs") as "$". }
iApply fupd_trans.
iApply (fupd_mask_mono (userE borN inhN)); first solve_ndisj.
iMod (lfts_kill A I K K' with "[$HI $HinvD] HinvK") as "[[HI HinvD] HinvK]".
{ done. }
{ intros κ κ' [Λ [? [??]%elem_of_kill_set]]%elem_of_set_bind ??.
apply elem_of_set_bind; exists Λ; split; [assumption|].
apply elem_of_kill_set.
split; last done. by eapply gmultiset_elem_of_subseteq. }
{ intros κ ???. rewrite elem_of_difference elem_of_filter elem_of_dom. auto. }
iModIntro. iMod ("Hclose" with "[-]") as "_"; last first.
{ iModIntro.
iApply (big_sepS_impl with "HΛs").
iIntros "!>" (Λ) "??".
rewrite /lft_dead. iExists Λ.
rewrite gmultiset_elem_of_singleton. auto. }
iNext. iExists ((gset_to_gmap false Λs A) : gmap atomic_lft bool), I.
rewrite /own_alft_auth /to_alftUR map_fmap_union fmap_gset_to_gmap.
iFrame "HA HI".
rewrite HI !big_sepS_union //.
iSplitL "HinvK HinvD"; first iSplitL "HinvK".
- iApply (@big_sepS_impl with "[$HinvK]"); iIntros "!>".
iIntros (κ [Λ [? [? _]%elem_of_kill_set]]%elem_of_set_bind) "Hdead". rewrite /lft_inv.
iRight. iFrame. iPureIntro. by eapply lft_dead_in_union_false'.
- iApply (@big_sepS_impl with "[$HinvD]"); iIntros "!>".
iIntros (κ [[ HκI]%elem_of_filter HκK]%elem_of_difference) "Halive".
rewrite /lft_inv. iLeft. iFrame "Halive". iPureIntro.
apply lft_alive_in_union_false; [|done].
set_solver.
- iApply (@big_sepS_impl with "[$Hinv]"); iIntros "!>".
rewrite /lft_inv. iIntros (κ ) "[[? %]|[? %]]".
+ iLeft. iFrame. iPureIntro.
apply lft_alive_in_union_false; last done. intros ??.
set_solver.
+ iRight. iFrame. iPureIntro. by apply lft_dead_in_union_false.
Qed. Qed.
Lemma down_close_disjoint A I K : K down_close A I K.
Proof. intros κ. rewrite elem_of_down_close. naive_solver. Qed.
Lemma down_close_subseteq A I K :
down_close A I K dom (gset lft) I.
Proof. intros κ [??]%elem_of_down_close. by apply elem_of_dom. Qed.
Lemma lft_create E : Lemma lft_create_strong P E :
lftN E pred_infinite P lftN E
lft_ctx ={E}=∗ κ, 1.[κ] (1.[κ] ={,⊤∖↑lftN}▷=∗ [κ]). lft_ctx ={E}=∗
p : positive, let Λ := positive_to_atomic_lft p in P p (1).[@Λ].
Proof. Proof.
iIntros (?) "#LFT". assert (userE_lftN_disj:=userE_lftN_disj). iIntros (HP ?) "#LFT".
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose".
destruct (exist_fresh (dom (gset _) A)) as [Λ %not_elem_of_dom]. rewrite ->(pred_infinite_set (C:=gset _)) in HP.
destruct (HP (dom A)) as [Λ [HPx %not_elem_of_dom]].
iMod (own_update with "HA") as "[HA HΛ]". iMod (own_update with "HA") as "[HA HΛ]".
{ apply auth_update_alloc, (alloc_singleton_local_update _ Λ (Cinl 1%Qp))=>//. { apply auth_update_alloc, (alloc_singleton_local_update _ Λ (Cinl 1%Qp))=>//.
by rewrite lookup_fmap . } by rewrite lookup_fmap . }
...@@ -143,62 +213,23 @@ Proof. ...@@ -143,62 +213,23 @@ Proof.
{ iNext. rewrite /lfts_inv /own_alft_auth. { iNext. rewrite /lfts_inv /own_alft_auth.
iExists (<[Λ:=true]>A), I. rewrite /to_alftUR fmap_insert; iFrame. iExists (<[Λ:=true]>A), I. rewrite /to_alftUR fmap_insert; iFrame.
iApply (@big_sepS_impl with "[$Hinv]"). iApply (@big_sepS_impl with "[$Hinv]").
iAlways. rewrite /lft_inv. iIntros (κ ?) "[[Hκ %]|[Hκ %]]". iModIntro. rewrite /lft_inv. iIntros (κ ?) "[[Hκ %]|[Hκ %]]".
- iLeft. iFrame "Hκ". iPureIntro. by apply lft_alive_in_insert. - iLeft. iFrame "Hκ". iPureIntro. by apply lft_alive_in_insert.
- iRight. iFrame "Hκ". iPureIntro. by apply lft_dead_in_insert. } - iRight. iFrame "Hκ". iPureIntro. by apply lft_dead_in_insert. }
iModIntro; iExists {[ Λ ]}. iModIntro; iExists Λ.
rewrite {1}/lft_tok big_sepMS_singleton. iFrame "HΛ". rewrite {1}/lft_tok big_sepMS_singleton. iSplit; first done. iFrame "HΛ".
clear I A . iIntros "!# HΛ". Qed.
iApply (step_fupd_mask_mono _ _ (⊤∖↑mgmtN)); [solve_ndisj..|].
iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". Lemma lft_fresh κ κ' :
rewrite /lft_tok big_sepMS_singleton. i : positive, m : positive, (i < m)%positive (atomic_lft_to_lft (positive_to_atomic_lft m)) κ' κ.
iDestruct (own_valid_2 with "HA HΛ") Proof.
as %[[s [?%leibniz_equiv ?]]%singleton_included _]%auth_valid_discrete_2. unfold meet, lft_intersect.
iMod (own_update_2 with "HA HΛ") as "[HA HΛ]". destruct (minimal_exists (flip Pos.le) (dom κ)) as (i & Ha).
{ by eapply auth_update, singleton_local_update, exists (i + 1)%positive.
(exclusive_local_update _ (Cinr ())). } intros k Hik Heq. subst κ.
iDestruct "HΛ" as "#HΛ". iModIntro; iNext. ospecialize (Ha k _ _); simpl in *; [ | lia..].
pose (K := kill_set I Λ); pose (K' := down_close A I K). apply gmultiset_elem_of_dom.
assert (K K') by (by apply down_close_disjoint). apply gmultiset_elem_of_disj_union.
destruct (proj1 (subseteq_disjoint_union_L (K K') left. by apply gmultiset_elem_of_singleton.
(dom (gset lft) I))) as (K''&HI&?).
{ apply union_least. apply kill_set_subseteq. apply down_close_subseteq. }
rewrite HI !big_sepS_union //. iDestruct "Hinv" as "[[HinvK HinvD] Hinv]".
iAssert ([ set] κ K', lft_inv_alive κ)%I with "[HinvD]" as "HinvD".
{ iApply (@big_sepS_impl with "[$HinvD]"); iIntros "!#".
iIntros (κ ) "?". iApply lft_inv_alive_in; last done.
eauto using down_close_lft_alive_in. }
iAssert ([ set] κ K, lft_inv A κ [ κ])%I with "[HinvK]" as "HinvK".
{ iApply (@big_sepS_impl with "[$HinvK]"); iIntros "!#".
iIntros (κ [? _]%elem_of_kill_set) "$". rewrite /lft_dead. eauto. }
iMod (lfts_kill A I K K' with "[$HI $HinvD] HinvK") as "[[HI HinvD] HinvK]".
{ intros κ κ' [??]%elem_of_kill_set ??. apply elem_of_kill_set.
split; last done. by eapply gmultiset_elem_of_subseteq. }
{ intros κ κ' ?????. apply elem_of_down_close; eauto 10. }
iMod ("Hclose" with "[-]") as "_"; last first.
{ iModIntro. rewrite /lft_dead. iExists Λ.
rewrite elem_of_singleton. auto. }
iNext. iExists (<[Λ:=false]>A), I.
rewrite /own_alft_auth /to_alftUR fmap_insert. iFrame "HA HI".
rewrite HI !big_sepS_union //.
iSplitL "HinvK HinvD"; first iSplitL "HinvK".
- iApply (@big_sepS_impl with "[$HinvK]"); iIntros "!#".
iIntros (κ [? _]%elem_of_kill_set) "Hdead". rewrite /lft_inv.
iRight. iFrame. iPureIntro. by apply lft_dead_in_insert_false'.
- iApply (@big_sepS_impl with "[$HinvD]"); iIntros "!#".
iIntros (κ ) "Halive". rewrite /lft_inv.
iLeft. iFrame "Halive". iPureIntro.
assert (lft_alive_in A κ) as Halive by (by eapply down_close_lft_alive_in).
apply lft_alive_in_insert_false; last done.
apply elem_of_down_close in as (?&HFOO&_).
move: HFOO. rewrite elem_of_kill_set. tauto.
- iApply (@big_sepS_impl with "[$Hinv]"); iIntros "!#".
rewrite /lft_inv. iIntros (κ ) "[[? %]|[? %]]".
+ iLeft. iFrame. iPureIntro.
apply lft_alive_in_insert_false; last done.
assert (κ K). intros ?. eapply H5. 2: eauto. rewrite elem_of_union. eauto.
move: H7. rewrite elem_of_kill_set not_and_l. intros [?|?]. done.
contradict H7. apply elem_of_dom. set_solver +HI .
+ iRight. iFrame. iPureIntro. by apply lft_dead_in_insert_false.
Qed. Qed.
End creation. End creation.
From iris.algebra Require Import csum auth frac gmap dec_agree gset. From iris.algebra Require Import csum auth frac gmap agree gset numbers.
From iris.prelude Require Export gmultiset strings.
From iris.base_logic.lib Require Export invariants.
From iris.base_logic.lib Require Import boxes. From iris.base_logic.lib Require Import boxes.
From iris.base_logic Require Import big_op. From lrust.lifetime Require Export lifetime_sig.
From iris.prelude Require Import options.
Import uPred. Import uPred.
Definition lftN : namespace := nroot .@ "lft".
Definition borN : namespace := lftN .@ "bor".
Definition inhN : namespace := lftN .@ "inh". Definition inhN : namespace := lftN .@ "inh".
Definition mgmtN : namespace := lftN .@ "mgmt". Definition mgmtN : namespace := lftN .@ "mgmt".
Definition borN : namespace := lftN .@ "bor".
Definition atomic_lft := positive. Definition atomic_lft := positive.
Notation lft := (gmultiset positive). (* HACK : We need to make sure this is not in the top-level context,
Definition static : lft := ∅. so that it does not conflict with the *definition* of [lft] that we
use in lifetime.v. *)
Module Export lft_notation.
Notation lft := (gmultiset atomic_lft).
End lft_notation.
Definition static : lft := ( : gmultiset _).
Global Instance lft_intersect : Meet lft := λ κ κ', κ κ'.
Definition positive_to_atomic_lft (p : positive) : atomic_lft := p.
Definition atomic_lft_to_lft (Λ : atomic_lft) : lft := {[+ Λ +]}.
Inductive bor_state := Inductive bor_state :=
| Bor_in | Bor_in
| Bor_open (q : frac) | Bor_open (q : frac)
| Bor_rebor (κ : lft). | Bor_rebor (κ : lft).
Instance bor_state_eq_dec : EqDecision bor_state. Canonical bor_stateO := leibnizO bor_state.
Proof. solve_decision. Defined.
Definition bor_filled (s : bor_state) : bool :=
match s with Bor_in => true | _ => false end.
Definition lft_stateR := csumR fracR unitR.
Definition to_lft_stateR (b : bool) : lft_stateR :=
if b then Cinl 1%Qp else Cinr ().
Record lft_names := LftNames { Record lft_names := LftNames {
bor_name : gname; bor_name : gname;
cnt_name : gname; cnt_name : gname;
inh_name : gname inh_name : gname
}. }.
Instance lft_names_eq_dec : EqDecision lft_names. Canonical lft_namesO := leibnizO lft_names.
Proof. solve_decision. Defined.
Definition lft_stateR := csumR fracR unitR.
Definition alftUR := gmapUR atomic_lft lft_stateR. Definition alftUR := gmapUR atomic_lft lft_stateR.
Definition to_alftUR : gmap atomic_lft bool alftUR := fmap to_lft_stateR. Definition ilftUR := gmapUR lft (agreeR lft_namesO).
Definition borUR := gmapUR slice_name (prodR fracR (agreeR bor_stateO)).
Definition ilftUR := gmapUR lft (dec_agreeR lft_names).
Definition to_ilftUR : gmap lft lft_names ilftUR := fmap DecAgree.
Definition borUR := gmapUR slice_name (prodR fracR (dec_agreeR bor_state)).
Definition to_borUR : gmap slice_name bor_state borUR := fmap ((1%Qp,) DecAgree).
Definition inhUR := gset_disjUR slice_name. Definition inhUR := gset_disjUR slice_name.
Class lftG Σ := LftG { Class lftGS Σ (userE : coPset) := LftG {
lft_box :> boxG Σ; lft_box :: boxG Σ;
alft_inG :> inG Σ (authR alftUR); alft_inG :: inG Σ (authR alftUR);
alft_name : gname; alft_name : gname;
ilft_inG :> inG Σ (authR ilftUR); ilft_inG :: inG Σ (authR ilftUR);
ilft_name : gname; ilft_name : gname;
lft_bor_box :> inG Σ (authR borUR); lft_bor_inG :: inG Σ (authR borUR);
lft_cnt_inG :> inG Σ (authR natUR); lft_cnt_inG :: inG Σ (authR natUR);
lft_inh_box :> inG Σ (authR inhUR); lft_inh_inG :: inG Σ (authR inhUR);
userE_lftN_disj : lftN ## userE;
}.
Definition lftGS' := lftGS.
Class lftGpreS Σ := LftGPreS {
lft_preG_box :: boxG Σ;
alft_preG_inG :: inG Σ (authR alftUR);
ilft_preG_inG :: inG Σ (authR ilftUR);
lft_preG_bor_inG :: inG Σ (authR borUR);
lft_preG_cnt_inG :: inG Σ (authR natUR);
lft_preG_inh_inG :: inG Σ (authR inhUR);
}. }.
Definition lftGpreS' := lftGpreS.
Definition lftΣ : gFunctors :=
#[ boxΣ; GFunctor (authR alftUR); GFunctor (authR ilftUR);
GFunctor (authR borUR); GFunctor (authR natUR); GFunctor (authR inhUR) ].
Global Instance subG_lftGpreS Σ :
subG lftΣ Σ lftGpreS Σ.
Proof. solve_inG. Qed.
Definition bor_filled (s : bor_state) : bool :=
match s with Bor_in => true | _ => false end.
Definition to_lft_stateR (b : bool) : lft_stateR :=
if b then Cinl 1%Qp else Cinr ().
Definition to_alftUR : gmap atomic_lft bool alftUR := fmap to_lft_stateR.
Definition to_ilftUR : gmap lft lft_names ilftUR := fmap to_agree.
Definition to_borUR : gmap slice_name bor_state borUR := fmap ((1%Qp,.) to_agree).
Section defs. Section defs.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Definition lft_tok (q : Qp) (κ : lft) : iProp Σ := Definition lft_tok (q : Qp) (κ : lft) : iProp Σ :=
([ mset] Λ κ, own alft_name ( {[ Λ := Cinl q ]}))%I. ([ mset] Λ κ, own alft_name ( {[ Λ := Cinl q ]}))%I.
...@@ -73,19 +95,19 @@ Section defs. ...@@ -73,19 +95,19 @@ Section defs.
own ilft_name ( to_ilftUR I). own ilft_name ( to_ilftUR I).
Definition own_bor (κ : lft) Definition own_bor (κ : lft)
(x : auth (gmap slice_name (frac * dec_agree bor_state))) : iProp Σ := (x : authR borUR) : iProp Σ :=
( γs, ( γs,
own ilft_name ( {[ κ := DecAgree γs ]}) own ilft_name ( {[ κ := to_agree γs ]})
own (bor_name γs) x)%I. own (bor_name γs) x)%I.
Definition own_cnt (κ : lft) (x : auth nat) : iProp Σ := Definition own_cnt (κ : lft) (x : authR natUR) : iProp Σ :=
( γs, ( γs,
own ilft_name ( {[ κ := DecAgree γs ]}) own ilft_name ( {[ κ := to_agree γs ]})
own (cnt_name γs) x)%I. own (cnt_name γs) x)%I.
Definition own_inh (κ : lft) (x : auth (gset_disj slice_name)) : iProp Σ := Definition own_inh (κ : lft) (x : authR inhUR) : iProp Σ :=
( γs, ( γs,
own ilft_name ( {[ κ := DecAgree γs ]}) own ilft_name ( {[ κ := to_agree γs ]})
own (inh_name γs) x)%I. own (inh_name γs) x)%I.
Definition bor_cnt (κ : lft) (s : bor_state) : iProp Σ := Definition bor_cnt (κ : lft) (s : bor_state) : iProp Σ :=
...@@ -103,19 +125,18 @@ Section defs. ...@@ -103,19 +125,18 @@ Section defs.
Definition lft_bor_dead (κ : lft) : iProp Σ := Definition lft_bor_dead (κ : lft) : iProp Σ :=
( (B: gset slice_name) (Pb : iProp Σ), ( (B: gset slice_name) (Pb : iProp Σ),
own_bor κ ( to_gmap (1%Qp, DecAgree Bor_in) B) own_bor κ ( gset_to_gmap (1%Qp, to_agree Bor_in) B)
box borN (to_gmap false B) Pb)%I. box borN (gset_to_gmap false B) Pb)%I.
Definition lft_inh (κ : lft) (s : bool) (Pi : iProp Σ) : iProp Σ := Definition lft_inh (κ : lft) (s : bool) (Pi : iProp Σ) : iProp Σ :=
( E : gset slice_name, ( E : gset slice_name,
own_inh κ ( GSet E) own_inh κ ( GSet E)
box inhN (to_gmap s E) Pi)%I. box inhN (gset_to_gmap s E) Pi)%I.
Definition lft_vs_inv_go (κ : lft) (lft_inv_alive : κ', κ' κ iProp Σ) Definition lft_vs_inv_go (κ : lft) (lft_inv_alive : κ', κ' κ iProp Σ)
(I : gmap lft lft_names) : iProp Σ := (I : gmap lft lft_names) : iProp Σ :=
(lft_bor_dead κ (own_ilft_auth I
own_ilft_auth I [ set] κ' dom I, : κ' κ, lft_inv_alive κ' )%I.
[ set] κ' dom _ I, : κ' κ, lft_inv_alive κ' )%I.
Definition lft_vs_go (κ : lft) (lft_inv_alive : κ', κ' κ iProp Σ) Definition lft_vs_go (κ : lft) (lft_inv_alive : κ', κ' κ iProp Σ)
(Pb Pi : iProp Σ) : iProp Σ := (Pb Pi : iProp Σ) : iProp Σ :=
...@@ -123,7 +144,7 @@ Section defs. ...@@ -123,7 +144,7 @@ Section defs.
own_cnt κ ( n) own_cnt κ ( n)
I : gmap lft lft_names, I : gmap lft lft_names,
lft_vs_inv_go κ lft_inv_alive I -∗ Pb -∗ lft_dead κ lft_vs_inv_go κ lft_inv_alive I -∗ Pb -∗ lft_dead κ
={⊤∖↑mgmtN}=∗ ={userE borN}=∗
lft_vs_inv_go κ lft_inv_alive I Pi own_cnt κ ( n))%I. lft_vs_inv_go κ lft_inv_alive I Pi own_cnt κ ( n))%I.
Definition lft_inv_alive_go (κ : lft) Definition lft_inv_alive_go (κ : lft)
...@@ -159,7 +180,7 @@ Section defs. ...@@ -159,7 +180,7 @@ Section defs.
( (A : gmap atomic_lft bool) (I : gmap lft lft_names), ( (A : gmap atomic_lft bool) (I : gmap lft lft_names),
own_alft_auth A own_alft_auth A
own_ilft_auth I own_ilft_auth I
[ set] κ dom _ I, lft_inv A κ)%I. [ set] κ dom I, lft_inv A κ)%I.
Definition lft_ctx : iProp Σ := inv mgmtN lfts_inv. Definition lft_ctx : iProp Σ := inv mgmtN lfts_inv.
...@@ -171,39 +192,43 @@ Section defs. ...@@ -171,39 +192,43 @@ Section defs.
Definition bor_idx := (lft * slice_name)%type. Definition bor_idx := (lft * slice_name)%type.
Definition idx_bor_own (q : frac) (i : bor_idx) : iProp Σ := Definition idx_bor_own (q : frac) (i : bor_idx) : iProp Σ :=
own_bor (i.1) ( {[ i.2 := (q,DecAgree Bor_in) ]}). own_bor (i.1) ( {[ i.2 := (q, to_agree Bor_in) ]}).
Definition idx_bor (κ : lft) (i : bor_idx) (P : iProp Σ) : iProp Σ := Definition idx_bor (κ : lft) (i : bor_idx) (P : iProp Σ) : iProp Σ :=
(lft_incl κ (i.1) slice borN (i.2) P)%I. (lft_incl κ (i.1) P', (P' P) slice borN (i.2) P')%I.
Definition raw_bor (κ : lft) (P : iProp Σ) : iProp Σ := Definition raw_bor (κ : lft) (P : iProp Σ) : iProp Σ :=
( s : slice_name, idx_bor_own 1 (κ, s) slice borN s P)%I. ( s : slice_name, idx_bor_own 1 (κ, s) P', (P' P) slice borN s P')%I.
Definition bor (κ : lft) (P : iProp Σ) : iProp Σ := Definition bor (κ : lft) (P : iProp Σ) : iProp Σ :=
( κ', lft_incl κ κ' raw_bor κ' P)%I. ( κ', lft_incl κ κ' raw_bor κ' P)%I.
End defs. End defs.
Instance: Params (@lft_bor_alive) 4. Global Instance: Params (@lft_bor_alive) 4 := {}.
Instance: Params (@lft_inh) 5. Global Instance: Params (@lft_inh) 5 := {}.
Instance: Params (@lft_vs) 4. Global Instance: Params (@lft_vs) 4 := {}.
Instance: Params (@idx_bor) 5. Global Instance idx_bor_params : Params (@idx_bor) 5 := {}.
Instance: Params (@raw_bor) 4. Global Instance raw_bor_params : Params (@raw_bor) 4 := {}.
Instance: Params (@bor) 4. Global Instance bor_params : Params (@bor) 4 := {}.
Notation "q .[ κ ]" := (lft_tok q κ) Notation "q .[ κ ]" := (lft_tok q κ)
(format "q .[ κ ]", at level 0) : uPred_scope. (format "q .[ κ ]", at level 2, left associativity) : bi_scope.
Notation "[† κ ]" := (lft_dead κ) (format "[† κ ]"): uPred_scope. Notation "q .[@ Λ ]" := (lft_tok q (atomic_lft_to_lft Λ))
(format "q .[@ Λ ]", at level 2, left associativity) : bi_scope.
Notation "[† κ ]" := (lft_dead κ) (format "[† κ ]"): bi_scope.
Notation "[†@ Λ ]" := (lft_dead (atomic_lft_to_lft Λ)) (format "[†@ Λ ]"): bi_scope.
Notation "&{ κ } P" := (bor κ P) Notation "&{ κ }" := (bor κ) (format "&{ κ }") : bi_scope.
(format "&{ κ } P", at level 20, right associativity) : uPred_scope. Notation "&{ κ , i }" := (idx_bor κ i) (format "&{ κ , i }") : bi_scope.
Notation "&{ κ , i } P" := (idx_bor κ i P)
(format "&{ κ , i } P", at level 20, right associativity) : uPred_scope.
Infix "⊑" := lft_incl (at level 70) : uPred_scope. Infix "⊑" := lft_incl (at level 70) : bi_scope.
Typeclasses Opaque lft_tok lft_dead bor_cnt lft_bor_alive lft_bor_dead (* TODO: Making all these things opaque is rather annoying, we should
find a way to avoid it, or *at least*, to avoid having to manually unfold
this because iDestruct et al don't look through these names any more. *)
Global Typeclasses Opaque lft_tok lft_dead bor_cnt lft_bor_alive lft_bor_dead
lft_inh lft_inv_alive lft_vs_inv lft_vs lft_inv_dead lft_inv lft_incl lft_inh lft_inv_alive lft_vs_inv lft_vs lft_inv_dead lft_inv lft_incl
idx_bor_own idx_bor raw_bor bor. idx_bor_own idx_bor raw_bor bor.
Section basic_properties. Section basic_properties.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft. Implicit Types κ : lft.
(* Unfolding lemmas *) (* Unfolding lemmas *)
...@@ -211,7 +236,7 @@ Lemma lft_vs_inv_go_ne κ (f f' : ∀ κ', κ' ⊂ κ → iProp Σ) I n : ...@@ -211,7 +236,7 @@ Lemma lft_vs_inv_go_ne κ (f f' : ∀ κ', κ' ⊂ κ → iProp Σ) I n :
( κ' ( : κ' κ), f κ' {n} f' κ' ) ( κ' ( : κ' κ), f κ' {n} f' κ' )
lft_vs_inv_go κ f I {n} lft_vs_inv_go κ f' I. lft_vs_inv_go κ f I {n} lft_vs_inv_go κ f' I.
Proof. Proof.
intros Hf. apply sep_ne, sep_ne, big_opS_ne=> // κ'. intros Hf. apply sep_ne, big_opS_ne=> // κ' _.
by apply forall_ne=> . by apply forall_ne=> .
Qed. Qed.
...@@ -240,9 +265,9 @@ Proof. ...@@ -240,9 +265,9 @@ Proof.
apply (Fix_F_proper _ (λ _, dist n)); auto using lft_inv_alive_go_ne. apply (Fix_F_proper _ (λ _, dist n)); auto using lft_inv_alive_go_ne.
Qed. Qed.
Lemma lft_vs_inv_unfold κ (I : gmap lft lft_names) : Lemma lft_vs_inv_unfold κ (I : gmap lft lft_names) :
lft_vs_inv κ I ⊣⊢ lft_bor_dead κ lft_vs_inv κ I ⊣⊢
own_ilft_auth I own_ilft_auth I
[ set] κ' dom _ I, κ' κ lft_inv_alive κ'. [ set] κ' dom I, κ' κ lft_inv_alive κ'.
Proof. Proof.
rewrite /lft_vs_inv /lft_vs_inv_go. by setoid_rewrite pure_impl_forall. rewrite /lft_vs_inv /lft_vs_inv_go. by setoid_rewrite pure_impl_forall.
Qed. Qed.
...@@ -250,10 +275,17 @@ Lemma lft_vs_unfold κ Pb Pi : ...@@ -250,10 +275,17 @@ Lemma lft_vs_unfold κ Pb Pi :
lft_vs κ Pb Pi ⊣⊢ n : nat, lft_vs κ Pb Pi ⊣⊢ n : nat,
own_cnt κ ( n) own_cnt κ ( n)
I : gmap lft lft_names, I : gmap lft lft_names,
lft_vs_inv κ I -∗ Pb -∗ lft_dead κ ={⊤∖↑mgmtN}=∗ lft_vs_inv κ I -∗ Pb -∗ lft_dead κ ={userE borN}=∗
lft_vs_inv κ I Pi own_cnt κ ( n). lft_vs_inv κ I Pi own_cnt κ ( n).
Proof. done. Qed. Proof. done. Qed.
Global Instance own_bor_proper κ : Proper (() ==> ()) (own_bor κ).
Proof. solve_proper. Qed.
Global Instance own_cnt_proper κ : Proper (() ==> ()) (own_cnt κ).
Proof. solve_proper. Qed.
Global Instance own_inh_proper κ : Proper (() ==> ()) (own_inh κ).
Proof. solve_proper. Qed.
Global Instance lft_bor_alive_ne κ n : Proper (dist n ==> dist n) (lft_bor_alive κ). Global Instance lft_bor_alive_ne κ n : Proper (dist n ==> dist n) (lft_bor_alive κ).
Proof. solve_proper. Qed. Proof. solve_proper. Qed.
Global Instance lft_bor_alive_proper κ : Proper (() ==> ()) (lft_bor_alive κ). Global Instance lft_bor_alive_proper κ : Proper (() ==> ()) (lft_bor_alive κ).
...@@ -272,36 +304,57 @@ Proof. apply (ne_proper_2 _). Qed. ...@@ -272,36 +304,57 @@ Proof. apply (ne_proper_2 _). Qed.
Global Instance idx_bor_ne κ i n : Proper (dist n ==> dist n) (idx_bor κ i). Global Instance idx_bor_ne κ i n : Proper (dist n ==> dist n) (idx_bor κ i).
Proof. solve_proper. Qed. Proof. solve_proper. Qed.
Global Instance idx_bor_contractive κ i : Contractive (idx_bor κ i).
Proof. solve_contractive. Qed.
Global Instance idx_bor_proper κ i : Proper (() ==> ()) (idx_bor κ i). Global Instance idx_bor_proper κ i : Proper (() ==> ()) (idx_bor κ i).
Proof. apply (ne_proper _). Qed. Proof. apply (ne_proper _). Qed.
Global Instance raw_bor_ne i n : Proper (dist n ==> dist n) (raw_bor i). Global Instance raw_bor_ne i n : Proper (dist n ==> dist n) (raw_bor i).
Proof. solve_proper. Qed. Proof. solve_proper. Qed.
Global Instance raw_bor_contractive i : Contractive (raw_bor i).
Proof. solve_contractive. Qed.
Global Instance raw_bor_proper i : Proper (() ==> ()) (raw_bor i). Global Instance raw_bor_proper i : Proper (() ==> ()) (raw_bor i).
Proof. apply (ne_proper _). Qed. Proof. apply (ne_proper _). Qed.
Global Instance bor_ne κ n : Proper (dist n ==> dist n) (bor κ). Global Instance bor_ne κ n : Proper (dist n ==> dist n) (bor κ).
Proof. solve_proper. Qed. Proof. solve_proper. Qed.
Global Instance bor_contractive κ : Contractive (bor κ).
Proof. solve_contractive. Qed.
Global Instance bor_proper κ : Proper (() ==> ()) (bor κ). Global Instance bor_proper κ : Proper (() ==> ()) (bor κ).
Proof. apply (ne_proper _). Qed. Proof. apply (ne_proper _). Qed.
(*** PersistentP and TimelessP and instances *) (*** Persistent and Timeless and instances *)
Global Instance lft_tok_timeless q κ : TimelessP q.[κ]. Global Instance lft_tok_timeless q κ : Timeless q.[κ].
Proof. rewrite /lft_tok. apply _. Qed. Proof. rewrite /lft_tok. apply _. Qed.
Global Instance lft_dead_persistent κ : PersistentP [κ]. Global Instance lft_dead_persistent κ : Persistent [κ].
Proof. rewrite /lft_dead. apply _. Qed.
Global Instance lft_dead_timeless κ : Timeless [κ].
Proof. rewrite /lft_dead. apply _. Qed. Proof. rewrite /lft_dead. apply _. Qed.
Global Instance lft_dead_timeless κ : PersistentP [κ].
Proof. rewrite /lft_tok. apply _. Qed.
Global Instance lft_incl_persistent κ κ' : PersistentP (κ κ'). Global Instance lft_incl_persistent κ κ' : Persistent (κ κ').
Proof. rewrite /lft_incl. apply _. Qed. Proof. rewrite /lft_incl. apply _. Qed.
Global Instance idx_bor_persistent κ i P : PersistentP (&{κ,i} P). Global Instance idx_bor_persistent κ i P : Persistent (&{κ,i} P).
Proof. rewrite /idx_bor. apply _. Qed. Proof. rewrite /idx_bor. apply _. Qed.
Global Instance idx_bor_own_timeless q i : TimelessP (idx_bor_own q i). Global Instance idx_bor_own_timeless q i : Timeless (idx_bor_own q i).
Proof. rewrite /idx_bor_own. apply _. Qed. Proof. rewrite /idx_bor_own. apply _. Qed.
Global Instance lft_ctx_persistent : PersistentP lft_ctx. Global Instance lft_ctx_persistent : Persistent lft_ctx.
Proof. rewrite /lft_ctx. apply _. Qed. Proof. rewrite /lft_ctx. apply _. Qed.
Global Instance atomic_lft_to_lft_inj : Inj eq eq atomic_lft_to_lft.
Proof.
unfold atomic_lft_to_lft. intros x y Hxy.
apply gmultiset_elem_of_singleton. rewrite -Hxy.
set_solver.
Qed.
Global Instance positive_to_atomic_lft_inj : Inj eq eq positive_to_atomic_lft.
Proof. apply _. Qed.
Global Definition atomic_lft_inhabited : Inhabited atomic_lft := _.
Global Definition atomic_lft_eq_dec : EqDecision atomic_lft := _.
Global Definition atomic_lft_countable : Countable atomic_lft := _.
End basic_properties. End basic_properties.
From lrust.lifetime Require Export primitive. From lrust.lifetime Require Export primitive.
From iris.algebra Require Import csum auth frac gmap dec_agree gset. From iris.algebra Require Import csum auth frac gmap agree gset numbers.
From iris.base_logic Require Import big_op.
From iris.base_logic.lib Require Import boxes. From iris.base_logic.lib Require Import boxes.
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Section faking. Section faking.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft. Implicit Types κ : lft.
Lemma ilft_create A I κ : Lemma ilft_create A I κ :
own_alft_auth A -∗ own_ilft_auth I -∗ ([ set] κ dom _ I, lft_inv A κ) own_alft_auth A -∗ own_ilft_auth I -∗ ([ set] κ dom I, lft_inv A κ)
==∗ A' I', is_Some (I' !! κ) ==∗ A' I', is_Some (I' !! κ)
own_alft_auth A' own_ilft_auth I' ([ set] κ dom _ I', lft_inv A' κ). own_alft_auth A' own_ilft_auth I' ([ set] κ dom I', lft_inv A' κ).
Proof. Proof.
iIntros "HA HI Hinv". iIntros "HA HI Hinv".
destruct (decide (is_Some (I !! κ))) as [?|HIκ%eq_None_not_Some]. destruct (decide (is_Some (I !! κ))) as [?|HIκ%eq_None_not_Some].
{ iModIntro. iExists A, I. by iFrame. } { iModIntro. iExists A, I. by iFrame. }
iMod (own_alloc ( 0 0)) as (γcnt) "[Hcnt Hcnt']"; first done. iMod (own_alloc ( 0 0)) as (γcnt) "[Hcnt Hcnt']"; first by apply auth_both_valid_discrete.
iMod (own_alloc (( ) :auth (gmap slice_name iMod (own_alloc ( ( : gmap slice_name
(frac * dec_agree bor_state)))) as (γbor) "[Hbor Hbor']"; (frac * agree bor_stateO)) )) as (γbor) "[Hbor Hbor']".
first by apply auth_valid_discrete_2. { by apply auth_both_valid_discrete. }
iMod (own_alloc (( ) :auth (gset_disj slice_name))) iMod (own_alloc ( (ε : gset_disj slice_name)))
as (γinh) "Hinh"; first by done. as (γinh) "Hinh"; first by apply auth_auth_valid.
set (γs := LftNames γbor γcnt γinh). set (γs := LftNames γbor γcnt γinh).
iMod (own_update with "HI") as "[HI Hγs]". iMod (own_update with "HI") as "[HI Hγs]".
{ apply auth_update_alloc, { apply auth_update_alloc,
(alloc_singleton_local_update _ κ (DecAgree γs)); last done. (alloc_singleton_local_update _ κ (to_agree γs)); last done.
by rewrite lookup_fmap HIκ. } by rewrite lookup_fmap HIκ. }
iDestruct "Hγs" as "#Hγs". iDestruct "Hγs" as "#Hγs".
iAssert (own_cnt κ ( 0)) with "[Hcnt]" as "Hcnt". iAssert (own_cnt κ ( 0)) with "[Hcnt]" as "Hcnt".
...@@ -33,15 +33,17 @@ Proof. ...@@ -33,15 +33,17 @@ Proof.
iAssert (own_cnt κ ( 0)) with "[Hcnt']" as "Hcnt'". iAssert (own_cnt κ ( 0)) with "[Hcnt']" as "Hcnt'".
{ rewrite /own_cnt. iExists γs. by iFrame. } { rewrite /own_cnt. iExists γs. by iFrame. }
iAssert ( b, lft_inh κ b True)%I with "[Hinh]" as "Hinh". iAssert ( b, lft_inh κ b True)%I with "[Hinh]" as "Hinh".
{ iIntros (b). rewrite /lft_inh. iExists ∅. rewrite to_gmap_empty. { iIntros (b). rewrite /lft_inh. iExists ∅. rewrite gset_to_gmap_empty.
iSplitL; [|iApply box_alloc]. rewrite /own_inh. iExists γs. by iFrame. } iSplitL; [|iApply box_alloc]. rewrite /own_inh. iExists γs. by iFrame. }
iAssert (lft_inv_dead κ lft_inv_alive κ)%I iAssert (lft_inv_dead κ lft_inv_alive κ)%I
with "[-HA HI Hinv]" as "Hdeadandalive". with "[-HA HI Hinv]" as "Hdeadandalive".
{ iSplit. { iSplit.
- rewrite /lft_inv_dead. iExists True%I. iFrame "Hcnt". - rewrite /lft_inv_dead. iExists True%I. iFrame "Hcnt".
iSplitL "Hbor"; last by iApply "Hinh". iSplitL "Hbor"; last by iApply "Hinh".
rewrite /lft_bor_dead. iExists , True%I. rewrite !to_gmap_empty. rewrite /lft_bor_dead. iExists , True%I. rewrite !gset_to_gmap_empty.
iSplitL "Hbor". iExists γs. by iFrame. iApply box_alloc. iSplitL "Hbor".
+ iExists γs. by iFrame.
+ iApply box_alloc.
- rewrite lft_inv_alive_unfold. iExists True%I, True%I. iSplitL "Hbor". - rewrite lft_inv_alive_unfold. iExists True%I, True%I. iSplitL "Hbor".
{ rewrite /lft_bor_alive. iExists ∅. { rewrite /lft_bor_alive. iExists ∅.
rewrite /to_borUR !fmap_empty big_sepM_empty. rewrite /to_borUR !fmap_empty big_sepM_empty.
...@@ -63,7 +65,7 @@ Proof. ...@@ -63,7 +65,7 @@ Proof.
{ by iDestruct "Hdeadandalive" as "[? _]". } { by iDestruct "Hdeadandalive" as "[? _]". }
iPureIntro. exists Λ. rewrite lookup_insert; auto. } iPureIntro. exists Λ. rewrite lookup_insert; auto. }
iNext. iApply (@big_sepS_impl with "[$Hinv]"). iNext. iApply (@big_sepS_impl with "[$Hinv]").
rewrite /lft_inv. iIntros "!#"; iIntros (κ' ?%elem_of_dom) rewrite /lft_inv. iIntros "!>"; iIntros (κ' ?%elem_of_dom)
"[[HA HA']|[HA HA']]"; iDestruct "HA'" as %HA. "[[HA HA']|[HA HA']]"; iDestruct "HA'" as %HA.
+ iLeft. iFrame "HA". iPureIntro. by apply lft_alive_in_insert. + iLeft. iFrame "HA". iPureIntro. by apply lft_alive_in_insert.
+ iRight. iFrame "HA". iPureIntro. by apply lft_dead_in_insert. + iRight. iFrame "HA". iPureIntro. by apply lft_dead_in_insert.
...@@ -77,19 +79,18 @@ Proof. ...@@ -77,19 +79,18 @@ Proof.
+ iRight. by iDestruct "Hdeadandalive" as "[$ _]". + iRight. by iDestruct "Hdeadandalive" as "[$ _]".
Qed. Qed.
Lemma raw_bor_fake E q κ P : Lemma raw_bor_fake E κ P :
borN E borN E lft_bor_dead κ ={E}=∗ lft_bor_dead κ raw_bor κ P.
?q lft_bor_dead κ ={E}=∗ ?q lft_bor_dead κ raw_bor κ P.
Proof. Proof.
iIntros (?). rewrite /lft_bor_dead. iDestruct 1 as (B Pinh) "[>HB● Hbox]". iIntros (?). rewrite /lft_bor_dead. iDestruct 1 as (B Pinh) "[>HB● Hbox]".
iMod (slice_insert_empty _ _ _ _ P with "Hbox") as (γ) "(% & #Hslice & Hbox)". iMod (slice_insert_empty _ _ true _ P with "Hbox") as (γ) "(% & #Hslice & Hbox)".
iMod (own_bor_update with "HB●") as "[HB● H◯]". iMod (own_bor_update with "HB●") as "[HB● H◯]".
{ eapply auth_update_alloc, { eapply auth_update_alloc,
(alloc_singleton_local_update _ _ (1%Qp, DecAgree Bor_in)); last done. (alloc_singleton_local_update _ _ (1%Qp, to_agree Bor_in)); last done.
by do 2 eapply lookup_to_gmap_None. } by do 2 eapply lookup_gset_to_gmap_None. }
rewrite /bor /raw_bor /idx_bor_own /=. iModIntro. iSplitR "H◯". rewrite /bor /raw_bor /idx_bor_own /=. iModIntro. iSplitR "H◯".
- iExists ({[γ]} B), (P Pinh)%I. rewrite !to_gmap_union_singleton. by iFrame. - iExists ({[γ]} B), (P Pinh)%I. rewrite !gset_to_gmap_union_singleton. by iFrame.
- iExists γ. by iFrame. - iExists γ. iFrame. iExists P. rewrite -(bi.iff_refl emp). auto.
Qed. Qed.
Lemma raw_bor_fake' E κ P : Lemma raw_bor_fake' E κ P :
...@@ -105,14 +106,13 @@ Proof. ...@@ -105,14 +106,13 @@ Proof.
rewrite {1}/lft_inv; iDestruct "Hinv" as "[[_ >%]|[Hinv >%]]". rewrite {1}/lft_inv; iDestruct "Hinv" as "[[_ >%]|[Hinv >%]]".
{ unfold lft_alive_in in *; naive_solver. } { unfold lft_alive_in in *; naive_solver. }
rewrite /lft_inv_dead; iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)". rewrite /lft_inv_dead; iDestruct "Hinv" as (Pinh) "(Hdead & Hcnt & Hinh)".
iMod (raw_bor_fake _ true _ P with "Hdead") as "[Hdead Hbor]"; first solve_ndisj. iMod (raw_bor_fake _ _ P with "Hdead") as "[Hdead Hbor]"; first solve_ndisj.
iFrame. iApply "Hclose". iExists A', I'. iFrame. iNext. iApply "Hclose'". iFrame. iApply "Hclose". iExists A', I'. iFrame. iNext. iApply "Hclose'".
rewrite /lft_inv /lft_inv_dead. iRight. iFrame. eauto. rewrite /lft_inv /lft_inv_dead. iRight. iFrame. eauto.
Qed. Qed.
Lemma bor_fake E κ P : Lemma bor_fake E κ P :
lftN E lftN E lft_ctx -∗ [κ] ={E}=∗ &{κ}P.
lft_ctx -∗ [κ] ={E}=∗ &{κ}P.
Proof. Proof.
iIntros (?) "#LFT H†". iMod (raw_bor_fake' with "LFT H†"); first done. iIntros (?) "#LFT H†". iMod (raw_bor_fake' with "LFT H†"); first done.
iModIntro. unfold bor. iExists κ. iFrame. iApply lft_incl_refl. iModIntro. unfold bor. iExists κ. iFrame. iApply lft_incl_refl.
......
From lrust.lifetime Require Export definitions. From lrust.lifetime.model Require Export definitions.
From iris.algebra Require Import csum auth frac gmap dec_agree gset. From iris.algebra Require Import csum auth frac gmap agree gset proofmode_classes.
From iris.base_logic Require Import big_op. From iris.base_logic.lib Require Import boxes.
From iris.base_logic.lib Require Import boxes fractional. From iris.bi Require Import fractional.
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Import uPred. Import uPred.
Lemma lft_init `{!invGS Σ, !lftGpreS Σ} E userE :
lftN E lftN ## userE |={E}=> _ : lftGS Σ userE, lft_ctx.
Proof.
iIntros (? HuserE). rewrite /lft_ctx.
iMod (own_alloc ( : authR alftUR)) as (γa) "Ha"; first by apply auth_auth_valid.
iMod (own_alloc ( : authR ilftUR)) as (γi) "Hi"; first by apply auth_auth_valid.
set (HlftGS := LftG _ _ _ _ γa _ γi _ _ _ HuserE). iExists HlftGS.
iMod (inv_alloc _ _ lfts_inv with "[Ha Hi]") as "$"; last done.
iModIntro. rewrite /lfts_inv /own_alft_auth /own_ilft_auth. iExists , ∅.
rewrite /to_alftUR /to_ilftUR !fmap_empty. iFrame.
rewrite dom_empty_L big_sepS_empty. done.
Qed.
Section primitive. Section primitive.
Context `{invG Σ, lftG Σ}. Context `{!invGS Σ, !lftGS Σ userE}.
Implicit Types κ : lft. Implicit Types κ : lft.
Lemma to_borUR_included (B : gmap slice_name bor_state) i s q : Lemma to_borUR_included (B : gmap slice_name bor_state) i s q :
{[i := (q%Qp, DecAgree s)]} to_borUR B B !! i = Some s. {[i := (q%Qp, to_agree s)]} to_borUR B B !! i = Some s.
Proof. Proof.
rewrite singleton_included=> -[qs [/leibniz_equiv_iff]]. rewrite singleton_included_l=> -[qs []]. unfold_leibniz.
rewrite lookup_fmap fmap_Some=> -[s' [? ->]]. rewrite lookup_fmap fmap_Some_equiv=> -[s' [-> ->]].
by move=> /Some_pair_included [_] /Some_included_total /DecAgree_included=>->. by move=> /Some_pair_included [_] /Some_included_total /to_agree_included=>->.
Qed. Qed.
(** Ownership *) (** Ownership *)
Lemma own_ilft_auth_agree (I : gmap lft lft_names) κ γs : Lemma own_ilft_auth_agree (I : gmap lft lft_names) κ γs :
own_ilft_auth I -∗ own_ilft_auth I -∗
own ilft_name ( {[κ := DecAgree γs]}) -∗ is_Some (I !! κ)⌝. own ilft_name ( {[κ := to_agree γs]}) -∗ is_Some (I !! κ)⌝.
Proof. Proof.
iIntros "HI Hκ". iDestruct (own_valid_2 with "HI Hκ") iIntros "HI Hκ". iDestruct (own_valid_2 with "HI Hκ")
as %[[? [??]]%singleton_included _]%auth_valid_discrete_2. as %[[? [Hl ?]]%singleton_included_l _]%auth_both_valid_discrete.
unfold to_ilftUR in *. simplify_map_eq; simplify_option_eq; eauto. revert Hl. rewrite /to_ilftUR lookup_fmap.
case: (I !! _)=> [γs'|] Hl; first by eauto.
destruct (fmap_Some_equiv_1 _ _ _ Hl) as (?&?&?). eauto.
Qed. Qed.
Lemma own_alft_auth_agree (A : gmap atomic_lft bool) Λ b : Lemma own_alft_auth_agree (A : gmap atomic_lft bool) Λ b :
...@@ -32,8 +48,8 @@ Lemma own_alft_auth_agree (A : gmap atomic_lft bool) Λ b : ...@@ -32,8 +48,8 @@ Lemma own_alft_auth_agree (A : gmap atomic_lft bool) Λ b :
own alft_name ( {[Λ := to_lft_stateR b]}) -∗ A !! Λ = Some b⌝. own alft_name ( {[Λ := to_lft_stateR b]}) -∗ A !! Λ = Some b⌝.
Proof. Proof.
iIntros "HA HΛ". iIntros "HA HΛ".
iDestruct (own_valid_2 with "HA HΛ") as %[HA _]%auth_valid_discrete_2. iCombine "HA HΛ" gives %[HA _]%auth_both_valid_discrete.
iPureIntro. move: HA=> /singleton_included [qs [/leibniz_equiv_iff]]. iPureIntro. move: HA=> /singleton_included_l [qs [/leibniz_equiv_iff]].
rewrite lookup_fmap fmap_Some=> -[b' [? ->]] /Some_included. rewrite lookup_fmap fmap_Some=> -[b' [? ->]] /Some_included.
move=> [/leibniz_equiv_iff|/csum_included]; destruct b, b'; naive_solver. move=> [/leibniz_equiv_iff|/csum_included]; destruct b, b'; naive_solver.
Qed. Qed.
...@@ -43,25 +59,33 @@ Proof. ...@@ -43,25 +59,33 @@ Proof.
iIntros "HI"; iDestruct 1 as (γs) "[? _]". iIntros "HI"; iDestruct 1 as (γs) "[? _]".
by iApply (own_ilft_auth_agree with "HI"). by iApply (own_ilft_auth_agree with "HI").
Qed. Qed.
Lemma own_bor_op κ x y : own_bor κ (x y) ⊣⊢ own_bor κ x own_bor κ y. Lemma own_bor_op κ x y : own_bor κ (x y) ⊣⊢ own_bor κ x own_bor κ y.
Proof. Proof.
iSplit. iSplit.
{ iDestruct 1 as (γs) "[#? [Hx Hy]]"; iSplitL "Hx"; iExists γs; eauto. } { iDestruct 1 as (γs) "[#? [Hx Hy]]"; iSplitL "Hx"; iExists γs; eauto. }
iIntros "[Hx Hy]". iIntros "[Hx Hy]".
iDestruct "Hx" as (γs) "[Hγs Hx]"; iDestruct "Hy" as (γs') "[Hγs' Hy]". iDestruct "Hx" as (γs) "[Hγs Hx]"; iDestruct "Hy" as (γs') "[Hγs' Hy]".
iDestruct (own_valid_2 with "Hγs Hγs'") as %Hγs%auth_own_valid. iCombine "Hγs Hγs'" gives %Hγs. move : Hγs.
move: Hγs; rewrite /= op_singleton singleton_valid=> /dec_agree_op_inv [<-]. rewrite -auth_frag_op auth_frag_valid singleton_op singleton_valid=> /to_agree_op_inv_L <-.
iExists γs. iSplit. done. rewrite own_op; iFrame. iExists γs. iSplit; first done. rewrite own_op; iFrame.
Qed. Qed.
Global Instance own_bor_into_op κ x x1 x2 : Global Instance own_bor_into_op κ x x1 x2 :
IntoOp x x1 x2 IntoAnd false (own_bor κ x) (own_bor κ x1) (own_bor κ x2). IsOp x x1 x2 IntoSep (own_bor κ x) (own_bor κ x1) (own_bor κ x2).
Proof. Proof.
rewrite /IntoOp /IntoAnd=> /leibniz_equiv_iff->. by rewrite -own_bor_op. rewrite /IsOp /IntoSep=>-> /=. by rewrite -own_bor_op.
Qed. Qed.
Lemma own_bor_valid κ x : own_bor κ x -∗ x. Lemma own_bor_valid κ x : own_bor κ x -∗ x.
Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed. Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed.
Lemma own_bor_valid_2 κ x y : own_bor κ x -∗ own_bor κ y -∗ (x y). Lemma own_bor_valid_2 κ x y : own_bor κ x -∗ own_bor κ y -∗ (x y).
Proof. apply wand_intro_r. rewrite -own_bor_op. apply own_bor_valid. Qed. Proof. apply entails_wand, wand_intro_r. rewrite -own_bor_op. iApply own_bor_valid. Qed.
Global Instance own_bor_sep_gives κ x y :
CombineSepGives (own_bor κ x) (own_bor κ y) ( (x y)).
Proof.
rewrite /CombineSepGives. iIntros "[H1 H2]".
iDestruct (own_bor_valid_2 with "H1 H2") as "#?"; auto.
Qed.
Lemma own_bor_update κ x y : x ~~> y own_bor κ x ==∗ own_bor κ y. Lemma own_bor_update κ x y : x ~~> y own_bor κ x ==∗ own_bor κ y.
Proof. Proof.
iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update. iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update.
...@@ -69,7 +93,7 @@ Qed. ...@@ -69,7 +93,7 @@ Qed.
Lemma own_bor_update_2 κ x1 x2 y : Lemma own_bor_update_2 κ x1 x2 y :
x1 x2 ~~> y own_bor κ x1 own_bor κ x2 ==∗ own_bor κ y. x1 x2 ~~> y own_bor κ x1 own_bor κ x2 ==∗ own_bor κ y.
Proof. Proof.
intros. apply wand_intro_r. rewrite -own_bor_op. by apply own_bor_update. intros. apply wand_intro_r. rewrite -own_bor_op. by iApply own_bor_update.
Qed. Qed.
Lemma own_cnt_auth I κ x : own_ilft_auth I -∗ own_cnt κ x -∗ is_Some (I !! κ)⌝. Lemma own_cnt_auth I κ x : own_ilft_auth I -∗ own_cnt κ x -∗ is_Some (I !! κ)⌝.
...@@ -83,19 +107,25 @@ Proof. ...@@ -83,19 +107,25 @@ Proof.
{ iDestruct 1 as (γs) "[#? [Hx Hy]]"; iSplitL "Hx"; iExists γs; eauto. } { iDestruct 1 as (γs) "[#? [Hx Hy]]"; iSplitL "Hx"; iExists γs; eauto. }
iIntros "[Hx Hy]". iIntros "[Hx Hy]".
iDestruct "Hx" as (γs) "[Hγs Hx]"; iDestruct "Hy" as (γs') "[Hγs' Hy]". iDestruct "Hx" as (γs) "[Hγs Hx]"; iDestruct "Hy" as (γs') "[Hγs' Hy]".
iDestruct (own_valid_2 with "Hγs Hγs'") as %Hγs%auth_own_valid. iCombine "Hγs Hγs'" gives %Hγs. move: Hγs.
move: Hγs; rewrite /= op_singleton singleton_valid=> /dec_agree_op_inv [<-]. rewrite -auth_frag_op auth_frag_valid singleton_op singleton_valid=> /to_agree_op_inv_L=> <-.
iExists γs. iSplit; first done. rewrite own_op; iFrame. iExists γs. iSplit; first done. rewrite own_op; iFrame.
Qed. Qed.
Global Instance own_cnt_into_op κ x x1 x2 : Global Instance own_cnt_into_op κ x x1 x2 :
IntoOp x x1 x2 IntoAnd false (own_cnt κ x) (own_cnt κ x1) (own_cnt κ x2). IsOp x x1 x2 IntoSep (own_cnt κ x) (own_cnt κ x1) (own_cnt κ x2).
Proof. Proof.
rewrite /IntoOp /IntoAnd=> /leibniz_equiv_iff->. by rewrite -own_cnt_op. rewrite /IsOp /IntoSep=> ->. by rewrite -own_cnt_op.
Qed. Qed.
Lemma own_cnt_valid κ x : own_cnt κ x -∗ x. Lemma own_cnt_valid κ x : own_cnt κ x -∗ x.
Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed. Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed.
Lemma own_cnt_valid_2 κ x y : own_cnt κ x -∗ own_cnt κ y -∗ (x y). Lemma own_cnt_valid_2 κ x y : own_cnt κ x -∗ own_cnt κ y -∗ (x y).
Proof. apply wand_intro_r. rewrite -own_cnt_op. apply own_cnt_valid. Qed. Proof. apply entails_wand, wand_intro_r. rewrite -own_cnt_op. iApply own_cnt_valid. Qed.
Global Instance own_cnt_sep_gives κ x y :
CombineSepGives (own_cnt κ x) (own_cnt κ y) ( (x y)).
Proof.
rewrite /CombineSepGives. iIntros "[H1 H2]".
iDestruct (own_cnt_valid_2 with "H1 H2") as "#?"; auto.
Qed.
Lemma own_cnt_update κ x y : x ~~> y own_cnt κ x ==∗ own_cnt κ y. Lemma own_cnt_update κ x y : x ~~> y own_cnt κ x ==∗ own_cnt κ y.
Proof. Proof.
iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update. iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update.
...@@ -103,7 +133,7 @@ Qed. ...@@ -103,7 +133,7 @@ Qed.
Lemma own_cnt_update_2 κ x1 x2 y : Lemma own_cnt_update_2 κ x1 x2 y :
x1 x2 ~~> y own_cnt κ x1 -∗ own_cnt κ x2 ==∗ own_cnt κ y. x1 x2 ~~> y own_cnt κ x1 -∗ own_cnt κ x2 ==∗ own_cnt κ y.
Proof. Proof.
intros. apply wand_intro_r. rewrite -own_cnt_op. by apply own_cnt_update. intros. apply entails_wand, wand_intro_r. rewrite -own_cnt_op. by iApply own_cnt_update.
Qed. Qed.
Lemma own_inh_auth I κ x : own_ilft_auth I -∗ own_inh κ x -∗ is_Some (I !! κ)⌝. Lemma own_inh_auth I κ x : own_ilft_auth I -∗ own_inh κ x -∗ is_Some (I !! κ)⌝.
...@@ -117,19 +147,25 @@ Proof. ...@@ -117,19 +147,25 @@ Proof.
{ iDestruct 1 as (γs) "[#? [Hx Hy]]"; iSplitL "Hx"; iExists γs; eauto. } { iDestruct 1 as (γs) "[#? [Hx Hy]]"; iSplitL "Hx"; iExists γs; eauto. }
iIntros "[Hx Hy]". iIntros "[Hx Hy]".
iDestruct "Hx" as (γs) "[Hγs Hx]"; iDestruct "Hy" as (γs') "[Hγs' Hy]". iDestruct "Hx" as (γs) "[Hγs Hx]"; iDestruct "Hy" as (γs') "[Hγs' Hy]".
iDestruct (own_valid_2 with "Hγs Hγs'") as %Hγs%auth_own_valid. iCombine "Hγs Hγs'" gives %Hγs. move: Hγs.
move: Hγs; rewrite /= op_singleton singleton_valid=> /dec_agree_op_inv [<-]. rewrite -auth_frag_op auth_frag_valid singleton_op singleton_valid=> /to_agree_op_inv_L=> <-.
iExists γs. iSplit. done. rewrite own_op; iFrame. iExists γs. iSplit; first done. rewrite own_op; iFrame.
Qed. Qed.
Global Instance own_inh_into_op κ x x1 x2 : Global Instance own_inh_into_op κ x x1 x2 :
IntoOp x x1 x2 IntoAnd false (own_inh κ x) (own_inh κ x1) (own_inh κ x2). IsOp x x1 x2 IntoSep (own_inh κ x) (own_inh κ x1) (own_inh κ x2).
Proof. Proof.
rewrite /IntoOp /IntoAnd=> /leibniz_equiv_iff->. by rewrite -own_inh_op. rewrite /IsOp /IntoSep=> ->. by rewrite -own_inh_op.
Qed. Qed.
Lemma own_inh_valid κ x : own_inh κ x -∗ x. Lemma own_inh_valid κ x : own_inh κ x -∗ x.
Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed. Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed.
Lemma own_inh_valid_2 κ x y : own_inh κ x -∗ own_inh κ y -∗ (x y). Lemma own_inh_valid_2 κ x y : own_inh κ x -∗ own_inh κ y -∗ (x y).
Proof. apply wand_intro_r. rewrite -own_inh_op. apply own_inh_valid. Qed. Proof. apply entails_wand, wand_intro_r. rewrite -own_inh_op. iApply own_inh_valid. Qed.
Global Instance own_inh_sep_gives κ x y :
CombineSepGives (own_inh κ x) (own_inh κ y) ( (x y)).
Proof.
rewrite /CombineSepGives. iIntros "[H1 H2]".
iDestruct (own_inh_valid_2 with "H1 H2") as "#?"; auto.
Qed.
Lemma own_inh_update κ x y : x ~~> y own_inh κ x ==∗ own_inh κ y. Lemma own_inh_update κ x y : x ~~> y own_inh κ x ==∗ own_inh κ y.
Proof. Proof.
iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update. iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update.
...@@ -137,20 +173,20 @@ Qed. ...@@ -137,20 +173,20 @@ Qed.
Lemma own_inh_update_2 κ x1 x2 y : Lemma own_inh_update_2 κ x1 x2 y :
x1 x2 ~~> y own_inh κ x1 own_inh κ x2 ==∗ own_inh κ y. x1 x2 ~~> y own_inh κ x1 own_inh κ x2 ==∗ own_inh κ y.
Proof. Proof.
intros. apply wand_intro_r. rewrite -own_inh_op. by apply own_inh_update. intros. apply wand_intro_r. rewrite -own_inh_op. by iApply own_inh_update.
Qed. Qed.
(** Alive and dead *) (** Alive and dead *)
Global Instance lft_alive_in_dec A κ : Decision (lft_alive_in A κ). Global Instance lft_alive_in_dec A κ : Decision (lft_alive_in A κ).
Proof. Proof.
refine (cast_if (decide (set_Forall (λ Λ, A !! Λ = Some true) refine (cast_if (decide (set_Forall (λ Λ, A !! Λ = Some true)
(dom (gset atomic_lft) κ)))); (dom κ))));
rewrite /lft_alive_in; by setoid_rewrite <-gmultiset_elem_of_dom. rewrite /lft_alive_in; by setoid_rewrite <-gmultiset_elem_of_dom.
Qed. Qed.
Global Instance lft_dead_in_dec A κ : Decision (lft_dead_in A κ). Global Instance lft_dead_in_dec A κ : Decision (lft_dead_in A κ).
Proof. Proof.
refine (cast_if (decide (set_Exists (λ Λ, A !! Λ = Some false) refine (cast_if (decide (set_Exists (λ Λ, A !! Λ = Some false)
(dom (gset atomic_lft) κ)))); (dom κ))));
rewrite /lft_dead_in; by setoid_rewrite <-gmultiset_elem_of_dom. rewrite /lft_dead_in; by setoid_rewrite <-gmultiset_elem_of_dom.
Qed. Qed.
...@@ -158,9 +194,9 @@ Lemma lft_alive_or_dead_in A κ : ...@@ -158,9 +194,9 @@ Lemma lft_alive_or_dead_in A κ :
( Λ, Λ κ A !! Λ = None) (lft_alive_in A κ lft_dead_in A κ). ( Λ, Λ κ A !! Λ = None) (lft_alive_in A κ lft_dead_in A κ).
Proof. Proof.
rewrite /lft_alive_in /lft_dead_in. rewrite /lft_alive_in /lft_dead_in.
destruct (decide (set_Exists (λ Λ, A !! Λ = None) (dom (gset _) κ))) destruct (decide (set_Exists (λ Λ, A !! Λ = None) (dom κ)))
as [(Λ & ?%gmultiset_elem_of_dom & HAΛ)|HA%(not_set_Exists_Forall _)]; first eauto. as [(Λ & ?%gmultiset_elem_of_dom & HAΛ)|HA%(not_set_Exists_Forall _)]; first eauto.
destruct (decide (set_Exists (λ Λ, A !! Λ = Some false) (dom (gset _) κ))) destruct (decide (set_Exists (λ Λ, A !! Λ = Some false) (dom κ)))
as [(Λ & %gmultiset_elem_of_dom & ?)|HA'%(not_set_Exists_Forall _)]; first eauto. as [(Λ & %gmultiset_elem_of_dom & ?)|HA'%(not_set_Exists_Forall _)]; first eauto.
right; left. intros Λ %gmultiset_elem_of_dom. right; left. intros Λ %gmultiset_elem_of_dom.
move: (HA _ ) (HA' _ )=> /=. case: (A !! Λ)=>[[]|]; naive_solver. move: (HA _ ) (HA' _ )=> /=. case: (A !! Λ)=>[[]|]; naive_solver.
...@@ -185,6 +221,14 @@ Proof. ...@@ -185,6 +221,14 @@ Proof.
intros Halive Λ' HΛ'. intros Halive Λ' HΛ'.
rewrite lookup_insert_ne; last (by intros ->); auto. rewrite lookup_insert_ne; last (by intros ->); auto.
Qed. Qed.
Lemma lft_alive_in_union_false A κ Λs b :
( Λ, Λ Λs Λ κ) lft_alive_in A κ lft_alive_in (gset_to_gmap b Λs A) κ.
Proof.
intros Halive Λ' HΛ'.
rewrite lookup_union_r; [by apply Halive|].
apply lookup_gset_to_gmap_None.
by intros ?%.
Qed.
Lemma lft_dead_in_insert A κ Λ b : Lemma lft_dead_in_insert A κ Λ b :
A !! Λ = None lft_dead_in A κ lft_dead_in (<[Λ:=b]> A) κ. A !! Λ = None lft_dead_in A κ lft_dead_in (<[Λ:=b]> A) κ.
...@@ -200,12 +244,32 @@ Proof. ...@@ -200,12 +244,32 @@ Proof.
- exists Λ. by rewrite lookup_insert. - exists Λ. by rewrite lookup_insert.
- exists Λ'. by rewrite lookup_insert_ne. - exists Λ'. by rewrite lookup_insert_ne.
Qed. Qed.
Lemma lft_dead_in_union_false A κ Λs :
lft_dead_in A κ lft_dead_in (gset_to_gmap false Λs A) κ.
Proof.
intros (Λ&?&). destruct (decide (Λ Λs)) as [|].
- exists Λ. rewrite lookup_union_l'.
+ by rewrite lookup_gset_to_gmap_Some.
+ exists false; by rewrite lookup_gset_to_gmap_Some.
- exists Λ. rewrite lookup_union_r.
+ done.
+ by rewrite lookup_gset_to_gmap_None.
Qed.
Lemma lft_dead_in_insert_false' A κ Λ : Λ κ lft_dead_in (<[Λ:=false]> A) κ. Lemma lft_dead_in_insert_false' A κ Λ : Λ κ lft_dead_in (<[Λ:=false]> A) κ.
Proof. exists Λ. by rewrite lookup_insert. Qed. Proof. exists Λ. by rewrite lookup_insert. Qed.
Lemma lft_dead_in_union_false' A κ Λs Λ :
Λ κ Λ Λs lft_dead_in (gset_to_gmap false Λs A) κ.
Proof.
intros HΛκ .
exists Λ; split; [done|].
rewrite lookup_union_l'.
- by apply lookup_gset_to_gmap_Some.
- exists false; by apply lookup_gset_to_gmap_Some.
Qed.
Lemma lft_dead_in_alive_in_union A κ κ' : Lemma lft_dead_in_alive_in_union A κ κ' :
lft_dead_in A (κ κ') lft_alive_in A κ lft_dead_in A κ'. lft_dead_in A (κ κ') lft_alive_in A κ lft_dead_in A κ'.
Proof. Proof.
intros (Λ & [Hin|Hin]%elem_of_union & ) Halive. intros (Λ & [Hin|Hin]%gmultiset_elem_of_disj_union & ) Halive.
- contradict . rewrite (Halive _ Hin). done. - contradict . rewrite (Halive _ Hin). done.
- exists Λ. auto. - exists Λ. auto.
Qed. Qed.
...@@ -226,14 +290,6 @@ Proof. ...@@ -226,14 +290,6 @@ Proof.
by rewrite HAinsert. by rewrite HAinsert.
Qed. Qed.
Lemma lft_inv_alive_twice κ : lft_inv_alive κ -∗ lft_inv_alive κ -∗ False.
Proof.
rewrite lft_inv_alive_unfold /lft_inh.
iDestruct 1 as (P Q) "(_&_&Hinh)"; iDestruct 1 as (P' Q') "(_&_&Hinh')".
iDestruct "Hinh" as (E) "[HE _]"; iDestruct "Hinh'" as (E') "[HE' _]".
by iDestruct (own_inh_valid_2 with "HE HE'") as %?.
Qed.
Lemma lft_inv_alive_in A κ : lft_alive_in A κ lft_inv A κ -∗ lft_inv_alive κ. Lemma lft_inv_alive_in A κ : lft_alive_in A κ lft_inv A κ -∗ lft_inv_alive κ.
Proof. Proof.
rewrite /lft_inv. iIntros (?) "[[$ _]|[_ %]]". rewrite /lft_inv. iIntros (?) "[[$ _]|[_ %]]".
...@@ -246,10 +302,23 @@ Proof. ...@@ -246,10 +302,23 @@ Proof.
Qed. Qed.
(** Basic rules about lifetimes *) (** Basic rules about lifetimes *)
Lemma lft_tok_sep q κ1 κ2 : q.[κ1] q.[κ2] ⊣⊢ q.[κ1 κ2]. Local Instance lft_inhabited : Inhabited lft := _.
Proof. by rewrite /lft_tok -big_sepMS_union. Qed. Local Instance lft_eq_dec : EqDecision lft := _.
Local Instance lft_countable : Countable lft := _.
Lemma lft_dead_or κ1 κ2 : [κ1] [κ2] ⊣⊢ [ κ1 κ2]. Local Instance bor_idx_inhabited : Inhabited bor_idx := _.
Local Instance bor_idx_eq_dec : EqDecision bor_idx := _.
Local Instance bor_idx_countable : Countable bor_idx := _.
Local Instance lft_intersect_comm : Comm (A:=lft) eq () := _.
Local Instance lft_intersect_assoc : Assoc (A:=lft) eq () := _.
Local Instance lft_intersect_inj_1 κ : Inj eq eq (κ ⊓.) := _.
Local Instance lft_intersect_inj_2 κ : Inj eq eq (. κ) := _.
Local Instance lft_intersect_left_id : LeftId eq static () := _.
Local Instance lft_intersect_right_id : RightId eq static () := _.
Lemma lft_tok_sep q κ1 κ2 : q.[κ1] q.[κ2] ⊣⊢ q.[κ1 κ2].
Proof. by rewrite /lft_tok -big_sepMS_disj_union. Qed.
Lemma lft_dead_or κ1 κ2 : [κ1] [κ2] ⊣⊢ [ κ1 κ2].
Proof. Proof.
rewrite /lft_dead -or_exist. apply exist_proper=> Λ. rewrite /lft_dead -or_exist. apply exist_proper=> Λ.
rewrite -sep_or_r -pure_or. do 2 f_equiv. set_solver. rewrite -sep_or_r -pure_or. do 2 f_equiv. set_solver.
...@@ -258,110 +327,176 @@ Qed. ...@@ -258,110 +327,176 @@ Qed.
Lemma lft_tok_dead q κ : q.[κ] -∗ [ κ] -∗ False. Lemma lft_tok_dead q κ : q.[κ] -∗ [ κ] -∗ False.
Proof. Proof.
rewrite /lft_tok /lft_dead. iIntros "H"; iDestruct 1 as (Λ') "[% H']". rewrite /lft_tok /lft_dead. iIntros "H"; iDestruct 1 as (Λ') "[% H']".
iDestruct (big_sepMS_elem_of _ _ Λ' with "H") as "H"=> //. iDestruct (@big_sepMS_elem_of _ _ _ _ _ _ Λ' with "H") as "H"=> //.
iDestruct (own_valid_2 with "H H'") as %Hvalid. iCombine "H H'" gives %Hvalid.
move: Hvalid=> /auth_own_valid /=; by rewrite op_singleton singleton_valid. move: Hvalid=> /auth_frag_valid /=; by rewrite singleton_op singleton_valid.
Qed. Qed.
Lemma lft_tok_static q : q.[static]%I. Lemma lft_tok_static q : q.[static].
Proof. by rewrite /lft_tok big_sepMS_empty. Qed. Proof. by rewrite /lft_tok big_sepMS_empty. Qed.
Lemma lft_dead_static : [ static] -∗ False. Lemma lft_dead_static : [ static] -∗ False.
Proof. rewrite /lft_dead. iDestruct 1 as (Λ) "[% H']". set_solver. Qed. Proof. rewrite /lft_dead. iDestruct 1 as (Λ) "[% H']". set_solver. Qed.
Lemma lft_intersect_static_cancel_l κ κ' : κ κ' = static κ = static.
Proof.
rewrite !gmultiset_eq=>Heq Λ. move:(Heq Λ).
rewrite multiplicity_disj_union multiplicity_empty Nat.eq_add_0.
naive_solver.
Qed.
Lemma lft_tok_valid q κ :
q.[κ] -∗ (q 1)%Qp κ = static⌝.
Proof.
rewrite /lft_tok.
destruct (decide (κ = static)) as [-> | Hneq]; first by eauto.
edestruct (gmultiset_choose _ Hneq) as (Λ & Hel).
iIntros "Hm". iPoseProof (big_sepMS_elem_of with "Hm") as "Hm"; first done.
iPoseProof (own_valid with "Hm") as "%Hv".
iLeft. iPureIntro. rewrite auth_frag_valid in Hv. apply singleton_valid in Hv. done.
Qed.
(* Fractional and AsFractional instances *) (* Fractional and AsFractional instances *)
Global Instance lft_tok_fractional κ : Fractional (λ q, q.[κ])%I. Global Instance lft_tok_fractional κ : Fractional (λ q, q.[κ])%I.
Proof. Proof.
intros p q. rewrite /lft_tok -big_sepMS_sepMS. apply big_sepMS_proper. intros p q. rewrite /lft_tok -big_sepMS_sep. apply big_sepMS_proper.
intros Λ ?. rewrite -Cinl_op -op_singleton auth_frag_op own_op //. intros Λ ?. rewrite Cinl_op -singleton_op auth_frag_op own_op //.
Qed. Qed.
Global Instance lft_tok_as_fractional κ q : Global Instance lft_tok_as_fractional κ q :
AsFractional q.[κ] (λ q, q.[κ])%I q. AsFractional q.[κ] (λ q, q.[κ])%I q.
Proof. done. Qed. Proof. split; first done. apply _. Qed.
Global Instance frame_lft_tok p κ q1 q2 q :
FrameFractionalQp q1 q2 q
Frame p q1.[κ] q2.[κ] q.[κ] | 5.
(* FIXME(https://github.com/coq/coq/issues/17688): Φ is not inferred. *)
Proof. apply: (frame_fractional (λ q, q.[κ])%I). Qed.
Global Instance idx_bor_own_fractional i : Fractional (λ q, idx_bor_own q i)%I. Global Instance idx_bor_own_fractional i : Fractional (λ q, idx_bor_own q i)%I.
Proof. Proof.
intros p q. rewrite /idx_bor_own -own_bor_op /own_bor. f_equiv=>?. intros p q. rewrite /idx_bor_own -own_bor_op /own_bor. f_equiv=>?.
by rewrite -auth_frag_op op_singleton. rewrite -auth_frag_op singleton_op -pair_op agree_idemp. done.
Qed. Qed.
Global Instance idx_bor_own_as_fractional i q : Global Instance idx_bor_own_as_fractional i q :
AsFractional (idx_bor_own q i) (λ q, idx_bor_own q i)%I q. AsFractional (idx_bor_own q i) (λ q, idx_bor_own q i)%I q.
Proof. done. Qed. Proof. split; first done. apply _. Qed.
Global Instance frame_idx_bor_own p i q1 q2 q :
FrameFractionalQp q1 q2 q
Frame p (idx_bor_own q1 i) (idx_bor_own q2 i) (idx_bor_own q i) | 5.
(* FIXME(https://github.com/coq/coq/issues/17688): Φ is not inferred. *)
Proof. apply: (frame_fractional (λ q, idx_bor_own q i)%I). Qed.
(** Lifetime inclusion *) (** Lifetime inclusion *)
Lemma lft_le_incl κ κ' : κ' κ (κ κ')%I. Lemma lft_incl_acc E κ κ' q :
lftN E
κ κ' -∗ q.[κ] ={E}=∗ q', q'.[κ'] (q'.[κ'] ={E}=∗ q.[κ]).
Proof.
rewrite /lft_incl.
iIntros (?) "#[H _] Hq". iApply fupd_mask_mono; first done.
iMod ("H" with "Hq") as (q') "[Hq' Hclose]". iExists q'.
iIntros "{$Hq'} !> Hq". iApply fupd_mask_mono; first done. by iApply "Hclose".
Qed.
Lemma lft_incl_dead E κ κ' : lftN E κ κ' -∗ [κ'] ={E}=∗ [κ].
Proof. Proof.
iIntros (->%gmultiset_union_difference) "!#". iSplitR. rewrite /lft_incl.
iIntros (?) "#[_ H] Hq". iApply fupd_mask_mono; first done. by iApply "H".
Qed.
Lemma lft_incl_intro κ κ' :
(( q, q.[κ] ={lftN}=∗ q', q'.[κ'] (q'.[κ'] ={lftN}=∗ q.[κ]))
([κ'] ={lftN}=∗ [κ])) -∗ κ κ'.
Proof. iIntros "?". done. Qed.
Lemma lft_intersect_incl_l κ κ' : κ κ' κ.
Proof.
unfold lft_incl. iIntros "!>". iSplitR.
- iIntros (q). rewrite <-lft_tok_sep. iIntros "[H Hf]". iExists q. iFrame. - iIntros (q). rewrite <-lft_tok_sep. iIntros "[H Hf]". iExists q. iFrame.
rewrite <-lft_tok_sep. by iIntros "!>{$Hf}$". rewrite <-lft_tok_sep. by iIntros "!>{$Hf}$".
- iIntros "? !>". iApply lft_dead_or. auto. - iIntros "? !>". iApply lft_dead_or. auto.
Qed. Qed.
Lemma lft_incl_refl κ : (κ κ)%I. Lemma lft_intersect_incl_r κ κ' : κ κ' κ'.
Proof. by apply lft_le_incl. Qed. Proof. rewrite comm. apply lft_intersect_incl_l. Qed.
Lemma lft_incl_trans κ κ' κ'': κ κ' -∗ κ' κ'' -∗ κ κ''. Lemma lft_incl_refl κ : κ κ.
Proof. unfold lft_incl. iIntros "!>"; iSplitR; auto 10 with iFrame. Qed.
Lemma lft_incl_trans κ κ' κ'' : κ κ' -∗ κ' κ'' -∗ κ κ''.
Proof. Proof.
rewrite /lft_incl. iIntros "#[H1 H1†] #[H2 H2†] !#". iSplitR. rewrite /lft_incl. iIntros "#[H1 H1†] #[H2 H2†] !>". iSplitR.
- iIntros (q) "Hκ". - iIntros (q) "Hκ".
iMod ("H1" with "*Hκ") as (q') "[Hκ' Hclose]". iMod ("H1" with "Hκ") as (q') "[Hκ' Hclose]".
iMod ("H2" with "*Hκ'") as (q'') "[Hκ'' Hclose']". iMod ("H2" with "Hκ'") as (q'') "[Hκ'' Hclose']".
iExists q''. iIntros "{$Hκ''} !> Hκ''". iExists q''. iIntros "{$Hκ''} !> Hκ''".
iMod ("Hclose'" with "Hκ''") as "Hκ'". by iApply "Hclose". iMod ("Hclose'" with "Hκ''") as "Hκ'". by iApply "Hclose".
- iIntros "H†". iMod ("H2†" with "H†"). by iApply "H1†". - iIntros "H†". iMod ("H2†" with "H†"). by iApply "H1†".
Qed. Qed.
Lemma lft_incl_glb κ κ' κ'' : κ κ' -∗ κ κ'' -∗ κ κ' κ''. Lemma lft_intersect_acc κ κ' q q' :
q.[κ] -∗ q'.[κ'] -∗ q'', q''.[κ κ'] (q''.[κ κ'] -∗ q.[κ] q'.[κ']).
Proof.
iIntros "Hκ Hκ'".
destruct (Qp.lower_bound q q') as (qq & q0 & q'0 & -> & ->).
iExists qq. rewrite -lft_tok_sep.
iDestruct "Hκ" as "[$$]". iDestruct "Hκ'" as "[$$]". auto.
Qed.
Lemma lft_incl_glb κ κ' κ'' : κ κ' -∗ κ κ'' -∗ κ κ' κ''.
Proof. Proof.
rewrite /lft_incl. iIntros "#[H1 H1†] #[H2 H2†]!#". iSplitR. rewrite /lft_incl. iIntros "#[H1 H1†] #[H2 H2†]!>". iSplitR.
- iIntros (q) "[Hκ'1 Hκ'2]". - iIntros (q) "[Hκ'1 Hκ'2]".
iMod ("H1" with "*Hκ'1") as (q') "[Hκ' Hclose']". iMod ("H1" with "Hκ'1") as (q') "[Hκ' Hclose']".
iMod ("H2" with "*Hκ'2") as (q'') "[Hκ'' Hclose'']". iMod ("H2" with "Hκ'2") as (q'') "[Hκ'' Hclose'']".
destruct (Qp_lower_bound q' q'') as (qq & q'0 & q''0 & -> & ->). iDestruct (lft_intersect_acc with "Hκ' Hκ''") as (qq) "[Hqq Hclose]".
iExists qq. rewrite -lft_tok_sep. iExists qq. iFrame. iIntros "!> Hqq".
iDestruct "Hκ'" as "[$ Hκ']". iDestruct "Hκ''" as "[$ Hκ'']". iDestruct ("Hclose" with "Hqq") as "[Hκ' Hκ'']".
iIntros "!> [Hκ'0 Hκ''0]". iMod ("Hclose'" with "Hκ'") as "$". by iApply "Hclose''".
iMod ("Hclose'" with "[$Hκ' $Hκ'0]") as "$". - rewrite -lft_dead_or. iIntros "[H†|H†]".
iApply "Hclose''". iFrame. + by iApply "H1†".
- rewrite -lft_dead_or. iIntros "[H†|H†]". by iApply "H1†". by iApply "H2†". + by iApply "H2†".
Qed. Qed.
Lemma lft_incl_mono κ1 κ1' κ2 κ2' : Lemma lft_intersect_mono κ1 κ1' κ2 κ2' :
κ1 κ1' -∗ κ2 κ2' -∗ κ1 κ2 κ1' κ2'. κ1 κ1' -∗ κ2 κ2' -∗ κ1 κ2 κ1' κ2'.
Proof. Proof.
iIntros "#H1 #H2". iApply (lft_incl_glb with "[]"). iIntros "#H1 #H2". iApply (lft_incl_glb with "[]").
- iApply (lft_incl_trans with "[] H1"). - iApply (lft_incl_trans with "[] H1"). iApply lft_intersect_incl_l.
iApply lft_le_incl. apply gmultiset_union_subseteq_l. - iApply (lft_incl_trans with "[] H2"). iApply lft_intersect_incl_r.
- iApply (lft_incl_trans with "[] H2").
iApply lft_le_incl. apply gmultiset_union_subseteq_r.
Qed. Qed.
Lemma lft_incl_acc E κ κ' q : (** Basic rules about borrows *)
lftN E Lemma raw_bor_iff i P P' : (P P') -∗ raw_bor i P -∗ raw_bor i P'.
κ κ' -∗ q.[κ] ={E}=∗ q', q'.[κ'] (q'.[κ'] ={E}=∗ q.[κ]).
Proof. Proof.
rewrite /lft_incl. iIntros "#HPP' HP". unfold raw_bor. iDestruct "HP" as (s) "[HiP HP]".
iIntros (?) "#[H _] Hq". iApply fupd_mask_mono; first done. iExists s. iFrame. iDestruct "HP" as (P0) "[#Hiff ?]". iExists P0. iFrame.
iMod ("H" with "* Hq") as (q') "[Hq' Hclose]". iExists q'. iNext. iModIntro. iSplit; iIntros.
iIntros "{$Hq'} !> Hq". iApply fupd_mask_mono; first done. by iApply "Hclose". - by iApply "HPP'"; iApply "Hiff".
- by iApply "Hiff"; iApply "HPP'".
Qed. Qed.
Lemma lft_incl_dead E κ κ' : lftN E κ κ' -∗ [κ'] ={E}=∗ [κ]. Lemma idx_bor_iff κ i P P' : (P P') -∗ &{κ,i}P -∗ &{κ,i}P'.
Proof. Proof.
rewrite /lft_incl. unfold idx_bor. iIntros "#HPP' [$ HP]". iDestruct "HP" as (P0) "[#HP0P Hs]".
iIntros (?) "#[_ H] Hq". iApply fupd_mask_mono; first done. by iApply "H". iExists P0. iFrame. iNext. iModIntro. iSplit; iIntros.
- by iApply "HPP'"; iApply "HP0P".
- by iApply "HP0P"; iApply "HPP'".
Qed. Qed.
(** Basic rules about borrows *)
Lemma bor_unfold_idx κ P : &{κ}P ⊣⊢ i, &{κ,i}P idx_bor_own 1 i. Lemma bor_unfold_idx κ P : &{κ}P ⊣⊢ i, &{κ,i}P idx_bor_own 1 i.
Proof. Proof.
rewrite /bor /raw_bor /idx_bor /bor_idx. iProof; iSplit. rewrite /bor /raw_bor /idx_bor /bor_idx. iSplit.
- iDestruct 1 as (κ') "[? Hraw]". iDestruct "Hraw" as (s) "[??]". - iDestruct 1 as (κ') "[? Hraw]". iDestruct "Hraw" as (s) "[??]".
iExists (κ', s). by iFrame. iExists (κ', s). by iFrame.
- iDestruct 1 as ([κ' s]) "[[??]?]". - iDestruct 1 as ([κ' s]) "[[??]?]". iFrame.
iExists κ'. iFrame. iExists s. by iFrame. Qed.
Lemma bor_iff κ P P' : (P P') -∗ &{κ}P -∗ &{κ}P'.
Proof.
rewrite !bor_unfold_idx. iIntros "#HPP' HP". iDestruct "HP" as (i) "[??]".
iExists i. iFrame. by iApply (idx_bor_iff with "HPP'").
Qed. Qed.
Lemma bor_shorten κ κ' P: κ κ' -∗ &{κ'}P -∗ &{κ}P. Lemma bor_shorten κ κ' P : κ κ' -∗ &{κ'}P -∗ &{κ}P.
Proof. Proof.
rewrite /bor. iIntros "#Hκκ'". iDestruct 1 as (i) "[#? ?]". rewrite /bor. iIntros "#Hκκ'". iDestruct 1 as (i) "[#? ?]".
iExists i. iFrame. by iApply (lft_incl_trans with "Hκκ'"). iExists i. iFrame. by iApply (lft_incl_trans with "Hκκ'").
...@@ -397,25 +532,25 @@ Proof. ...@@ -397,25 +532,25 @@ Proof.
as (γE) "(% & #Hslice & Hbox)". as (γE) "(% & #Hslice & Hbox)".
iMod (own_inh_update with "HE") as "[HE HE◯]". iMod (own_inh_update with "HE") as "[HE HE◯]".
{ by eapply auth_update_alloc, (gset_disj_alloc_empty_local_update _ {[γE]}), { by eapply auth_update_alloc, (gset_disj_alloc_empty_local_update _ {[γE]}),
disjoint_singleton_l, lookup_to_gmap_None. } disjoint_singleton_l, lookup_gset_to_gmap_None. }
iModIntro. iSplitL "Hbox HE". iModIntro. iSplitL "Hbox HE".
{ iNext. rewrite /lft_inh. iExists ({[γE]} PE). { iNext. rewrite /lft_inh. iExists ({[γE]} PE).
rewrite to_gmap_union_singleton. iFrame. } rewrite gset_to_gmap_union_singleton. iFrame. }
clear dependent PE. rewrite -(left_id_L op ( GSet {[γE]})). clear dependent PE. rewrite -(left_id ε op ( GSet {[γE]})).
iDestruct "HE◯" as "[HE◯' HE◯]". iSplitL "HE◯'". iDestruct "HE◯" as "[HE◯' HE◯]". iSplitL "HE◯'".
{ iIntros (I) "HI". iApply (own_inh_auth with "HI HE◯'"). } { iIntros (I) "HI". iApply (own_inh_auth with "HI HE◯'"). }
iIntros (Q'). rewrite {1}/lft_inh. iDestruct 1 as (PE) "[>HE Hbox]". iIntros (Q'). rewrite {1}/lft_inh. iDestruct 1 as (PE) "[>HE Hbox]".
iDestruct (own_inh_valid_2 with "HE HE◯") iDestruct (own_inh_valid_2 with "HE HE◯")
as %[Hle%gset_disj_included _]%auth_valid_discrete_2. as %[Hle%gset_disj_included _]%auth_both_valid_discrete.
iMod (own_inh_update_2 with "HE HE◯") as "HE". iMod (own_inh_update_2 with "HE HE◯") as "HE".
{ apply auth_update_dealloc, gset_disj_dealloc_local_update. } { apply auth_update_dealloc, gset_disj_dealloc_local_update. }
iMod (slice_delete_full _ _ true with "Hslice Hbox") iMod (slice_delete_full _ _ true with "Hslice Hbox")
as (Q'') "($ & #? & Hbox)"; first by solve_ndisj. as (Q'') "($ & #? & Hbox)"; first by solve_ndisj.
{ rewrite lookup_to_gmap_Some. set_solver. } { rewrite lookup_gset_to_gmap_Some. set_solver. }
iModIntro. iExists Q''. iSplit; first done. iModIntro. iExists Q''. iSplit; first done.
iNext. rewrite /lft_inh. iExists (PE {[γE]}). iFrame "HE". iNext. rewrite /lft_inh. iExists (PE {[γE]}). iFrame "HE".
rewrite {1}(union_difference_L {[ γE ]} PE); last set_solver. rewrite {1}(union_difference_L {[ γE ]} PE); last set_solver.
rewrite to_gmap_union_singleton delete_insert // lookup_to_gmap_None. rewrite gset_to_gmap_union_singleton delete_insert // lookup_gset_to_gmap_None.
set_solver. set_solver.
Qed. Qed.
...@@ -426,7 +561,7 @@ Proof. ...@@ -426,7 +561,7 @@ Proof.
rewrite /lft_inh. iIntros (?) "[Hinh HQ]". rewrite /lft_inh. iIntros (?) "[Hinh HQ]".
iDestruct "Hinh" as (E') "[Hinh Hbox]". iDestruct "Hinh" as (E') "[Hinh Hbox]".
iMod (box_fill with "Hbox HQ") as "?"=>//. iMod (box_fill with "Hbox HQ") as "?"=>//.
rewrite fmap_to_gmap. iModIntro. iExists E'. by iFrame. rewrite fmap_gset_to_gmap. iModIntro. iExists E'. by iFrame.
Qed. Qed.
(* View shifts *) (* View shifts *)
...@@ -438,18 +573,16 @@ Proof. ...@@ -438,18 +573,16 @@ Proof.
iApply ("Hvs" $! I'' with "Hinv HPb H†"). iApply ("Hvs" $! I'' with "Hinv HPb H†").
Qed. Qed.
(* TODO RJ: Are there still places where this lemma Lemma lft_vs_cons κ Pb Pb' Pi :
is re-proven inline? *) ( Pb'-∗ [κ] ={userE borN}=∗ Pb) -∗
Lemma lft_vs_cons q κ Pb Pb' Pi : lft_vs κ Pb Pi -∗ lft_vs κ Pb' Pi.
(lft_bor_dead κ Pb' ={ mgmtN}=∗ lft_bor_dead κ Pb) -∗
?q lft_vs κ Pb Pi -∗ ?q lft_vs κ Pb' Pi.
Proof. Proof.
iIntros "Hcons Hvs". iNext. rewrite !lft_vs_unfold. iIntros "Hcons Hvs". rewrite !lft_vs_unfold.
iDestruct "Hvs" as (n) "[Hn● Hvs]". iExists n. iFrame "Hn●". iDestruct "Hvs" as (n) "[Hn● Hvs]". iExists n. iFrame "Hn●".
iIntros (I). rewrite {1}lft_vs_inv_unfold. iIntros (I). rewrite {1}lft_vs_inv_unfold.
iIntros "(Hdead & Hinv & Hκs) HPb #Hκ†". iIntros "(Hinv & Hκs) HPb #Hκ†".
iMod ("Hcons" with "[$Hdead $HPb]") as "[Hdead HPb]". iMod ("Hcons" with "HPb Hκ†") as "HPb".
iApply ("Hvs" $! I with "[Hdead Hinv Hκs] HPb Hκ†"). iApply ("Hvs" $! I with "[Hinv Hκs] HPb Hκ†").
rewrite lft_vs_inv_unfold. by iFrame. rewrite lft_vs_inv_unfold. by iFrame.
Qed. Qed.
End primitive. End primitive.