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 3828 additions and 174 deletions
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing.
From lrust.typing.lib.refcell Require Import refcell.
From iris.prelude Require Import options.
Definition refcell_refN := refcellN .@ "ref".
Section ref.
Context `{!typeGS Σ, !refcellG Σ}.
(* The Rust type looks as follows (after some unfolding):
pub struct Ref<'b, T: ?Sized + 'b> {
value: &'b T,
borrow: &'b Cell<BorrowFlag>,
}
*)
Program Definition ref (α : lft) (ty : type) :=
tc_opaque
{| ty_size := 2;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc lv); #(LitLoc lrc) ] =>
ν q γ β ty', ty.(ty_shr) (α ν) tid lv
α β &na{β, tid, refcell_invN}(refcell_inv tid lrc γ β ty')
q.[ν] own γ ( reading_stR q ν)
| _ => False
end;
ty_shr κ tid l :=
ν q γ β ty' (lv lrc : loc),
κ ν &frac{κ} (λ q, l↦∗{q} [ #lv; #lrc])
ty.(ty_shr) (α ν) tid lv
(α β) &na{β, tid, refcell_invN}(refcell_inv tid lrc γ β ty')
&na{κ, tid, refcell_refN}(own γ ( reading_stR q ν)) |}%I.
Next Obligation. iIntros (???[|[[]|][|[[]|][]]]) "?"; auto. Qed.
Next Obligation.
iIntros (α ty E κ l tid q ?) "#LFT Hb Htok".
iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done.
iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done.
destruct vl as [|[[|lv|]|][|[[|lrc|]|][]]];
try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]".
iMod (bor_exists with "LFT Hb") as (ν) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (q') "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (γ) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (β) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (ty') "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[Hshr Hb]"; first done.
iMod (bor_persistent with "LFT Hshr Htok") as "[#Hshr Htok]"; first done.
iMod (bor_sep with "LFT Hb") as "[Hαβ Hb]"; first done.
iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ Htok]"; first done.
iMod (bor_sep with "LFT Hb") as "[Hinv Hb]"; first done.
iMod (bor_persistent with "LFT Hinv Htok") as "[#Hinv $]"; first done.
iMod (bor_sep with "LFT Hb") as "[Hκν Hb]"; first done.
iDestruct (frac_bor_lft_incl with "LFT [> Hκν]") as "#Hκν".
{ iApply bor_fracture; try done. by rewrite Qp.mul_1_r. }
iMod (bor_na with "Hb") as "#Hb"; first done. eauto 20.
Qed.
Next Obligation.
iIntros (??????) "#? H". iDestruct "H" as (ν q γ β ty' lv lrc) "H".
iExists _, _, _, _, _, _, _. iDestruct "H" as "#(? & ? & $ & $ & $ & ?)".
iSplit; last iSplit.
- by iApply lft_incl_trans.
- by iApply frac_bor_shorten.
- by iApply na_bor_shorten.
Qed.
Global Instance ref_wf α ty `{!TyWf ty} : TyWf (ref α ty) :=
{ ty_lfts := [α]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty α }.
Global Instance ref_type_contractive α : TypeContractive (ref α).
Proof. solve_type_proper. Qed.
Global Instance ref_ne α : NonExpansive (ref α).
Proof. apply type_contractive_ne, _. Qed.
Global Instance ref_mono E L :
Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) ref.
Proof.
iIntros (α1 α2 ty1 ty2 Hty qmax qL) "HL".
iDestruct (Hty with "HL") as "#Hty". iDestruct ( with "HL") as "#Hα".
iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2.
iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro].
- done.
- iIntros (tid [|[[]|][|[[]|][]]]) "H"=>//=.
iDestruct "H" as (ν q' γ β ty') "(#Hshr & #H⊑ & #Hinv & Htok & Hown)".
iExists ν, q', γ, β, ty'. iFrame "∗#". iSplit.
+ iApply ty_shr_mono; last by iApply "Hs".
iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl.
+ iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done.
- iIntros (κ tid l) "H /=". iDestruct "H" as (ν q' γ β ty' lv lrc) "H".
iExists ν, q', γ, β, ty', lv, lrc. iDestruct "H" as "#($&$&?&?&$&$)". iSplit.
+ iApply ty_shr_mono; last by iApply "Hs".
iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl.
+ iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done.
Qed.
Global Instance ref_mono_flip E L :
Proper (lctx_lft_incl E L ==> flip (subtype E L) ==> flip (subtype E L)) ref.
Proof. intros ??????. by apply ref_mono. Qed.
Lemma ref_mono' E L α1 α2 ty1 ty2 :
lctx_lft_incl E L α2 α1 subtype E L ty1 ty2
subtype E L (ref α1 ty1) (ref α2 ty2).
Proof. intros. by eapply ref_mono. Qed.
Global Instance ref_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) ref.
Proof. intros ??[]?? EQ. split; apply ref_mono'; try done; apply EQ. Qed.
Lemma ref_proper' E L α1 α2 ty1 ty2 :
lctx_lft_eq E L α1 α2 eqtype E L ty1 ty2
eqtype E L (ref α1 ty1) (ref α2 ty2).
Proof. intros. by eapply ref_proper. Qed.
End ref.
Global Hint Resolve refcell_mono' refcell_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree numbers.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import lifetime na_borrow.
From lrust.typing Require Import typing.
From lrust.typing.lib.refcell Require Import refcell ref.
From iris.prelude Require Import options.
Section ref_functions.
Context `{!typeGS Σ, !refcellG Σ}.
Lemma refcell_inv_reading_inv tid l γ α ty q ν :
refcell_inv tid l γ α ty -∗
own γ ( reading_stR q ν) -∗
(q' : Qp) n, l #(Zpos n) (q q')%Qp
own γ ( (refcell_st_to_R $ Some (ν, false, q', n)) reading_stR q ν)
((1).[ν] ={lftN lft_userE}[lft_userE]▷=∗ &{α}((l + 1) ↦∗: ty_own ty tid))
( q'', (q' + q'' = 1)%Qp q''.[ν])
ty.(ty_shr) (α ν) tid (l + 1).
Proof.
iIntros "INV H◯".
iDestruct "INV" as (st) "(H↦lrc & H● & INV)".
iAssert ( q' n, st Some (ν, false, q', n) q q'⌝%Qp)%I with
"[#]" as (q' n) "[%%]".
{ destruct st as [[[[??]?]?]|];
iDestruct (own_valid_2 with "H● H◯")
as %[[[=]|(?&[[? q'] n]&[=<-]&[=]&[[[Eq_ag ?%leibniz_equiv]_]|Hle])]
%option_included Hv]%auth_both_valid_discrete; simpl in *; subst.
- apply (inj to_agree), (inj2 pair) in Eq_ag.
destruct Eq_ag. setoid_subst. eauto.
- revert Hle=> /prod_included [/= /prod_included
[/= /to_agree_included [/= ??] /frac_included_weak ?] _].
setoid_subst. eauto. }
setoid_subst. iExists q', n. rewrite own_op. by iFrame.
Qed.
(* Cloning a ref. We need to increment the counter. *)
Definition ref_clone : val :=
fn: ["x"] :=
let: "x'" := !"x" in
let: "rc" := !("x'" + #1) in
let: "n" := !"rc" in
"rc" <- "n" + #1;;
letalloc: "r" <-{2} !"x'" in
delete [ #1; "x"];; return: ["r"].
Lemma ref_clone_type ty `{!TyWf ty} :
typed_val ref_clone (fn( '(α, β), ; &shr{α}(ref β ty)) ref β ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=.
iDestruct "Hx'" as (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hβδ & Hinv & H◯inv)".
wp_op.
iMod (lctx_lft_alive_tok β with "HE HL") as () "(Hβ & HL & Hclose)"; [solve_typing..|].
iMod (lft_incl_acc with "Hβδ Hβ") as () "[Hδ Hcloseβ]"; first done.
iMod (lctx_lft_alive_tok_noend α with "HE HL") as () "([Hα1 Hα2] & HL & Hclose')";
[solve_typing..|].
iMod (frac_bor_acc with "LFT Hfrac Hα1") as (qlx') "[H↦ Hcloseα1]"; first done.
iMod (na_bor_acc with "LFT Hinv Hδ Hna") as "(INV & Hna & Hcloseδ)"; [done..|].
iMod (na_bor_acc with "LFT H◯inv Hα2 Hna") as "(H◯ & Hna & Hcloseα2)"; [solve_ndisj..|].
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iMod "H↦" as "[H↦1 H↦2]". wp_read. wp_let.
iDestruct (refcell_inv_reading_inv with "INV H◯")
as (q' n) "(H↦lrc & _ & [H● H◯] & H† & Hq' & Hshr')".
wp_read. wp_let. wp_op. wp_write. iDestruct "Hq'" as (q'') "(Hq'q'' & Hν1 & Hν2)".
iDestruct "Hq'q''" as %Hq'q''. iMod (own_update with "H●") as "[H● ?]".
{ apply auth_update_alloc,
(op_local_update_discrete _ _ (reading_stR (q''/2)%Qp ν))=>-[Hagv _].
split; [split|done].
- by rewrite /= agree_idemp.
- apply frac_valid. rewrite /= -Hq'q'' comm_L.
by apply Qp.add_le_mono_l, Qp.div_le. }
wp_apply wp_new; [done..|]. iIntros (lr) "(?&Hlr)".
iAssert (lx' ↦∗{qlx'} [ #lv; #lrc])%I with "[H↦1 H↦2]" as "H↦".
{ rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton. iFrame. }
wp_let. wp_apply (wp_memcpy with "[$Hlr $H↦]"); [done..|].
iIntros "[Hlr H↦]". wp_seq. iMod ("Hcloseα2" with "[$H◯] Hna") as "[Hα1 Hna]".
iMod ("Hcloseδ" with "[H↦lrc H● Hν1 Hshr' H†] Hna") as "[Hδ Hna]".
{ iExists (Some (_, false, _, _)). rewrite Z.add_comm -Some_op -!pair_op agree_idemp.
iFrame. rewrite (comm Qp.add) (assoc Qp.add) Qp.div_2 (comm Qp.add). auto. }
iMod ("Hcloseβ" with "Hδ") as "Hβ". iMod ("Hcloseα1" with "[$H↦]") as "Hα2".
iMod ("Hclose'" with "[$Hα1 $Hα2] HL") as "HL". iMod ("Hclose" with "Hβ HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(ref β ty)); #lr box(ref β ty)]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
rewrite /= freeable_sz_full. iFrame "∗#". simpl. iFrame "∗#". }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Turning a ref into a shared borrow. *)
Definition ref_deref : val :=
fn: ["x"] :=
let: "x'" := !"x" in
let: "r'" := !"x'" in
letalloc: "r" <- "r'" in
delete [ #1; "x"];; return: ["r"].
Lemma ref_deref_type ty `{!TyWf ty} :
typed_val ref_deref
(fn( '(α, β), ; &shr{α}(ref β ty)) &shr{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=.
iDestruct "Hx'" as (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hx')".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (frac_bor_acc with "LFT Hfrac Hα") as (qlx') "[H↦ Hcloseα]"; first done.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iMod "H↦" as "[H↦1 H↦2]". wp_read. wp_let.
iMod ("Hcloseα" with "[$H↦1 $H↦2]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|].
iApply (type_type _ _ _
[ x box (&shr{α}(ref β ty)); #lv &shr{α}ty]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
iFrame. iApply (ty_shr_mono with "[] Hshr").
iApply lft_incl_glb; last done. by iApply lft_incl_syn_sem. }
iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|].
iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Dropping a ref and decrement the count in the corresponding refcell. *)
Definition ref_drop : val :=
fn: ["x"] :=
let: "rc" := !("x" + #1) in
let: "n" := !"rc" in
"rc" <- "n" - #1;;
Endlft;;
delete [ #2; "x"];;
let: "r" := new [ #0] in return: ["r"].
Lemma ref_drop_type ty `{!TyWf ty} :
typed_val ref_drop (fn( α, ; ref α ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val.
destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]".
iDestruct "Hx" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hx"; simpl;
try iDestruct "Hx" as "[_ >[]]".
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hx" as "[[Hx↦1 Hx↦2] Hx]". wp_op. wp_read. wp_let.
iDestruct "Hx" as (ν q γ β ty') "(_ & #Hαβ & #Hinv & Hν & H◯)".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[Hβ Hcloseα]"; first done.
iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hcloseβ)"; [done..|].
iDestruct (refcell_inv_reading_inv with "INV [$H◯]")
as (q' n) "(H↦lrc & >% & H●◯ & H† & Hq' & Hshr)".
iDestruct "Hq'" as (q'') ">[% Hν']".
wp_read. wp_let. wp_op. wp_write.
iAssert (|={lftN lft_userE}[lft_userE]▷=> refcell_inv tid lrc γ β ty')%I
with "[H↦lrc H●◯ Hν Hν' Hshr H†]" as "INV".
{ iDestruct (own_valid with "H●◯") as
%[[[[_ ?]?]|[[_ [q0 Hq0]]%prod_included [n' Hn']]%prod_included]
%Some_included _]%auth_both_valid_discrete.
- simpl in *. setoid_subst.
iExists None. iFrame. iMod (own_update with "H●◯") as "$".
{ apply auth_update_dealloc. rewrite -(right_id None op (Some _)).
apply cancel_local_update_unit, _. }
iApply "H†". replace 1%Qp with (q'+q'')%Qp by naive_solver. iFrame.
- simpl in *. setoid_subst. iExists (Some (ν, false, q0, n')). iFrame "Hshr H†".
iAssert (lrc #(Z.pos n'))%I with "[H↦lrc]" as "$".
{ rewrite pos_op_add Z.sub_1_r -Pos2Z.inj_pred; last lia.
by rewrite Pos.add_1_l Pos.pred_succ. }
iMod (own_update with "H●◯") as "$".
{ apply auth_update_dealloc.
rewrite -(agree_idemp (to_agree _)) !pair_op Some_op.
apply (cancel_local_update_unit (reading_stR q ν)), _. }
iApply step_fupd_intro; first set_solver+. iExists (q+q'')%Qp. iFrame.
by rewrite assoc (comm _ q0 q). }
wp_bind Endlft. iApply (wp_mask_mono _ (lftN lft_userE)); first done.
iApply (wp_step_fupd with "INV"); [set_solver+..|]. wp_seq.
iIntros "INV !>". wp_seq. iMod ("Hcloseβ" with "[$INV] Hna") as "[Hβ Hna]".
iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ #lx box (uninit 2)] with "[] LFT HE Hna HL Hk");
first last.
{ rewrite tctx_interp_singleton tctx_hasty_val' //. iFrame. iExists [ #lv;#lrc].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton uninit_own. iFrame. auto. }
iApply type_delete; [solve_typing..|].
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_jump; solve_typing.
Qed.
(* Apply a function within the ref, typically for accessing a component. *)
Definition ref_map (call_once : val) : val :=
fn: ["ref"; "f"] :=
let: "call_once" := call_once in
let: "x'" := !"ref" in
letalloc: "x" <- "x'" in
letcall: "r" := "call_once" ["f"; "x"]%E in
let: "r'" := !"r" in delete [ #1; "r"];;
"ref" <- "r'";;
return: ["ref"].
Lemma ref_map_type ty1 ty2 call_once fty `{!TyWf ty1, !TyWf ty2, !TyWf fty} :
(* fty : for<'a>, FnOnce(&'a ty1) -> &'a ty2, as witnessed by the impl call_once *)
typed_val call_once (fn( α, ; fty, &shr{α}ty1) &shr{α}ty2)
typed_val (ref_map call_once) (fn( α, ; ref α ty1, fty) ref α ty2).
Proof.
intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg).
inv_vec arg=>ref env. simpl_subst.
iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)".
rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done.
iDestruct "Href" as "[Href Href†]".
iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl;
try iDestruct "Href" as "[_ >[]]".
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Href" as "[[Href↦1 Href↦2] Href]".
iDestruct "Href" as (ν γ β ty') "(#Hshr & #Hαβ & #Hinv & >Hν & Hγ)".
wp_read. wp_let. wp_apply wp_new; try done.
iIntros (lx) "(H† & Hlx)". rewrite heap_pointsto_vec_singleton.
wp_let. wp_write. wp_let. rewrite tctx_hasty_val.
iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|].
iMod (lctx_lft_alive_tok_noend ϝ with "HE HL") as (?) "(Hϝ & HL & Hclose2)";[solve_typing..|].
iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]".
iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]".
rewrite -[ϝ in (α ν) ϝ]right_id_L.
iApply (type_call_iris _ [α ν; ϝ] (α ν) [_; _]
with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H†]"); [solve_typing|done| |].
{ rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full.
iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. by iFrame. }
iIntros ([[|r|]|]) "Hna Hανϝ Hr //".
iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]".
iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]".
iMod ("Hclose2" with "Hϝ HL") as "HL". iMod ("Hclose1" with "Hα HL") as "HL".
wp_rec. iDestruct "Hr" as "[Hr Hr†]".
iDestruct "Hr" as ([|r'[]]) "[Hr #Hr']";
try by iDestruct (ty_size_eq with "Hr'") as "%".
rewrite heap_pointsto_vec_singleton. wp_read. wp_let.
wp_apply (wp_delete _ _ _ [_] with "[Hr Hr†]")=>//.
{ rewrite heap_pointsto_vec_singleton freeable_sz_full. iFrame. } iIntros "_".
wp_seq. wp_write.
iApply (type_type _ [_] _ [ #lref box (ref α ty2) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_singleton tctx_hasty_val' //. simpl. iFrame.
iExists [_;_]. rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iFrame. destruct r' as [[]|]=>//=. auto 10 with iFrame. }
iApply type_jump; solve_typing.
Qed.
(* Apply a function within the ref, and create two ref,
typically for splitting the reference. *)
Definition ref_map_split (call_once : val) : val :=
fn: ["ref"; "f"] :=
let: "call_once" := call_once in
let: "x'" := !"ref" in
letalloc: "x" <- "x'" in
letcall: "r" := "call_once" ["f"; "x"]%E in
let: "a" := !"r" in
let: "b" := !("r" + #1) in
delete [ #2; "r"];;
let: "rc" := !("ref" + #1) in
let: "n" := !"rc" in
"rc" <- "n" + #1;;
delete [ #2; "ref"];;
let: "refs" := new [ #4 ] in
"refs" <- "a";;
"refs" + #1 <- "rc";;
"refs" + #2 <- "b";;
"refs" + #3 <- "rc";;
return: ["refs"].
Lemma ref_map_split_type ty ty1 ty2 call_once fty
`{!TyWf ty, !TyWf ty1, !TyWf ty2, !TyWf fty} :
(* fty : for<'a>, FnOnce(&'a ty) -> (&'a ty1, &'a ty2),
as witnessed by the impl call_once *)
typed_val call_once
(fn( α, ; fty, &shr{α}ty) Π[&shr{α}ty1; &shr{α}ty2])
typed_val (ref_map_split call_once)
(fn( α, ; ref α ty, fty) Π[ref α ty1; ref α ty2]).
Proof.
intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg).
inv_vec arg=>ref env. simpl_subst.
iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)".
rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done.
iDestruct "Href" as "[Href Href†]".
iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl;
try iDestruct "Href" as "[_ >[]]".
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Href" as "[[Href↦1 Href↦2] Href]".
iDestruct "Href" as (ν γ β ty') "(#Hshr & #Hαβ & #Hinv & >Hν & Hγ)".
wp_read. wp_let. wp_apply wp_new; [done..|].
iIntros (lx) "(H† & Hlx)". rewrite heap_pointsto_vec_singleton.
wp_let. wp_write. wp_let. rewrite tctx_hasty_val.
iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|].
iMod (lctx_lft_alive_tok_noend ϝ with "HE HL") as (?) "(Hϝ & HL & Hclose2)";[solve_typing..|].
iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]".
iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]".
rewrite -[ϝ in (α ν) ϝ]right_id_L.
iApply (type_call_iris _ [α ν; ϝ] (α ν) [_; _]
with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H†]"); [solve_typing|done| |].
{ rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full.
iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. by iFrame. }
iIntros ([[|r|]|]) "Hna Hανϝ Hr //".
iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]".
iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]".
iMod ("Hclose2" with "Hϝ HL") as "HL".
wp_rec. iDestruct "Hr" as "[Hr Hr†]".
iDestruct "Hr" as (?) "[Hr H]". iDestruct "H" as (vl1 vl1' ->) "[#Hr1' H]".
iDestruct "H" as (vl2 vl2' ->) "[#Hr2' ->]".
destruct vl1 as [|[[|lr1|]|] []], vl2 as [|[[|lr2|]|] []]=>//=.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hr" as "[Hr1 Hr2]". wp_read. wp_let. wp_op. wp_read. wp_let.
wp_apply (wp_delete _ _ _ [_; _] with "[Hr1 Hr2 Hr†]")=>//.
{ rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton /= freeable_sz_full.
iFrame. }
iIntros "_".
iMod (lft_incl_acc with "Hαβ Hα") as (?) "[Hβ Hβclose]"; first done.
iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hclosena)"; [done..|].
wp_seq. wp_op. wp_read.
iDestruct (refcell_inv_reading_inv with "INV Hγ")
as (q1 n) "(H↦lrc & _ & [H● H◯] & H† & Hq1 & Hshr')".
iDestruct "Hq1" as (q2) "(Hq1q2 & Hν1 & Hν2)".
iDestruct "Hq1q2" as %Hq1q2. iMod (own_update with "H●") as "[H● ?]".
{ apply auth_update_alloc,
(op_local_update_discrete _ _ (reading_stR (q2/2)%Qp ν))=>-[Hagv _].
split; [split|done].
- by rewrite /= agree_idemp.
- apply frac_valid. rewrite /= -Hq1q2 comm_L.
by apply Qp.add_le_mono_l, Qp.div_le. }
wp_let. wp_read. wp_let. wp_op. wp_write.
wp_apply (wp_delete _ _ _ [_; _] with "[Href↦1 Href↦2 Href†]")=>//.
{ rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton /= freeable_sz_full.
iFrame. }
iIntros "_". wp_seq. wp_apply wp_new; [done..|]. iIntros (lrefs) "[Hrefs† Hrefs]".
rewrite 3!heap_pointsto_vec_cons heap_pointsto_vec_singleton. wp_let.
iDestruct "Hrefs" as "(Hrefs1 & Hrefs2 & Hrefs3 & Hrefs4)".
rewrite !shift_loc_assoc. wp_write. do 3 (wp_op; wp_write).
iMod ("Hclosena" with "[H↦lrc H● Hν1 Hshr' H†] Hna") as "[Hβ Hna]".
{ iExists (Some (_, false, _, _)). rewrite Z.add_comm -Some_op -!pair_op agree_idemp.
iFrame. rewrite (comm Qp.add) (assoc Qp.add) Qp.div_2 (comm Qp.add). auto. }
iMod ("Hβclose" with "Hβ") as "Hα". iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ [_] _ [ #lrefs box (Π[ref α ty1; ref α ty2]) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_singleton tctx_hasty_val' //= -freeable_sz_full.
simpl. iFrame. iExists [_;_;_;_].
rewrite 3!heap_pointsto_vec_cons heap_pointsto_vec_singleton !shift_loc_assoc.
iFrame. iExists [_;_], [_;_]. iSplit=>//=.
iSplitL "Hν H◯"; [by auto 10 with iFrame|]. iExists [_;_], [].
iSplitR=>//=. rewrite right_id. by iFrame "∗#". }
iApply type_jump; solve_typing.
Qed.
End ref_functions.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree numbers.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Definition refcell_stR :=
optionUR (prodR (prodR
(agreeR (prodO lftO boolO))
fracR)
positiveR).
Class refcellG Σ :=
RefCellG :: inG Σ (authR refcell_stR).
Definition refcellΣ : gFunctors := #[GFunctor (authR refcell_stR)].
Global Instance subG_refcellΣ {Σ} : subG refcellΣ Σ refcellG Σ.
Proof. solve_inG. Qed.
Definition refcell_st := option ((lft * Datatypes.bool) * frac * positive).
Coercion refcell_st_to_R (st : refcell_st) : refcell_stR :=
match st with
| None => None
| Some (x, q, n) => Some (to_agree x, q, n)
end.
Definition Z_of_refcell_st (st : refcell_st) : Z :=
match st with
| None => 0
| Some (_, false, _, n) => Zpos n (* Immutably borrowed *)
| Some (_, true, _, n) => Zneg n (* Mutably borrowed *)
end.
Definition reading_stR (q : frac) (ν : lft) : refcell_stR :=
refcell_st_to_R $ Some (ν, false, q, xH).
Definition writing_stR (q : frac) (ν : lft) : refcell_stR :=
refcell_st_to_R $ Some (ν, true, q, xH).
Definition refcellN := lrustN .@ "refcell".
Definition refcell_invN := refcellN .@ "inv".
Section refcell_inv.
Context `{!typeGS Σ, !refcellG Σ}.
Definition refcell_inv tid (l : loc) (γ : gname) (α : lft) ty : iProp Σ :=
( st, l #(Z_of_refcell_st st) own γ ( (refcell_st_to_R st))
match st return _ with
| None =>
(* Not borrowed. *)
&{α}((l + 1) ↦∗: ty.(ty_own) tid)
| Some (ν, rw, q, _) =>
(1.[ν] ={lftN lft_userE}[lft_userE]▷=∗ &{α}((l + 1) ↦∗: ty.(ty_own) tid))
( q', (q + q')%Qp = 1%Qp q'.[ν])
if rw then (* Mutably borrowed. *) True
else (* Immutably borrowed. *) ty.(ty_shr) (α ν) tid (l + 1)
end)%I.
Global Instance refcell_inv_type_ne n tid l γ α :
Proper (type_dist2 n ==> dist n) (refcell_inv tid l γ α).
Proof.
solve_proper_core
ltac:(fun _ => exact: type_dist2_S || f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance refcell_inv_ne tid l γ α : NonExpansive (refcell_inv tid l γ α).
Proof.
intros n ???. (* TODO: f_equiv takes forever here, what is going on? *)
eapply refcell_inv_type_ne, type_dist_dist2. done.
Qed.
Lemma refcell_inv_proper E L ty1 ty2 qmax qL :
eqtype E L ty1 ty2
llctx_interp_noend qmax L qL -∗ tid l γ α, (elctx_interp E -∗
refcell_inv tid l γ α ty1 -∗ refcell_inv tid l γ α ty2).
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic. *)
(* It would easily gain from some automation. *)
rewrite eqtype_unfold. iIntros (Hty) "HL".
iDestruct (Hty with "HL") as "#Hty". iIntros (tid l γ α) "!> #HE H".
iDestruct ("Hty" with "HE") as "(% & #Hown & #Hshr)".
iAssert ( (&{α}((l + 1) ↦∗: ty_own ty1 tid) -∗
&{α}((l + 1) ↦∗: ty_own ty2 tid)))%I as "#Hb".
{ iIntros "!> H". iApply bor_iff; last done.
iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl;
iFrame; by iApply "Hown". }
iDestruct "H" as (st) "H"; iExists st;
iDestruct "H" as "($&$&H)"; destruct st as [[[[ν rw] q' ] n]|]; try done;
last by iApply "Hb".
iDestruct "H" as "(Hh & $ & H)". iSplitL "Hh".
{ iIntros "Hν". iMod ("Hh" with "Hν") as "Hh". iModIntro. iNext.
iMod "Hh". by iApply "Hb". }
destruct rw; try done. by iApply "Hshr".
Qed.
End refcell_inv.
Section refcell.
Context `{!typeGS Σ, !refcellG Σ}.
(* Original Rust type:
pub struct RefCell<T: ?Sized> {
borrow: Cell<BorrowFlag>,
value: UnsafeCell<T>,
}
*)
Program Definition refcell (ty : type) :=
tc_opaque
{| ty_size := S ty.(ty_size);
ty_own tid vl :=
match vl return _ with
| #(LitInt z) :: vl' => ty.(ty_own) tid vl'
| _ => False
end%I;
ty_shr κ tid l :=
( α γ, κ α &na{α, tid, refcell_invN}(refcell_inv tid l γ α ty))%I |}.
Next Obligation.
iIntros (??[|[[]|]]); try iIntros "[]". rewrite ty_size_eq /=. by iIntros (->).
Qed.
Next Obligation.
iIntros (ty E κ l tid q ?) "#LFT Hb Htok".
iMod (bor_acc_cons with "LFT Hb Htok") as "[H Hclose]"; first done.
iDestruct "H" as ([|[[| |n]|]vl]) "[H↦ H]"; try iDestruct "H" as ">[]".
iDestruct "H" as "Hown".
iMod ("Hclose" $! (( n:Z, l #n)
(l + 1) ↦∗: ty.(ty_own) tid) with "[] [-]")%I as "[H [Htok Htok']]".
{ iIntros "!> [Hn Hvl] !>". iDestruct "Hn" as (n') "Hn".
iDestruct "Hvl" as (vl') "[H↦ Hvl]".
iExists (#n'::vl'). rewrite heap_pointsto_vec_cons. iFrame. }
{ iNext. rewrite heap_pointsto_vec_cons. iDestruct "H↦" as "[Hn Hvl]".
iSplitL "Hn"; eauto with iFrame. }
iMod (bor_sep with "LFT H") as "[Hn Hvl]"; first done.
iMod (bor_acc_cons with "LFT Hn Htok") as "[H Hclose]"; first done.
iAssert ((q / 2).[κ] γ, refcell_inv tid l γ κ ty)%I with "[> -Hclose]"
as "[$ HQ]"; last first.
{ iMod ("Hclose" with "[] HQ") as "[Hb $]".
- iIntros "!> H !>". iNext. iDestruct "H" as (γ st) "(H & _ & _)". eauto.
- iMod (bor_exists with "LFT Hb") as (γ) "Hb"; first done.
iExists κ, γ. iSplitR.
+ by iApply lft_incl_refl.
+ by iApply bor_na. }
clear dependent n. iDestruct "H" as ([|n|n]) "Hn"; try lia.
- iFrame. iMod (own_alloc ( None)) as (γ) "Hst"; first by apply auth_auth_valid.
iExists γ, None. by iFrame.
- iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]"; first done.
iMod (own_alloc ( (refcell_st_to_R $ Some (ν, false, (1/2)%Qp, n)))) as (γ) "Hst".
{ by apply auth_auth_valid. }
iApply (fupd_mask_mono (lftN)); first done.
iMod (rebor _ _ (κ ν) with "LFT [] Hvl") as "[Hvl Hh]"; first done.
{ iApply lft_intersect_incl_l. }
iDestruct (lft_intersect_acc with "Htok' Htok1") as (q') "[Htok Hclose]".
iMod (ty_share with "LFT Hvl Htok") as "[Hshr Htok]"; first done.
iDestruct ("Hclose" with "Htok") as "[$ Htok]".
iExists γ, _. iFrame "Hst Hn Hshr".
iSplitR "Htok2"; last by iExists _; iFrame; rewrite Qp.div_2.
iIntros "!> !> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν".
iApply fupd_mask_mono; last iApply "Hh"; first set_solver+. rewrite -lft_dead_or. auto.
- iMod (own_alloc ( (refcell_st_to_R $ Some (static, true, (1/2)%Qp, n)))) as (γ) "Hst".
{ by apply auth_auth_valid. }
iFrame "Htok'". iExists γ, _. iFrame. iSplitR.
{ rewrite -step_fupd_intro; first auto. set_solver+. }
iSplitR; [|done]. iExists (1/2)%Qp. rewrite Qp.div_2. iSplitR; [done|].
iApply lft_tok_static.
Qed.
Next Obligation.
iIntros (?????) "#Hκ H". iDestruct "H" as (α γ) "[#??]".
iExists _, _. iFrame. iApply lft_incl_trans; auto.
Qed.
Global Instance refcell_wf ty `{!TyWf ty} : TyWf (refcell ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance refcell_type_ne : TypeNonExpansive refcell.
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || (eapply refcell_inv_type_ne; try reflexivity) ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance refcell_ne : NonExpansive refcell.
Proof.
constructor; solve_proper_core ltac:(fun _ => (eapply ty_size_ne; try reflexivity) || f_equiv).
Qed.
Global Instance refcell_mono E L : Proper (eqtype E L ==> subtype E L) refcell.
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic.
It would easily gain from some automation. *)
iIntros (ty1 ty2 EQ qmax qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'.
iDestruct (EQ' with "HL") as "#EQ'".
iDestruct (refcell_inv_proper with "HL") as "#Hty1ty2"; first done.
iDestruct (refcell_inv_proper with "HL") as "#Hty2ty1"; first by symmetry.
iIntros "!> #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)".
iSplit; last iSplit.
- iPureIntro. simpl. congruence.
- iIntros "!> %tid %vl H". destruct vl as [|[[]|]]=>//=. by iApply "Hown".
- iIntros "!> %α %tid %l H". simpl.
iDestruct "H" as (a γ) "[Ha H]". iExists a, γ. iFrame.
iApply na_bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H".
+ by iApply "Hty1ty2".
+ by iApply "Hty2ty1".
Qed.
Lemma refcell_mono' E L ty1 ty2 :
eqtype E L ty1 ty2 subtype E L (refcell ty1) (refcell ty2).
Proof. eapply refcell_mono. Qed.
Global Instance refcell_proper E L : Proper (eqtype E L ==> eqtype E L) refcell.
Proof. by split; apply refcell_mono. Qed.
Lemma refcell_proper' E L ty1 ty2 :
eqtype E L ty1 ty2 eqtype E L (refcell ty1) (refcell ty2).
Proof. eapply refcell_proper. Qed.
Global Instance refcell_send ty :
Send ty Send (refcell ty).
Proof. move=>???[|[[]|]]//=. Qed.
End refcell.
Global Hint Resolve refcell_mono' refcell_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lang.lib Require Import memcpy.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing option.
From lrust.typing.lib.refcell Require Import refcell ref refmut.
From iris.prelude Require Import options.
Section refcell_functions.
Context `{!typeGS Σ, !refcellG Σ}.
(* Constructing a refcell. *)
Definition refcell_new ty : val :=
fn: ["x"] :=
let: "r" := new [ #(S ty.(ty_size))] in
"r" + #0 <- #0;;
"r" + #1 <-{ty.(ty_size)} !"x";;
delete [ #ty.(ty_size) ; "x"];; return: ["r"].
Lemma refcell_new_type ty `{!TyWf ty} :
typed_val (refcell_new ty) (fn(; ty) refcell ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new (S ty.(ty_size))); [solve_typing..|].
iIntros (r tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hr Hx]".
iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x.
iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr↦ & Hr†)". subst r.
inv_vec vlr=>??. iDestruct (heap_pointsto_vec_cons with "Hr↦") as "[Hr↦0 Hr↦1]". wp_op.
rewrite shift_loc_0. wp_write. wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz.
wp_apply (wp_memcpy with "[$Hr↦1 $Hx↦]"); [by auto using length_vec_to_list..|].
iIntros "[Hr↦1 Hx↦]". wp_seq.
iApply (type_type _ _ _ [ #lx box (uninit (ty_size ty)); #lr box (refcell ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=.
iFrame. iSplitR; first done.
simpl. iFrame. iExists (_::_). rewrite heap_pointsto_vec_cons. by iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* The other direction: getting ownership out of a refcell. *)
Definition refcell_into_inner ty : val :=
fn: ["x"] :=
let: "r" := new [ #ty.(ty_size)] in
"r" <-{ty.(ty_size)} !("x" + #1);;
(* TODO: Can we make it so that the parenthesis above are mandatory?
Leaving them away is inconsistent with `let ... := !"x" +ₗ #1`. *)
delete [ #(S ty.(ty_size)) ; "x"];; return: ["r"].
Lemma refcell_into_inner_type ty `{!TyWf ty} :
typed_val (refcell_into_inner ty) (fn(; refcell ty) ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new ty.(ty_size)); [solve_typing..|].
iIntros (r tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hr Hx]".
iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & >Hx↦ & Hx & Hx†)". subst x.
inv_vec vlx=>-[[|?|?]|????] vl; simpl; try iDestruct "Hx" as ">[]".
iDestruct (heap_pointsto_vec_cons with "Hx↦") as "[Hx↦0 Hx↦1]".
iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr↦ & Hr†)". subst r.
wp_op. wp_apply (wp_memcpy with "[$Hr↦ $Hx↦1]"); [by auto using length_vec_to_list..|].
iIntros "[Hr↦ Hx↦1]". wp_seq.
iApply (type_type _ _ _ [ #lx box (uninit (S (ty_size ty))); #lr box ty]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame.
iExists (_::_). rewrite heap_pointsto_vec_cons uninit_own. iFrame.
by rewrite /= length_vec_to_list. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition refcell_get_mut : val :=
fn: ["x"] :=
let: "x'" := !"x" in
"x" <- "x'" + #1;; (* Get the second field *)
return: ["x"].
Lemma refcell_get_mut_type ty `{!TyWf ty} :
typed_val refcell_get_mut
(fn( α, ; &uniq{α} (refcell ty)) &uniq{α} ty)%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL HC HT".
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try iDestruct "Hx'" as "[]".
iAssert (&{α} ( (z : Z), lx' #z)
(&uniq{α} ty).(ty_own) tid [ #(lx' + 1)])%I with "[> Hx']" as "[_ Hx']".
{ iApply bor_sep; [done..|]. iApply (bor_proper with "Hx'"). iSplit.
- iIntros "[H1 H2]". iDestruct "H1" as (z) "?". iDestruct "H2" as (vl) "[??]".
iExists (_::_). rewrite heap_pointsto_vec_cons /=. iFrame=>//=.
- iIntros "H".
iDestruct "H" as ([|[[| |z]|]vl]) "[H↦ H]";
simpl; try iDestruct "H" as "[]".
rewrite heap_pointsto_vec_cons. iDestruct "H↦" as "[H↦1 H↦2]".
iSplitL "H↦1"; eauto. iExists _. iFrame. }
destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]".
iDestruct "Hx" as (vl) "[Hx↦ Hx]". rewrite uninit_own. wp_op.
iApply (type_type _ _ _
[ #lx box (uninit 1); #(lx' + 1) &uniq{α}ty]
with "[] LFT HE Hna HL HC [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Shared borrowing. *)
Definition refcell_try_borrow : val :=
fn: ["x"] :=
let: "r" := new [ #3 ] in
withcont: "k":
let: "x'" := !"x" in
let: "n" := !"x'" in
if: "n" #-1 then
"r" <-{Σ none} ();;
"k" []
else
"x'" <- "n" + #1;;
let: "ref" := new [ #2 ] in
"ref" <- "x'" + #1;;
"ref" + #1 <- "x'";;
"r" <-{2,Σ some} !"ref";;
delete [ #2; "ref"];;
"k" []
cont: "k" [] :=
delete [ #1; "x"];; return: ["r"].
Lemma refcell_try_borrow_type ty `{!TyWf ty} :
typed_val refcell_try_borrow
(fn( α, ; &shr{α}(refcell ty)) option (ref α ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply (type_cont [] [ϝ []] (λ _, [x box (&shr{α}(refcell ty));
r box (option (ref α ty))]));
[iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; simpl_subst; last first.
{ iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iApply type_deref; [solve_typing..|].
iIntros (x' tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]=>//=.
iDestruct "Hx'" as (β γ) "#[Hαβ Hinv]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[[Hβtok1 Hβtok2] Hclose']"; first done.
iMod (na_bor_acc with "LFT Hinv Hβtok1 Hna") as "(INV & Hna & Hclose'')"; try done.
iDestruct "INV" as (st) "(Hlx & Hownst & Hst)". wp_read. wp_let. wp_op; case_bool_decide; wp_if.
- iMod ("Hclose''" with "[Hlx Hownst Hst] Hna") as "[Hβtok1 Hna]";
first by iExists st; iFrame.
iMod ("Hclose'" with "[$Hβtok1 $Hβtok2]") as "Hα".
iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(refcell ty)); r box (uninit 3) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply (type_sum_unit (option $ ref α ty)); [solve_typing..|]; first last.
simpl. iApply type_jump; solve_typing.
- wp_op. wp_write. wp_apply wp_new; [done..|].
iIntros (lref) "(H† & Hlref)". wp_let.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hlref" as "[Hlref0 Hlref1]". wp_op. wp_write. wp_op. wp_write.
iAssert ( ν, ( / 2).[β] ().[ν]
ty_shr ty (β ν) tid (lx + 1)
own γ ( reading_stR ν) refcell_inv tid lx γ β ty)%I
with "[> Hlx Hownst Hst Hβtok2]" as (q' ν) "(Hβtok2 & Hν & Hshr & Hreading & INV)".
{ destruct st as [[[[ν []] s] n]|].
- simpl in *; lia.
- iDestruct "Hst" as "(H† & Hst & #Hshr)".
iDestruct "Hst" as (q' Hqq') "[Hν1 Hν2]".
iExists _, _. iFrame "Hν1". iMod (own_update with "Hownst") as "[Hownst ?]".
{ apply auth_update_alloc,
(op_local_update_discrete _ _ (reading_stR (q'/2)%Qp ν)) => ?.
split; [split|].
- by rewrite /= agree_idemp.
- apply frac_valid. rewrite /= -Hqq' comm_L.
by apply Qp.add_le_mono_l, Qp.div_le.
- done. }
iFrame "∗#". iExists (Some (ν, false, _, _)). iFrame "∗#".
rewrite [_ _]comm -Some_op -!pair_op agree_idemp. iFrame.
rewrite -(assoc Qp.add) Qp.div_2 //.
- iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]"; first done.
iMod (own_update with "Hownst") as "[Hownst Hreading]"; first by apply
auth_update_alloc, (op_local_update_discrete _ _ (reading_stR (1/2)%Qp ν)).
rewrite (right_id None). iExists _, _. iFrame "Htok1 Hreading".
iDestruct (lft_intersect_acc with "Hβtok2 Htok2") as (q) "[Htok Hclose]".
iApply (fupd_mask_mono (lftN)); first done.
iMod (rebor _ _ (β ν) with "LFT [] Hst") as "[Hst Hh]"; first done.
{ iApply lft_intersect_incl_l. }
iMod (ty_share with "LFT Hst Htok") as "[#Hshr Htok]"; first done. iFrame "Hshr".
iDestruct ("Hclose" with "Htok") as "[$ Htok2]". iExists _. iFrame "∗#".
iSplitR "Htok2".
+ iIntros "!> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro.
iNext. iMod "Hν".
iMod fupd_mask_subseteq as "Hclose"; last iMod ("Hh" with "[Hν]") as "$".
{ set_solver+. }
* rewrite -lft_dead_or. auto.
* done.
+ iExists _. iFrame. by rewrite Qp.div_2. }
iMod ("Hclose''" with "[$INV] Hna") as "[Hβtok1 Hna]".
iMod ("Hclose'" with "[$Hβtok1 $Hβtok2]") as "Hα".
iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(refcell ty)); r box (uninit 3); #lref box (ref α ty)]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame.
rewrite tctx_hasty_val' //. rewrite /= freeable_sz_full. iFrame.
iExists [_; _]. rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iFrame. simpl. iExists _, _, _, _, _. iFrame "∗#". iApply ty_shr_mono; try by auto.
iApply lft_intersect_mono; first done. iApply lft_incl_refl. }
iApply (type_sum_memcpy (option $ ref α ty)); [solve_typing..|].
simpl. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Unique borrowing. *)
Definition refcell_try_borrow_mut : val :=
fn: ["x"] :=
let: "r" := new [ #3 ] in
withcont: "k":
let: "x'" := !"x" in
let: "n" := !"x'" in
if: "n" = #0 then
"x'" <- #(-1);;
let: "ref" := new [ #2 ] in
"ref" <- "x'" + #1;;
"ref" + #1 <- "x'";;
"r" <-{2,Σ some} !"ref";;
delete [ #2; "ref"];;
"k" []
else
"r" <-{Σ none} ();;
"k" []
cont: "k" [] :=
delete [ #1; "x"];; return: ["r"].
Lemma refcell_try_borrow_mut_type ty `{!TyWf ty} :
typed_val refcell_try_borrow_mut
(fn( α, ; &shr{α}(refcell ty)) option (refmut α ty))%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply (type_cont [] [ϝ []] (λ _, [x box (&shr{α}(refcell ty));
r box (option (refmut α ty))]));
[iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg];
simpl_subst; last first.
{ iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iApply type_deref; [solve_typing..|].
iIntros (x' tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]=>//=.
iDestruct "Hx'" as (β γ) "#[Hαβ Hinv]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[Hβtok Hclose']"; first done.
iMod (na_bor_acc with "LFT Hinv Hβtok Hna") as "(INV & Hna & Hclose'')"; try done.
iDestruct "INV" as (st) "(Hlx & Hownst & Hb)". wp_read. wp_let. wp_op; case_bool_decide; wp_if.
- wp_write. wp_apply wp_new; [done..|].
iIntros (lref) "(H† & Hlref)". wp_let.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hlref" as "[Hlref0 Hlref1]". wp_op. wp_write. wp_op. wp_write.
destruct st as [[[[ν []] s] n]|]; try done.
iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]"; first done.
iMod (own_update with "Hownst") as "[Hownst ?]".
{ by eapply auth_update_alloc,
(op_local_update_discrete _ _ (refcell_st_to_R $ Some (ν, true, (1/2)%Qp, xH))). }
rewrite (right_id None).
iApply fupd_wp. iApply (fupd_mask_mono (lftN)); first done.
iMod (rebor _ _ (β ν) with "LFT [] Hb") as "[Hb Hbh]"; first done.
{ iApply lft_intersect_incl_l. }
iModIntro. iMod ("Hclose''" with "[Hlx Hownst Hbh Htok1] Hna") as "[Hβtok Hna]".
{ iExists _. iFrame. iNext. iSplitL "Hbh".
- iIntros "Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν".
iMod fupd_mask_subseteq as "Hclose"; last iMod ("Hbh" with "[Hν]") as "$".
{ set_solver+. }
* rewrite -lft_dead_or. auto.
* done.
- iSplitL; [|done]. iExists _. iFrame. by rewrite Qp.div_2. }
iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(refcell ty)); r box (uninit 3); #lref box (refmut α ty)]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame.
rewrite tctx_hasty_val' //. rewrite /= freeable_sz_full. iFrame.
iExists [_; _]. rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iFrame. simpl.
iExists _, _, _, _, _. iFrame "#∗". iApply (bor_shorten with "[] [$Hb]").
iApply lft_intersect_mono; first done. iApply lft_incl_refl. }
iApply (type_sum_memcpy (option $ refmut α ty)); [solve_typing..|].
simpl. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
- iMod ("Hclose''" with "[Hlx Hownst Hb] Hna") as "[Hβtok Hna]";
first by iExists st; iFrame.
iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(refcell ty)); r box (uninit 3) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply (type_sum_unit (option $ refmut α ty)); [solve_typing..|].
simpl. iApply type_jump; solve_typing.
Qed.
End refcell_functions.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing util.
From lrust.typing.lib.refcell Require Import refcell.
From iris.prelude Require Import options.
Section refmut.
Context `{!typeGS Σ, !refcellG Σ}.
(* The Rust type looks as follows (after some unfolding):
pub struct RefMut<'b, T: ?Sized + 'b> {
value: &'b mut T,
borrow: &'b Cell<BorrowFlag>,
}
In other words, we have a pointer to the data, and a pointer
to the refcount field of the RefCell. *)
Program Definition refmut (α : lft) (ty : type) :=
tc_opaque
{| ty_size := 2;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc lv); #(LitLoc lrc) ] =>
ν q γ β ty', &{α ν}(lv ↦∗: ty.(ty_own) tid)
α β &na{β, tid, refcell_invN}(refcell_inv tid lrc γ β ty')
q.[ν] own γ ( writing_stR q ν)
| _ => False
end;
ty_shr κ tid l :=
(lv lrc : loc),
&frac{κ} (λ q, l↦∗{q} [ #lv; #lrc])
F q, ⌜↑shrN lftN F -∗ q.[α κ]
={F}[F∖↑shrN]▷=∗ ty.(ty_shr) (α κ) tid lv q.[α κ] |}%I.
Next Obligation.
iIntros (???[|[[]|][|[[]|][]]]); try iIntros "[]". by iIntros "_".
Qed.
Next Obligation.
iIntros (α ty E κ l tid q HE) "#LFT Hb Htok".
iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done.
iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done.
destruct vl as [|[[|lv|]|][|[[|lrc|]|][]]];
try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]".
iMod (bor_exists with "LFT Hb") as (ν) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (q') "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (γ) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (β) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (ty') "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[Hb H]"; first done.
iMod (bor_sep with "LFT H") as "[Hαβ H]"; first done.
iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ $]"; first done.
iMod (bor_sep with "LFT H") as "[_ H]"; first done.
iMod (bor_sep with "LFT H") as "[H _]"; first done.
iMod (bor_fracture (λ q, (q' * q).[ν])%I with "LFT [H]") as "H"; first done.
{ by rewrite Qp.mul_1_r. }
iDestruct (frac_bor_lft_incl with "LFT H") as "#Hκν". iClear "H".
iExists _, _. iFrame "H↦". iApply delay_sharing_nested; try done.
rewrite -assoc. iApply lft_intersect_mono; first by iApply lft_incl_refl.
iApply lft_incl_glb; first done. iApply lft_incl_refl.
Qed.
Next Obligation.
iIntros (??????) "#? H". iDestruct "H" as (lv lrc) "[#Hf #H]".
iExists _, _. iSplit.
- by iApply frac_bor_shorten.
- iIntros "!> * % Htok".
iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
{ iApply lft_intersect_mono; last done. iApply lft_incl_refl. }
iMod ("H" with "[] Htok") as "Hshr"; first done. iModIntro. iNext.
iMod "Hshr" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$".
iApply ty_shr_mono; try done. iApply lft_intersect_mono; last done. iApply lft_incl_refl.
Qed.
Global Instance refmut_wf α ty `{!TyWf ty} : TyWf (refmut α ty) :=
{ ty_lfts := [α]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty α }.
Global Instance refmut_type_contractive α : TypeContractive (refmut α).
Proof. solve_type_proper. Qed.
Global Instance refmut_ne α : NonExpansive (refmut α).
Proof. apply type_contractive_ne, _. Qed.
Global Instance refmut_mono E L :
Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) refmut.
Proof.
intros α1 α2 ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (qmax qL) "HL".
iDestruct (Hty with "HL") as "#Hty". iDestruct ( with "HL") as "#Hα".
iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2.
iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro].
- done.
- iIntros (tid [|[[]|][|[[]|][]]]) "H"=>//=.
iDestruct "H" as (ν γ q' β ty') "(Hb & #H⊑ & #Hinv & Hν & Hown)".
iExists ν, γ, q', β, ty'. iFrame "∗#". iSplit.
+ iApply bor_shorten; last iApply bor_iff; last done.
* iApply lft_intersect_mono; first by iApply lft_incl_syn_sem.
iApply lft_incl_refl.
* iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]";
iExists vl; iFrame; by iApply "Ho".
+ iApply lft_incl_trans; last done. by iApply lft_incl_syn_sem.
- iIntros (κ tid l) "H /=". iDestruct "H" as (lv lrc) "H". iExists lv, lrc.
iDestruct "H" as "[$ #H]". iIntros "!> * % Htok".
iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
{ iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. }
iMod ("H" with "[] Htok") as "Hshr"; first done. iModIntro. iNext.
iMod "Hshr" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$".
iApply ty_shr_mono; try done.
+ iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl.
+ by iApply "Hs".
Qed.
Global Instance refmut_mono_flip E L :
Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) refmut.
Proof. intros ??????. by apply refmut_mono. Qed.
Lemma refmut_mono' E L α1 α2 ty1 ty2 :
lctx_lft_incl E L α2 α1 eqtype E L ty1 ty2
subtype E L (refmut α1 ty1) (refmut α2 ty2) .
Proof. intros. by eapply refmut_mono. Qed.
Global Instance refmut_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) refmut.
Proof. intros ??[]???. split; by apply refmut_mono'. Qed.
Lemma refmut_proper' E L α1 α2 ty1 ty2 :
lctx_lft_eq E L α1 α2 eqtype E L ty1 ty2
eqtype E L (refmut α1 ty1) (refmut α2 ty2).
Proof. intros. by eapply refmut_proper. Qed.
End refmut.
Global Hint Resolve refmut_mono' refmut_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree numbers.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing.
From lrust.typing.lib.refcell Require Import refcell refmut.
From iris.prelude Require Import options.
Section refmut_functions.
Context `{!typeGS Σ, !refcellG Σ}.
(* Turning a refmut into a shared borrow. *)
Definition refmut_deref : val :=
fn: ["x"] :=
let: "x'" := !"x" in
let: "r'" := !"x'" in
letalloc: "r" <- "r'" in
delete [ #1; "x"];; return: ["r"].
Lemma refmut_deref_type ty `{!TyWf ty} :
typed_val refmut_deref (fn( '(α, β), ; &shr{α}(refmut β ty)) &shr{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=.
iDestruct "Hx'" as (lv lrc) "#(Hfrac & #Hshr)".
iMod (lctx_lft_alive_tok α with "HE HL") as () "([Hα1 Hα2] & HL & Hclose')";
[solve_typing..|].
iMod (frac_bor_acc with "LFT Hfrac Hα1") as (qlx') "[H↦ Hcloseα1]"; first done.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iMod (lctx_lft_alive_tok_noend β with "HE HL") as () "(Hβ & HL & Hclose'')";
[solve_typing..|].
iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hcloseβα2]".
wp_bind (!#lx')%E. iApply (wp_step_fupd with "[Hα2β]");
[|by iApply ("Hshr" with "[] Hα2β")|]; first done.
iMod "H↦" as "[H↦1 H↦2]". wp_read. iIntros "[#Hshr' Hα2β] !>". wp_let.
iDestruct ("Hcloseβα2" with "Hα2β") as "[Hβ Hα2]".
iMod ("Hcloseα1" with "[$H↦1 $H↦2]") as "Hα1".
iMod ("Hclose''" with "Hβ HL") as "HL". iMod ("Hclose'" with "[$] HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|].
iApply (type_type _ _ _ [ x box (&shr{α}(refmut β ty)); #lv &shr{α}ty]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
iFrame. iApply (ty_shr_mono with "[] Hshr'").
iApply lft_incl_glb; last iApply lft_incl_refl. by iApply lft_incl_syn_sem. }
iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|].
iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Turning a refmut into a unique borrow. *)
Definition refmut_derefmut : val := refmut_deref.
Lemma refmut_derefmut_type ty `{!TyWf ty} :
typed_val refmut_derefmut
(fn( '(α, β), ; &uniq{α}(refmut β ty)) &uniq{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done.
iMod (bor_exists with "LFT Hx'") as (vl) "H"; first done.
iMod (bor_sep with "LFT H") as "[H↦ H]"; first done.
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose')";
[solve_typing..|].
destruct vl as [|[[|lv|]|][|[[|lrc|]|][]]]; simpl;
try by iMod (bor_persistent with "LFT H Hα") as "[>[] _]".
iMod (bor_exists with "LFT H") as (ν) "H"; first done.
iMod (bor_exists with "LFT H") as (q) "H"; first done.
iMod (bor_exists with "LFT H") as (γ) "H"; first done.
iMod (bor_exists with "LFT H") as (δ) "H"; first done.
iMod (bor_exists with "LFT H") as (ty') "H"; first done.
iMod (bor_sep with "LFT H") as "[Hb H]"; first done.
iMod (bor_sep with "LFT H") as "[Hβδ H]"; first done.
iMod (bor_persistent with "LFT Hβδ Hα") as "[#Hβδ Hα]"; first done.
iMod (bor_sep with "LFT H") as "[_ H]"; first done.
iMod (bor_sep with "LFT H") as "[H _]"; first done.
iMod (bor_fracture (λ q', (q * q').[ν])%I with "LFT [H]") as "H"; first done.
{ by rewrite Qp.mul_1_r. }
iDestruct (frac_bor_lft_incl _ _ q with "LFT H") as "#Hαν". iClear "H".
rewrite heap_pointsto_vec_cons. iMod (bor_sep with "LFT H↦") as "[H↦ _]"; first done.
iMod (bor_acc with "LFT H↦ Hα") as "[H↦ Hcloseα]"; first done.
wp_bind (!#lx')%E. iMod (bor_unnest with "LFT Hb") as "Hb"; first done.
wp_read. wp_let. iMod "Hb".
iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]". iMod ("Hclose'" with "Hα HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|].
iApply (type_type _ _ _ [ x box (uninit 1); #lv &uniq{α}ty]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
iFrame. iApply (bor_shorten with "[] Hb").
iApply lft_incl_glb; [iApply lft_incl_glb|iApply lft_incl_refl];
try by iApply lft_incl_syn_sem. done. }
iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|].
iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Dropping a refmut and set the count to 0 in the corresponding refcell. *)
Definition refmut_drop : val :=
fn: ["x"] :=
let: "rc" := !("x" + #1) in
let: "n" := !"rc" in
"rc" <- "n" + #1;;
Endlft;;
delete [ #2; "x"];;
let: "r" := new [ #0] in return: ["r"].
Lemma refmut_drop_type ty `{!TyWf ty} :
typed_val refmut_drop (fn( α, ; refmut α ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val.
destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]".
iDestruct "Hx" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hx"; simpl;
try iDestruct "Hx" as "[_ >[]]".
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hx" as "[[Hx↦1 Hx↦2] Hx]". wp_op. wp_read. wp_let.
iDestruct "Hx" as (ν q γ β ty') "(Hb & #Hαβ & #Hinv & Hν & H◯)".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)";
[solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[Hβ Hcloseα]"; first done.
iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hcloseβ)"; [done..|].
iDestruct "INV" as (st) "(H↦lrc & H● & INV)". wp_read. wp_let. wp_op. wp_write.
iAssert (|={lftN lft_userE}[lft_userE]▷=> refcell_inv tid lrc γ β ty')%I
with "[H↦lrc H● H◯ Hν INV]" as "INV".
{ iCombine "H● H◯" gives
%[[[=]|(? & [[? q'] ?] & [= <-] & Hst & INCL)]%option_included _]
%auth_both_valid_discrete.
destruct st as [[[[??]?]?]|]; [|done]. move: Hst=>-[= ???]. subst.
destruct INCL as [[[[ν' ?]%to_agree_inj ?] ?]|
[[[??]%to_agree_included [q'' Hq'']]%prod_included [n' Hn']]%prod_included].
- simpl in *. setoid_subst. iExists None. iFrame.
iMod (own_update_2 with "H● H◯") as "$".
{ apply auth_update_dealloc. rewrite -(right_id None op (Some _)).
apply cancel_local_update_unit, _. }
iDestruct "INV" as "(H† & Hq & _)".
iApply "H†".
iDestruct "Hq" as (q) "(<- & ?)". iFrame.
- simpl in *. setoid_subst. iExists (Some (_, _, _, _)).
iMod (own_update_2 with "H● H◯") as "$".
{ apply auth_update_dealloc. rewrite -(agree_idemp (to_agree _)) !pair_op Some_op.
apply (cancel_local_update_unit (writing_stR q _)), _. }
iDestruct "INV" as "(H† & Hq & _)".
rewrite /= (_:Z.neg (1%positive n') + 1 = Z.neg n');
last (rewrite pos_op_add; lia). iFrame "H↦lrc H†".
iApply step_fupd_intro; [set_solver+|]. iSplitL; [|done].
iDestruct "Hq" as (q' ?) "?". iExists (q+q')%Qp. iFrame.
rewrite assoc (comm _ q'' q) //. }
wp_bind Endlft. iApply (wp_mask_mono _ (lftN lft_userE)); first done.
iApply (wp_step_fupd with "INV"); [set_solver+|]. wp_seq. iIntros "{Hb} Hb !>".
iMod ("Hcloseβ" with "Hb Hna") as "[Hβ Hna]".
iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". wp_seq.
iApply (type_type _ _ _ [ #lx box (uninit 2)]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_singleton tctx_hasty_val' //. iFrame. iExists [ #lv;#lrc].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton uninit_own. iFrame. auto. }
iApply type_delete; [solve_typing..|].
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_jump; solve_typing.
Qed.
(* Apply a function within the refmut, typically for accessing a component. *)
Definition refmut_map (call_once : val) : val :=
fn: ["ref"; "f"] :=
let: "call_once" := call_once in
let: "x'" := !"ref" in
letalloc: "x" <- "x'" in
letcall: "r" := "call_once" ["f"; "x"]%E in
let: "r'" := !"r" in delete [ #1; "r"];;
"ref" <- "r'";;
return: ["ref"].
Lemma refmut_map_type ty1 ty2 call_once fty `{!TyWf ty1, !TyWf ty2, !TyWf fty} :
(* fty : for<'a>, FnOnce(&'a ty1) -> &'a ty2, as witnessed by the impl call_once *)
typed_val call_once (fn( α, ; fty, &uniq{α}ty1) &uniq{α}ty2)
typed_val (refmut_map call_once) (fn( α, ; refmut α ty1, fty) refmut α ty2).
Proof.
intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg).
inv_vec arg=>ref env. simpl_subst.
iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)".
rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done.
iDestruct "Href" as "[Href Href†]".
iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl;
try iDestruct "Href" as "[_ >[]]".
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Href" as "[[Href↦1 Href↦2] Href]".
iDestruct "Href" as (ν q γ β ty') "(Hbor & #Hαβ & #Hinv & >Hν & Hγ)".
wp_read. wp_let. wp_apply wp_new; [done..|].
iIntros (lx) "(H† & Hlx)". rewrite heap_pointsto_vec_singleton.
wp_let. wp_write. wp_let. rewrite tctx_hasty_val.
iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|].
iMod (lctx_lft_alive_tok_noend ϝ with "HE HL") as (?) "(Hϝ & HL & Hclose2)";[solve_typing..|].
iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]".
iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]".
rewrite -[ϝ in (α ν) ϝ](right_id_L).
iApply (type_call_iris _ [α ν; ϝ] (α ν) [_; _]
with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H† Hbor]"); [solve_typing|done| |].
{ rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full.
iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. by iFrame. }
iIntros ([[|r|]|]) "Hna Hανϝ Hr //".
iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]".
iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]".
iMod ("Hclose2" with "Hϝ HL") as "HL". iMod ("Hclose1" with "Hα HL") as "HL".
wp_rec. iDestruct "Hr" as "[Hr Hr†]".
iDestruct "Hr" as ([|r'[]]) "[Hr Hr']";
try by iDestruct (ty_size_eq with "Hr'") as "%".
rewrite heap_pointsto_vec_singleton. wp_read. wp_let.
wp_apply (wp_delete _ _ _ [_] with "[Hr Hr†]")=>//.
{ rewrite heap_pointsto_vec_singleton freeable_sz_full. iFrame. } iIntros "_".
wp_seq. wp_write.
iApply (type_type _ [_] _ [ #lref box (refmut α ty2) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_singleton tctx_hasty_val' //. simpl. iFrame.
iExists [_;_]. rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iFrame. destruct r' as [[]|]=>//=. auto 10 with iFrame. }
iApply type_jump; solve_typing.
Qed.
(* Apply a function within the refmut, and create two refmuts,
typically for splitting the reference. *)
Definition refmut_map_split (call_once : val) : val :=
fn: ["refmut"; "f"] :=
let: "call_once" := call_once in
let: "x'" := !"refmut" in
letalloc: "x" <- "x'" in
letcall: "r" := "call_once" ["f"; "x"]%E in
let: "a" := !"r" in
let: "b" := !("r" + #1) in
delete [ #2; "r"];;
let: "rc" := !("refmut" + #1) in
let: "n" := !"rc" in
"rc" <- "n" - #1;;
delete [ #2; "refmut"];;
let: "refmuts" := new [ #4 ] in
"refmuts" <- "a";;
"refmuts" + #1 <- "rc";;
"refmuts" + #2 <- "b";;
"refmuts" + #3 <- "rc";;
return: ["refmuts"].
Lemma refmut_map_split_type ty ty1 ty2 call_once fty
`{!TyWf ty, !TyWf ty1, !TyWf ty2, !TyWf fty} :
(* fty : for<'a>, FnOnce(&mut'a ty) -> (&mut'a ty1, &mut'a ty2),
as witnessed by the impl call_once *)
typed_val call_once
(fn( α, ; fty, &uniq{α}ty) Π[&uniq{α}ty1; &uniq{α}ty2])
typed_val (refmut_map_split call_once)
(fn( α, ; refmut α ty, fty) Π[refmut α ty1; refmut α ty2]).
Proof.
intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg).
inv_vec arg=>refmut env. simpl_subst.
iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Hrefmut & Henv & _)".
rewrite (tctx_hasty_val _ refmut). destruct refmut as [[|lrefmut|]|]; try done.
iDestruct "Hrefmut" as "[Hrefmut Hrefmut†]".
iDestruct "Hrefmut" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hrefmut"; simpl;
try iDestruct "Hrefmut" as "[_ >[]]".
rewrite {1}heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hrefmut" as "[[Hrefmut↦1 Hrefmut↦2] Hrefmut]".
iDestruct "Hrefmut" as (ν γ β ty') "(Hbor & #Hαβ & #Hinv & >Hν & Hγ)".
wp_read. wp_let. wp_apply wp_new; [done..|].
iIntros (lx) "(H† & Hlx)". rewrite heap_pointsto_vec_singleton.
wp_let. wp_write. wp_let. rewrite tctx_hasty_val.
iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|].
iMod (lctx_lft_alive_tok_noend ϝ with "HE HL") as (?) "(Hϝ & HL & Hclose2)";[solve_typing..|].
iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]".
iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]".
rewrite -[ϝ in (α ν) ϝ]right_id_L.
iApply (type_call_iris _ [α ν; ϝ] (α ν) [_; _]
with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H† Hbor]"); [solve_typing|done| |].
{ rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full.
iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. by iFrame. }
iIntros ([[|r|]|]) "Hna Hανϝ Hr //".
iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]".
iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]".
iMod ("Hclose2" with "Hϝ HL") as "HL".
wp_rec. iDestruct "Hr" as "[Hr Hr†]".
iDestruct "Hr" as (?) "[Hr H]". iDestruct "H" as (vl1 vl1' ->) "[Hr1' H]".
iDestruct "H" as (vl2 vl2' ->) "[Hr2' ->]".
destruct vl1 as [|[[|lr1|]|] []], vl2 as [|[[|lr2|]|] []]=>//=.
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
iDestruct "Hr" as "[Hr1 Hr2]". wp_read. wp_let. wp_op. wp_read. wp_let.
wp_apply (wp_delete _ _ _ [_; _] with "[Hr1 Hr2 Hr†]")=>//.
{ rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton /= freeable_sz_full.
iFrame. }
iIntros "_".
iMod (lft_incl_acc with "Hαβ Hα") as (?) "[Hβ Hβclose]"; first done.
iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hclosena)"; [done..|].
wp_seq. wp_op. wp_read.
iDestruct "INV" as (st) "(Hlrc & H● & Hst)".
iCombine "H● Hγ" gives %[Hst _]%auth_both_valid_discrete.
destruct st as [[[[??]?]?]|]; last by destruct Hst as [[?|] Hst]; inversion Hst.
move:Hst=>/Some_pair_included [/Some_pair_included_total_1
[/to_agree_included /(leibniz_equiv _ _) [= <- <-] _] _].
iDestruct "Hst" as "(H† & Hq & _)". iDestruct "Hq" as (q1 Hqq1) "[Hν1 Hν2]".
iMod (own_update with "H●") as "[H● ?]".
{ apply auth_update_alloc,
(op_local_update_discrete _ _ (writing_stR (q1/2)%Qp ν))=>-[Hagv _].
split; [split|done].
- by rewrite /= agree_idemp.
- apply frac_valid. rewrite /= -Hqq1 comm_L.
by apply Qp.add_le_mono_l, Qp.div_le. }
wp_let. wp_read. wp_let. wp_op. wp_write.
wp_apply (wp_delete _ _ _ [_; _] with "[Hrefmut↦1 Hrefmut↦2 Hrefmut†]")=>//.
{ rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton /= freeable_sz_full.
iFrame. }
iIntros "_". wp_seq. wp_apply wp_new; [done..|]. iIntros (lrefmuts) "[Hrefmuts† Hrefmuts]".
rewrite 3!heap_pointsto_vec_cons heap_pointsto_vec_singleton. wp_let.
iDestruct "Hrefmuts" as "(Hrefmuts1 & Hrefmuts2 & Hrefmuts3 & Hrefmuts4)".
rewrite !shift_loc_assoc. wp_write. do 3 (wp_op; wp_write).
iMod ("Hclosena" with "[Hlrc H● Hν1 H†] Hna") as "[Hβ Hna]".
{ iExists (Some (_, true, _, _)).
rewrite -Some_op -!pair_op agree_idemp /= (comm _ xH _).
iFrame.
rewrite (comm Qp.add) (assoc Qp.add) Qp.div_2 (comm Qp.add). auto. }
iMod ("Hβclose" with "Hβ") as "Hα". iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ [_] _ [ #lrefmuts box (Π[refmut α ty1; refmut α ty2]) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_singleton tctx_hasty_val' //= -freeable_sz_full.
simpl. iFrame. iExists [_;_;_;_].
rewrite 3!heap_pointsto_vec_cons heap_pointsto_vec_singleton !shift_loc_assoc.
iFrame. iExists [_;_], [_;_]. iSplit=>//=.
iSplitL "Hν Hγ Hr1'"; [by auto 10 with iFrame|]. iExists [_;_], [].
iSplitR=>//=. rewrite right_id. auto 20 with iFrame. }
iApply type_jump; solve_typing.
Qed.
End refmut_functions.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree excl numbers.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import at_borrow.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Definition rwlock_stR :=
optionUR (csumR (exclR unitO) (prodR (prodR (agreeR lftO) fracR) positiveR)).
Class rwlockG Σ :=
RwLockG :: inG Σ (authR rwlock_stR).
Definition rwlockΣ : gFunctors := #[GFunctor (authR rwlock_stR)].
Global Instance subG_rwlockΣ {Σ} : subG rwlockΣ Σ rwlockG Σ.
Proof. solve_inG. Qed.
Definition Z_of_rwlock_st (st : rwlock_stR) : Z :=
match st with
| None => 0
| Some (Cinr (_, _, n)) => Zpos n
| Some _ => -1
end.
Definition reading_st (q : frac) (ν : lft) : rwlock_stR :=
Some (Cinr (to_agree ν, q, xH)).
Definition writing_st : rwlock_stR :=
Some (Cinl (Excl ())).
Definition rwlockN := lrustN .@ "rwlock".
Section rwlock_inv.
Context `{!typeGS Σ, !rwlockG Σ}.
Definition rwlock_inv tid_own tid_shr (l : loc) (γ : gname) (α : lft) ty
: iProp Σ :=
( st, l #(Z_of_rwlock_st st) own γ ( st)
match st return _ with
| None =>
(* Not locked. *)
&{α}((l + 1) ↦∗: ty.(ty_own) tid_own)
| Some (Cinr (agν, q, n)) =>
(* Locked for read. *)
(ν : lft) q', agν to_agree ν
(1.[ν] ={lftN lft_userE}[lft_userE]▷=∗ [ν])
([ν] ={lftN}=∗ &{α}((l + 1) ↦∗: ty.(ty_own) tid_own))
ty.(ty_shr) (α ν) tid_shr (l + 1)
(q + q')%Qp = 1%Qp q'.[ν]
| _ => (* Locked for write. *) True
end)%I.
Global Instance rwlock_inv_type_ne n tid_own tid_shr l γ α :
Proper (type_dist2 n ==> dist n) (rwlock_inv tid_own tid_shr l γ α).
Proof.
solve_proper_core
ltac:(fun _ => exact: type_dist2_S || f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance rwlock_inv_ne tid_own tid_shr l γ α :
NonExpansive (rwlock_inv tid_own tid_shr l γ α).
Proof.
intros n ???. eapply rwlock_inv_type_ne, type_dist_dist2. done.
Qed.
Lemma rwlock_inv_proper E L ty1 ty2 qmax qL :
eqtype E L ty1 ty2
llctx_interp_noend qmax L qL -∗ tid_own tid_shr l γ α, (elctx_interp E -∗
rwlock_inv tid_own tid_shr l γ α ty1 -∗
rwlock_inv tid_own tid_shr l γ α ty2).
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic.
It would easily gain from some automation. *)
rewrite eqtype_unfold. iIntros (Hty) "HL".
iDestruct (Hty with "HL") as "#Hty". iIntros "%tid_own %tid_shr %l %γ %α !> #HE H".
iDestruct ("Hty" with "HE") as "(% & #Hown & #Hshr)".
iAssert ( (&{α}((l + 1) ↦∗: ty_own ty1 tid_own) -∗
&{α}((l + 1) ↦∗: ty_own ty2 tid_own)))%I as "#Hb".
{ iIntros "!> H". iApply bor_iff; last done.
iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl;
iFrame; by iApply "Hown". }
iDestruct "H" as (st) "H"; iExists st;
iDestruct "H" as "($&$&H)"; destruct st as [[|[[agν ?]?]|]|]; try done;
last by iApply "Hb".
iDestruct "H" as (ν q') "(Hag & #Hend & Hh & ? & ? & ?)". iExists ν, q'.
iFrame. iSplitR; first done. iSplitL "Hh"; last by iApply "Hshr".
iIntros "Hν". iApply "Hb". iApply ("Hh" with "Hν").
Qed.
Lemma rwlock_inv_change_tid_own tid_own1 tid_own2 tid_shr l γ α ty :
Send ty
rwlock_inv tid_own1 tid_shr l γ α ty rwlock_inv tid_own2 tid_shr l γ α ty.
Proof.
intros ?. apply bi.exist_proper=>?; do 7 f_equiv; first do 7 f_equiv.
- do 5 f_equiv. iApply send_change_tid'.
- iApply send_change_tid'.
Qed.
Lemma rwlock_inv_change_tid_shr tid_own tid_shr1 tid_shr2 l γ α ty :
Sync ty
rwlock_inv tid_own tid_shr1 l γ α ty rwlock_inv tid_own tid_shr2 l γ α ty.
Proof.
intros ?. apply bi.exist_proper=>?; do 7 f_equiv; first do 7 f_equiv.
iApply sync_change_tid'.
Qed.
End rwlock_inv.
Section rwlock.
Context `{!typeGS Σ, !rwlockG Σ}.
(* Original Rust type (we ignore poisoning):
pub struct RwLock<T: ?Sized> {
inner: Box<sys::RWLock>,
poison: poison::Flag,
data: UnsafeCell<T>,
}
*)
Program Definition rwlock (ty : type) :=
{| ty_size := S ty.(ty_size);
ty_own tid vl :=
match vl return _ with
| #(LitInt z) :: vl' => ⌜-1 z ty.(ty_own) tid vl'
| _ => False
end%I;
ty_shr κ tid l :=
( α γ, κ α &at{α,rwlockN}(rwlock_inv tid tid l γ α ty))%I |}.
Next Obligation.
iIntros (??[|[[]|]]); try iIntros "[]". rewrite ty_size_eq.
iIntros "[_ %] !% /=". congruence.
Qed.
Next Obligation.
iIntros (ty E κ l tid q ?) "#LFT Hb Htok".
iMod (bor_acc_cons with "LFT Hb Htok") as "[H Hclose]"; first done.
iDestruct "H" as ([|[[| |n]|]vl]) "[H↦ H]"; try iDestruct "H" as ">[]".
iDestruct "H" as "[>% Hown]".
iMod ("Hclose" $! (( n:Z, l #n ⌜-1 n)
(l + 1) ↦∗: ty.(ty_own) tid) with "[] [-]")%I as "[H [Htok Htok']]".
{ iIntros "!> [Hn Hvl] !>". iDestruct "Hn" as (n') "[Hn >%]".
iDestruct "Hvl" as (vl') "[H↦ Hvl]".
iExists (#n'::vl'). rewrite heap_pointsto_vec_cons. iFrame "∗%". }
{ iNext. rewrite heap_pointsto_vec_cons. iDestruct "H↦" as "[Hn Hvl]".
iSplitL "Hn"; [eauto|iExists _; iFrame]. }
iMod (bor_sep with "LFT H") as "[Hn Hvl]"; first done.
iMod (bor_acc_cons with "LFT Hn Htok") as "[H Hclose]"; first done.
iAssert ((q / 2).[κ] γ, rwlock_inv tid tid l γ κ ty)%I with "[> -Hclose]"
as "[$ HQ]"; last first.
{ iMod ("Hclose" with "[] HQ") as "[Hb $]".
- iIntros "!> H !>". iNext. iDestruct "H" as (γ st) "(H & _ & _)".
iExists _. iIntros "{$H}!%". destruct st as [[|[[]?]|]|]; simpl; lia.
- iMod (bor_exists with "LFT Hb") as (γ) "Hb"; first done.
iExists κ, γ. iSplitR; first by iApply lft_incl_refl. iApply bor_share; try done.
solve_ndisj. }
clear dependent n. iDestruct "H" as ([|n|[]]) "[Hn >%]"; try lia.
- iFrame. iMod (own_alloc ( None)) as (γ) "Hst"; first by apply auth_auth_valid.
iExists γ, None. by iFrame.
- iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]"; first done.
iMod (own_alloc ( Some (Cinr (to_agree ν, (1/2)%Qp, n)))) as (γ) "Hst".
{ by apply auth_auth_valid. }
iMod (rebor _ _ (κ ν) with "LFT [] Hvl") as "[Hvl Hh]"; first done.
{ iApply lft_intersect_incl_l. }
iDestruct (lft_intersect_acc with "Htok' Htok1") as (q') "[Htok Hclose]".
iMod (ty_share with "LFT Hvl Htok") as "[Hshr Htok]"; first done.
iDestruct ("Hclose" with "Htok") as "[$ Htok]".
iExists γ, _. iFrame "Hst Hn". iExists _, _. iIntros "{$Hshr}".
iSplitR; first by auto. iFrame "Htok2". iSplitR; first done.
rewrite Qp.div_2. iSplitL; last by auto.
iIntros "!> !> Hν". iDestruct (lft_tok_dead with "Htok Hν") as "[]".
- iMod (own_alloc ( writing_st)) as (γ) "Hst". { by apply auth_auth_valid. }
by iFrame.
Qed.
Next Obligation.
iIntros (?????) "#Hκ H". iDestruct "H" as (α γ) "[#??]".
iExists _, _. iFrame. iApply lft_incl_trans; auto.
Qed.
Global Instance rwlock_wf ty `{!TyWf ty} : TyWf (rwlock ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance rwlock_type_ne : TypeNonExpansive rwlock.
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || (eapply rwlock_inv_type_ne; try reflexivity) ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance rwlock_ne : NonExpansive rwlock.
Proof.
constructor; solve_proper_core ltac:(fun _ => (eapply ty_size_ne; try reflexivity) || f_equiv).
Qed.
Global Instance rwlock_mono E L : Proper (eqtype E L ==> subtype E L) rwlock.
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic.
It would easily gain from some automation. *)
iIntros (ty1 ty2 EQ qmax qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'.
iDestruct (EQ' with "HL") as "#EQ'".
iDestruct (rwlock_inv_proper with "HL") as "#Hty1ty2"; first done.
iDestruct (rwlock_inv_proper with "HL") as "#Hty2ty1"; first by symmetry.
iIntros "!> #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)".
iSplit; [|iSplit].
- iPureIntro. simpl. congruence.
- iIntros "!> %tid %vl H". destruct vl as [|[[]|]]; try done.
iDestruct "H" as "[$ H]". by iApply "Hown".
- iIntros "!> %α %tid %l H". simpl.
iDestruct "H" as (a γ) "[Ha H]". iExists a, γ. iFrame.
iApply at_bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H".
+ by iApply "Hty1ty2".
+ by iApply "Hty2ty1".
Qed.
Lemma rwlock_mono' E L ty1 ty2 :
eqtype E L ty1 ty2 subtype E L (rwlock ty1) (rwlock ty2).
Proof. eapply rwlock_mono. Qed.
Global Instance rwlock_proper E L : Proper (eqtype E L ==> eqtype E L) rwlock.
Proof. by split; apply rwlock_mono. Qed.
Lemma rwlock_proper' E L ty1 ty2 :
eqtype E L ty1 ty2 eqtype E L (rwlock ty1) (rwlock ty2).
Proof. eapply rwlock_proper. Qed.
(* Apparently, we don't need to require ty to be sync,
contrarily to Rust's implementation. *)
Global Instance rwlock_send ty :
Send ty Send (rwlock ty).
Proof. move=>???[|[[]|]]//=??. iIntros "[$?]". by iApply send_change_tid. Qed.
Global Instance rwlock_sync ty :
Send ty Sync ty Sync (rwlock ty).
Proof.
move=>??????/=. do 2 apply bi.exist_mono=>?. apply bi.sep_mono_r.
apply bi.equiv_entails. f_equiv.
by rewrite rwlock_inv_change_tid_own rwlock_inv_change_tid_shr.
Qed.
End rwlock.
Global Hint Resolve rwlock_mono' rwlock_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lang.lib Require Import memcpy.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing option.
From lrust.typing.lib.rwlock Require Import rwlock rwlockreadguard rwlockwriteguard.
From iris.prelude Require Import options.
Section rwlock_functions.
Context `{!typeGS Σ, !rwlockG Σ}.
(* Constructing a rwlock. *)
Definition rwlock_new ty : val :=
fn: ["x"] :=
let: "r" := new [ #(S ty.(ty_size))] in
"r" + #0 <- #0;;
"r" + #1 <-{ty.(ty_size)} !"x";;
delete [ #ty.(ty_size) ; "x"];; return: ["r"].
Lemma rwlock_new_type ty `{!TyWf ty} :
typed_val (rwlock_new ty) (fn(; ty) rwlock ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new (S ty.(ty_size))); [solve_typing..|].
iIntros (r tid qmax) "#LFT HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done.
iDestruct "Hx" as "[Hx Hx†]". iDestruct "Hx" as (vl) "[Hx↦ Hx]".
destruct r as [[|lr|]|]; try done. iDestruct "Hr" as "[Hr Hr†]".
iDestruct "Hr" as (vl') "Hr". rewrite uninit_own.
iDestruct "Hr" as "[Hr↦ >SZ]". destruct vl' as [|]; iDestruct "SZ" as %[=].
rewrite heap_pointsto_vec_cons. iDestruct "Hr↦" as "[Hr↦0 Hr↦1]". wp_op.
rewrite shift_loc_0. wp_write. wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz.
wp_apply (wp_memcpy with "[$Hr↦1 $Hx↦]"); [done||lia..|].
iIntros "[Hr↦1 Hx↦]". wp_seq.
iApply (type_type _ _ _ [ #lx box (uninit (ty_size ty)); #lr box (rwlock ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=. iFrame.
iSplitR; first done.
iExists (_::_). rewrite heap_pointsto_vec_cons. iFrame. simpl. iFrame. auto. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* The other direction: getting ownership out of a rwlock.
Note: as we ignore poisonning, this cannot fail. *)
Definition rwlock_into_inner ty : val :=
fn: ["x"] :=
let: "r" := new [ #ty.(ty_size)] in
"r" <-{ty.(ty_size)} !("x" + #1);;
delete [ #(S ty.(ty_size)) ; "x"];; return: ["r"].
Lemma rwlock_into_inner_type ty `{!TyWf ty} :
typed_val (rwlock_into_inner ty) (fn(; rwlock ty) ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new ty.(ty_size)); [solve_typing..|].
iIntros (r tid qmax) "#LFT HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done.
iDestruct "Hx" as "[Hx Hx†]".
iDestruct "Hx" as ([|[[]|]vl]) "[Hx↦ Hx]"; try iDestruct "Hx" as ">[]".
destruct r as [[|lr|]|]; try done. iDestruct "Hr" as "[Hr Hr†]".
iDestruct "Hr" as (vl') "Hr". rewrite uninit_own heap_pointsto_vec_cons.
iDestruct "Hr" as "[Hr↦ >%]". iDestruct "Hx↦" as "[Hx↦0 Hx↦1]". wp_op.
iDestruct "Hx" as "[% Hx]". iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz.
wp_apply (wp_memcpy with "[$Hr↦ $Hx↦1]"); [done||lia..|].
iIntros "[Hr↦ Hx↦1]". wp_seq.
iApply (type_type _ _ _ [ #lx box (uninit (S (ty_size ty))); #lr box ty]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame.
iExists (_::_). rewrite heap_pointsto_vec_cons uninit_own -Hsz. iFrame. auto. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rwlock_get_mut : val :=
fn: ["x"] :=
let: "x'" := !"x" in
"x" <- "x'" + #1;;
return: ["x"].
Lemma rwlock_get_mut_type ty `{!TyWf ty} :
typed_val rwlock_get_mut (fn( α, ; &uniq{α} (rwlock ty)) &uniq{α} ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iIntros (tid qmax) "#LFT HE Hna HL HC HT".
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try iDestruct "Hx'" as "[]".
iAssert (&{α} ( (z : Z), lx' #z ⌜-1 z)
(&uniq{α} ty).(ty_own) tid [ #(lx' + 1)])%I with "[> Hx']" as "[_ Hx']".
{ iApply bor_sep; [done..|]. iApply (bor_proper with "Hx'"). iSplit.
- iIntros "[H1 H2]". iDestruct "H1" as (z) "[??]". iDestruct "H2" as (vl) "[??]".
iExists (_::_). rewrite heap_pointsto_vec_cons. iFrame. iFrame.
- iIntros "H". iDestruct "H" as ([|[[| |z]|]vl]) "[H↦ H]"; try done.
rewrite heap_pointsto_vec_cons.
iDestruct "H" as "[H1 H2]". iDestruct "H↦" as "[H↦1 H↦2]".
iSplitL "H1 H↦1"; eauto. iExists _. iFrame. }
destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]".
iDestruct "Hx" as (vl) "[Hx↦ Hx]". rewrite uninit_own. wp_op.
iApply (type_type _ _ _
[ #lx box (uninit 1); #(lx' + 1) &uniq{α}ty]
with "[] LFT HE Hna HL HC [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Acquiring a read lock. *)
Definition rwlock_try_read : val :=
fn: ["x"] :=
let: "r" := new [ #2 ] in
let: "x'" := !"x" in
withcont: "k":
withcont: "loop":
"loop" []
cont: "loop" [] :=
let: "n" := !ˢᶜ"x'" in
if: "n" #-1 then
"r" <-{Σ none} ();;
"k" []
else
if: CAS "x'" "n" ("n" + #1) then
"r" <-{Σ some} "x'";;
"k" []
else "loop" []
cont: "k" [] :=
delete [ #1; "x"];; return: ["r"].
Lemma rwlock_try_read_type ty `{!TyWf ty} :
typed_val rwlock_try_read
(fn( α, ; &shr{α}(rwlock ty)) option (rwlockreadguard α ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iApply (type_cont [] [ϝ []] (λ _, [x box (&shr{α}(rwlock ty));
r box (option (rwlockreadguard α ty))]));
[iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg];
simpl_subst; last first.
{ iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iApply (type_cont [] [ϝ []] (λ _, [x _; x' _; r _]));
[iIntros (loop)|iIntros "/= !>"; iIntros (loop arg); inv_vec arg];
simpl_subst.
{ iApply type_jump; solve_typing. }
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT".
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]; try done.
iDestruct "Hx'" as (β γ) "#[Hαβ Hinv]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)";
[solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[[Hβtok1 Hβtok2] Hclose']"; first done.
wp_bind (!ˢᶜ(LitV lx))%E.
iMod (at_bor_acc_tok with "LFT Hinv Hβtok1") as "[INV Hclose'']"; try done.
iDestruct "INV" as (st) "(Hlx & INV)". wp_read.
iMod ("Hclose''" with "[Hlx INV]") as "Hβtok1"; first by iExists _; iFrame.
iModIntro. wp_let. wp_op; case_bool_decide as Hm1; wp_if.
- iMod ("Hclose'" with "[$]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(rwlock ty)); r box (uninit 2) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply (type_sum_unit (option $ rwlockreadguard α ty));
[solve_typing..|]; first last.
simpl. iApply type_jump; solve_typing.
- wp_op. wp_bind (CAS _ _ _).
iMod (at_bor_acc_tok with "LFT Hinv Hβtok1") as "[INV Hclose'']"; try done.
iDestruct "INV" as (st') "(Hlx & Hownst & Hst)". revert Hm1.
destruct (decide (Z_of_rwlock_st st = Z_of_rwlock_st st')) as [->|?]=>?.
+ iApply (wp_cas_int_suc with "Hlx"). iNext. iIntros "Hlx".
iAssert ( ν, ( / 2).[β] ().[ν]
ty_shr ty (β ν) tid (lx + 1)
own γ ( reading_st ν) rwlock_inv tid tid lx γ β ty
((1).[ν] ={lftN lft_userE}[lft_userE]▷=∗ [ν]))%I
with "[> Hlx Hownst Hst Hβtok2]"
as (q' ν) "(Hβtok2 & Hν & Hshr & Hreading & INV & H†)".
{ destruct st' as [[|[[agν q] n]|]|]; try done.
- iDestruct "Hst" as (ν q') "(Hag & #H† & Hh & #Hshr & #Hqq' & [Hν1 Hν2])".
iExists _, _. iFrame "Hν1". iDestruct "Hag" as %Hag. iDestruct "Hqq'" as %Hqq'.
iMod (own_update with "Hownst") as "[Hownst ?]".
{ apply auth_update_alloc,
(op_local_update_discrete _ _ (reading_st (q'/2)%Qp ν))=>-[Hagv _].
split; [split|].
- by rewrite /= Hag agree_idemp.
- apply frac_valid. rewrite /= -Hqq' comm_L.
by apply Qp.add_le_mono_l, Qp.div_le.
- done. }
iFrame "∗#". rewrite Z.add_comm /=. iFrame.
iSplitR; first by rewrite /= Hag agree_idemp.
rewrite (comm Qp.add) (assoc Qp.add) Qp.div_2 (comm Qp.add). auto.
- iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]"; first solve_ndisj.
iMod (own_update with "Hownst") as "[Hownst Hreading]"; first by apply
auth_update_alloc, (op_local_update_discrete _ _ (reading_st (1/2)%Qp ν)).
rewrite (right_id None). iExists _, _. iFrame "Htok1 Hreading".
iDestruct (lft_intersect_acc with "Hβtok2 Htok2") as (q) "[Htok Hclose]".
iApply (fupd_mask_mono (lftN)); first solve_ndisj.
iMod (rebor _ _ (β ν) with "LFT [] Hst") as "[Hst Hh]"; first solve_ndisj.
{ iApply lft_intersect_incl_l. }
iMod (ty_share with "LFT Hst Htok") as "[#Hshr Htok]"; first solve_ndisj.
iFrame "#". iDestruct ("Hclose" with "Htok") as "[$ Htok2]".
iExists _. iFrame. iExists _, _. iSplitR; first done. iFrame "#∗".
rewrite Qp.div_2. iSplitL; last done.
iIntros "!> Hν". iApply "Hh". rewrite -lft_dead_or. auto. }
iMod ("Hclose''" with "[$INV]") as "Hβtok1". iModIntro. wp_case.
iMod ("Hclose'" with "[$]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(rwlock ty)); r box (uninit 2);
#lx rwlockreadguard α ty]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. iFrame "∗#".
iApply ty_shr_mono; last done.
iApply lft_intersect_mono; first done. iApply lft_incl_refl. }
iApply (type_sum_assign (option $ rwlockreadguard α ty)); [solve_typing..|].
simpl. iApply type_jump; solve_typing.
+ iApply (wp_cas_int_fail with "Hlx"); try done. iNext. iIntros "Hlx".
iMod ("Hclose''" with "[Hlx Hownst Hst]") as "Hβtok1"; first by iExists _; iFrame.
iModIntro. wp_case. iMod ("Hclose'" with "[$]") as "Hα".
iMod ("Hclose" with "Hα HL") as "HL".
iSpecialize ("Hk" with "[]"); first solve_typing.
iApply ("Hk" $! [#] with "Hna HL").
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame.
iExists _. iSplit; first done. simpl. eauto.
Qed.
(* Acquiring a write lock. *)
Definition rwlock_try_write : val :=
fn: ["x"] :=
withcont: "k":
let: "r" := new [ #2 ] in
let: "x'" := !"x" in
if: CAS "x'" #0 #(-1) then
"r" <-{Σ some} "x'";;
"k" ["r"]
else
"r" <-{Σ none} ();;
"k" ["r"]
cont: "k" ["r"] :=
delete [ #1; "x"];; return: ["r"].
Lemma rwlock_try_write_type ty `{!TyWf ty} :
typed_val rwlock_try_write
(fn( α, ; &shr{α}(rwlock ty)) option (rwlockwriteguard α ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_cont [_] [ϝ []] (λ r, [x box (&shr{α}(rwlock ty));
(r!!!0%fin:val) box (option (rwlockwriteguard α ty))]));
[iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg=>r];
simpl_subst; last first.
{ iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. }
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_deref; [solve_typing..|].
iIntros (x' tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]; try done.
iDestruct "Hx'" as (β γ) "#[Hαβ Hinv]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[Hβtok Hclose']"; first done.
wp_bind (CAS _ _ _).
iMod (at_bor_acc_tok with "LFT Hinv Hβtok") as "[INV Hclose'']"; try done.
iDestruct "INV" as (st) "(Hlx & >Hownst & Hst)". destruct st as [c|].
- iApply (wp_cas_int_fail with "Hlx"); first by destruct c as [|[[]]|].
iNext. iIntros "Hlx".
iMod ("Hclose''" with "[Hlx Hownst Hst]") as "Hβtok"; first by iExists _; iFrame.
iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iModIntro. wp_case.
iApply (type_type _ _ _
[ x box (&shr{α}(rwlock ty)); r box (uninit 2) ]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply (type_sum_unit (option $ rwlockwriteguard α ty));
[solve_typing..|]; first last.
rewrite /option /=. iApply type_jump; solve_typing.
- iApply (wp_cas_int_suc with "Hlx"). iIntros "!> Hlx".
iMod (own_update with "Hownst") as "[Hownst ?]".
{ by eapply auth_update_alloc, (op_local_update_discrete _ _ writing_st). }
iMod ("Hclose''" with "[Hlx Hownst]") as "Hβtok". { iExists _. iFrame. }
iModIntro. wp_case. iMod ("Hclose'" with "Hβtok") as "Hα".
iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _
[ x box (&shr{α}(rwlock ty)); r box (uninit 2);
#lx rwlockwriteguard α ty]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. iFrame "∗#". }
iApply (type_sum_assign (option $ rwlockwriteguard α ty)); [solve_typing..|].
simpl. iApply type_jump; solve_typing.
Qed.
End rwlock_functions.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing.
From lrust.typing.lib.rwlock Require Import rwlock.
From iris.prelude Require Import options.
Section rwlockreadguard.
Context `{!typeGS Σ, !rwlockG Σ}.
(* Original Rust type:
pub struct RwLockReadGuard<'a, T: ?Sized + 'a> {
__lock: &'a RwLock<T>,
}
*)
Program Definition rwlockreadguard (α : lft) (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] =>
ν q γ β tid_own, ty.(ty_shr) (α ν) tid (l + 1)
α β &at{β,rwlockN}(rwlock_inv tid_own tid l γ β ty)
q.[ν] own γ ( reading_st q ν)
(1.[ν] ={lftN lft_userE}[lft_userE]▷=∗ [ν])
| _ => False
end;
ty_shr κ tid l :=
(l' : loc),
&frac{κ} (λ q, l↦∗{q} [ #l'])
ty.(ty_shr) (α κ) tid (l' + 1) |}%I.
Next Obligation. intros α ty tid [|[[]|] []]; auto. Qed.
Next Obligation.
iIntros (α ty E κ l tid q ?) "#LFT Hb Htok".
iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done.
iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done.
destruct vl as [|[[|l'|]|][]];
try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]".
iMod (bor_exists with "LFT Hb") as (ν) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (q') "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (γ) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (β) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (tid_own) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[Hshr Hb]"; first done.
iMod (bor_persistent with "LFT Hshr Htok") as "[#Hshr Htok]"; first done.
iMod (bor_sep with "LFT Hb") as "[Hαβ Hb]"; first done.
iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ Htok]"; first done.
iMod (bor_sep with "LFT Hb") as "[Hinv Hb]"; first done.
iMod (bor_persistent with "LFT Hinv Htok") as "[#Hinv $]"; first done.
iMod (bor_sep with "LFT Hb") as "[Hκν _]"; first done.
iDestruct (frac_bor_lft_incl with "LFT [> Hκν]") as "#Hκν".
{ iApply bor_fracture; try done. by rewrite Qp.mul_1_r. }
iExists _. iFrame "#". iApply ty_shr_mono; last done.
iApply lft_intersect_mono; last done. iApply lft_incl_refl.
Qed.
Next Obligation.
iIntros (α ty κ κ' tid l) "#Hκ'κ H". iDestruct "H" as (l') "[#Hf #Hshr]".
iExists l'. iSplit; first by iApply frac_bor_shorten.
iApply ty_shr_mono; last done. iApply lft_intersect_mono; last done.
iApply lft_incl_refl.
Qed.
Global Instance rwlockreadguard_wf α ty `{!TyWf ty} : TyWf (rwlockreadguard α ty) :=
{ ty_lfts := [α]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty α }.
Global Instance rwlockreadguard_type_contractive α : TypeContractive (rwlockreadguard α).
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || (eapply rwlock_inv_type_ne; try reflexivity) ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance rwlockreadguard_ne α : NonExpansive (rwlockreadguard α).
Proof. apply type_contractive_ne, _. Qed.
(* The rust type is not covariant, although it probably could (cf. refcell).
This would require changing the definition of the type, though. *)
Global Instance rwlockreadguard_mono E L :
Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockreadguard.
Proof.
iIntros (α1 α2 ty1 ty2 Hty qmax qL) "HL".
iDestruct (proj1 Hty with "HL") as "#Hty". iDestruct ( with "HL") as "#Hα".
iDestruct (rwlock_inv_proper with "HL") as "#Hty1ty2"; first done.
iDestruct (rwlock_inv_proper with "HL") as "#Hty2ty1"; first by symmetry.
iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2.
iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro].
- done.
- iIntros (tid [|[[]|][]]) "H"; try done.
iDestruct "H" as (ν q' γ β tid_own) "(#Hshr & #H⊑ & #Hinv & Htok & Hown)".
iExists ν, q', γ, β, tid_own. iFrame "∗#". iSplit; last iSplit.
+ iApply ty_shr_mono; last by iApply "Hs".
iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl.
+ iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done.
+ iApply (at_bor_iff with "[] Hinv").
iIntros "!> !>"; iSplit; iIntros "H".
* by iApply "Hty1ty2".
* by iApply "Hty2ty1".
- iIntros (κ tid l) "H". iDestruct "H" as (l') "[Hf Hshr]". iExists l'.
iFrame. iApply ty_shr_mono; last by iApply "Hs".
iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl.
Qed.
Global Instance rwlockreadguard_mono_flip E L :
Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L))
rwlockreadguard.
Proof. intros ??????. by apply rwlockreadguard_mono. Qed.
Lemma rwlockreadguard_mono' E L α1 α2 ty1 ty2 :
lctx_lft_incl E L α2 α1 eqtype E L ty1 ty2
subtype E L (rwlockreadguard α1 ty1) (rwlockreadguard α2 ty2).
Proof. intros. by eapply rwlockreadguard_mono. Qed.
Global Instance rwlockreadguard_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) rwlockreadguard.
Proof. intros ??[]?? EQ. split; apply rwlockreadguard_mono'; try done; apply EQ. Qed.
Lemma rwlockreadguard_proper' E L α1 α2 ty1 ty2 :
lctx_lft_eq E L α1 α2 eqtype E L ty1 ty2
eqtype E L (rwlockreadguard α1 ty1) (rwlockreadguard α2 ty2).
Proof. intros. by eapply rwlockreadguard_proper. Qed.
(* Rust requires the type to also be Send. Not sure why. *)
Global Instance rwlockreadguard_sync α ty :
Sync ty Sync (rwlockreadguard α ty).
Proof.
move=>?????/=. apply bi.exist_mono=>?. by rewrite sync_change_tid.
Qed.
(* POSIX requires the unlock to occur from the thread that acquired
the lock, so Rust does not implement Send for RwLockWriteGuard. We can
prove this for our spinlock implementation, however. *)
Global Instance rwlockreadguard_send α ty :
Sync ty Send (rwlockreadguard α ty).
Proof.
iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H".
iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_entails.
repeat lazymatch goal with
| |- (ty_shr _ _ _ _) (ty_shr _ _ _ _) => by apply sync_change_tid'
| |- (rwlock_inv _ _ _ _ _ _) _ => by apply rwlock_inv_change_tid_shr
| |- _ => f_equiv
end.
Qed.
End rwlockreadguard.
Global Hint Resolve rwlockreadguard_mono' rwlockreadguard_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree numbers.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import lifetime na_borrow.
From lrust.typing Require Import typing.
From lrust.typing.lib.rwlock Require Import rwlock rwlockreadguard.
From iris.prelude Require Import options.
Section rwlockreadguard_functions.
Context `{!typeGS Σ, !rwlockG Σ}.
(* Turning a rwlockreadguard into a shared borrow. *)
Definition rwlockreadguard_deref : val :=
fn: ["x"] :=
let: "x'" := !"x" in
let: "r'" := !"x'" + #1 in
letalloc: "r" <- "r'" in
delete [ #1; "x"];; return: ["r"].
Lemma rwlockreadguard_deref_type ty `{!TyWf ty} :
typed_val rwlockreadguard_deref
(fn( '(α, β), ; &shr{α}(rwlockreadguard β ty)) &shr{α} ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done.
iDestruct "Hx'" as (l') "#[Hfrac Hshr]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (frac_bor_acc with "LFT Hfrac Hα") as (qlx') "[H↦ Hcloseα]"; first done.
rewrite heap_pointsto_vec_singleton. wp_read. wp_op. wp_let.
iMod ("Hcloseα" with "[$H↦]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|].
iApply (type_type _ _ _ [ x box (&shr{α}(rwlockreadguard β ty));
#(l' + 1) &shr{α}ty]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
iFrame. iApply (ty_shr_mono with "[] Hshr").
iApply lft_incl_glb; first by iApply lft_incl_syn_sem.
by iApply lft_incl_refl. }
iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|].
iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Dropping a rwlockreadguard and releasing the corresponding lock. *)
Definition rwlockreadguard_drop : val :=
fn: ["x"] :=
let: "x'" := !"x" in
withcont: "loop":
"loop" []
cont: "loop" [] :=
let: "n" := !ˢᶜ"x'" in
if: CAS "x'" "n" ("n" - #1) then
delete [ #1; "x"];;
let: "r" := new [ #0] in return: ["r"]
else "loop" [].
Lemma rwlockreadguard_drop_type ty `{!TyWf ty} :
typed_val rwlockreadguard_drop (fn( α, ; rwlockreadguard α ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iApply (type_cont [] [ϝ []] (λ _, [x _; x' _]));
[iIntros (loop)|iIntros "/= !>"; iIntros (loop arg); inv_vec arg];
simpl_subst.
{ iApply type_jump; solve_typing. }
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT".
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']".
destruct x' as [[|lx'|]|]; try done. simpl.
iDestruct "Hx'" as (ν q γ β tid_own) "(Hx' & #Hαβ & #Hinv & Hν & H◯ & H†)".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[Hβ Hcloseα]"; first done.
wp_bind (!ˢᶜ#lx')%E.
iMod (at_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|].
iDestruct "INV" as (st) "[H↦ INV]". wp_read.
iMod ("Hcloseβ" with "[H↦ INV]") as "Hβ"; first by iExists _; iFrame.
iModIntro. wp_let. wp_op. wp_bind (CAS _ _ _).
iMod (at_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|].
iDestruct "INV" as (st') "(Hlx & >H● & Hst)".
destruct (decide (Z_of_rwlock_st st = Z_of_rwlock_st st')) as [->|?].
+ iAssert (|={ rwlockN}[ rwlockN lftN lft_userE]▷=>
(lx' #(Z_of_rwlock_st st'-1) ==∗ rwlock_inv tid_own tid lx' γ β ty))%I
with "[H● H◯ Hx' Hν Hst H†]" as "INV".
{ iCombine "H● H◯" gives %[[[=]| (? & st0 & [=<-] & -> & [Heq|Hle])]
%option_included Hv]%auth_both_valid_discrete.
- destruct st0 as [|[[agν q']n]|]; try by inversion Heq. revert Heq. simpl.
intros [[EQ <-%leibniz_equiv]%(inj2 pair) <-%leibniz_equiv]
%(inj Cinr)%(inj2 pair).
iDestruct "Hst" as (ν' q') "(>EQν & _ & Hh & _ & >Hq & >Hν')".
rewrite -EQ. iDestruct "EQν" as %<-%(inj to_agree)%leibniz_equiv.
iCombine "Hν" "Hν'" as "Hν". iDestruct "Hq" as %->.
iApply (step_fupd_mask_mono ((lftN lft_userE) ( rwlockN lftN lft_userE)));
last iApply (step_fupd_mask_frame_r _ (lft_userE)); [solve_ndisj..|].
iMod ("H†" with "Hν") as "H†". iModIntro. iNext. iMod "H†".
iMod fupd_mask_subseteq as "Hclose"; last iMod ("Hh" with "H†") as "Hb".
{ set_solver+. }
iMod "Hclose" as "_". iIntros "!> Hlx". iExists None. iFrame.
iApply (own_update_2 with "H● H◯"). apply auth_update_dealloc.
rewrite -(right_id None op (Some _)). apply cancel_local_update_unit, _.
- iApply step_fupd_intro; first set_solver. iNext. iIntros "Hlx".
apply csum_included in Hle.
destruct Hle as [|[(?&?&[=]&?)|(?&[[agν q']n]&[=<-]&->&Hle%prod_included)]];
[by subst|].
destruct Hle as [[Hag [q0 Hq]]%prod_included Hn%pos_included].
iDestruct "Hst" as (ν' q'') "(EQν & H†' & Hh & Hshr & Hq & Hν')".
iDestruct "EQν" as %EQν. revert Hag Hq. rewrite /= EQν to_agree_included.
intros <-%leibniz_equiv ->%leibniz_equiv.
iExists (Some (Cinr (to_agree ν, q0, Pos.pred n))).
iSplitL "Hlx"; first by destruct n.
replace (q q0 + q'')%Qp with (q0 + (q + q''))%Qp by
by rewrite (comm _ q q0) assoc. iCombine "Hν" "Hν'" as "Hν".
iSplitL "H● H◯"; last by eauto with iFrame.
iApply (own_update_2 with "H● H◯"). apply auth_update_dealloc.
assert (n = 1%positive Pos.pred n) as EQn.
{ rewrite pos_op_add -Pplus_one_succ_l Pos.succ_pred // =>?. by subst. }
rewrite {1}EQn -{1}(agree_idemp (to_agree _)) 2!pair_op Cinr_op Some_op.
apply (cancel_local_update_unit (reading_st q ν)) , _. }
iApply (wp_step_fupd with "INV"); first set_solver.
iApply (wp_cas_int_suc with "Hlx"); try done. iNext. iIntros "Hlx INV !>".
iMod ("INV" with "Hlx") as "INV". iMod ("Hcloseβ" with "[$INV]") as "Hβ".
iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iModIntro. wp_if.
iApply (type_type _ _ _ [ x box (uninit 1)]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_singleton tctx_hasty_val //. }
iApply type_delete; [solve_typing..|].
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_jump; solve_typing.
+ iApply (wp_cas_int_fail with "Hlx"); try done. iNext. iIntros "Hlx".
iMod ("Hcloseβ" with "[Hlx H● Hst]") as "Hβ".
{ by iExists _; iFrame. }
iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
iModIntro. wp_if.
iApply (type_type _ _ _ [ x box (uninit 1); #lx' rwlockreadguard α ty]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val
tctx_hasty_val' //=; simpl. auto 10 with iFrame. }
iApply type_jump; solve_typing.
Qed.
End rwlockreadguard_functions.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import util typing.
From lrust.typing.lib.rwlock Require Import rwlock.
From iris.prelude Require Import options.
Section rwlockwriteguard.
Context `{!typeGS Σ, !rwlockG Σ}.
(* Original Rust type (we ignore poisoning):
pub struct RwLockWriteGuard<'a, T: ?Sized + 'a> {
__lock: &'a RwLock<T>,
__poison: poison::Guard,
}
*)
Program Definition rwlockwriteguard (α : lft) (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] =>
γ β tid_shr, &{β}((l + 1) ↦∗: ty.(ty_own) tid)
α β &at{β,rwlockN}(rwlock_inv tid tid_shr l γ β ty)
own γ ( writing_st)
| _ => 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' + 1) q.[α κ] |}%I.
Next Obligation. by iIntros (???[|[[]|][]]) "?". Qed.
Next Obligation.
iIntros (α ty E κ l tid q HE) "#LFT Hb Htok".
iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done.
iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done.
destruct vl as [|[[|l'|]|][]];
try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]".
iMod (bor_exists with "LFT Hb") as (γ) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (β) "Hb"; first done.
iMod (bor_exists with "LFT Hb") as (tid_shr) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[Hb H]"; first done.
iMod (bor_sep with "LFT H") as "[Hαβ _]"; first done.
iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ $]"; first done.
iExists _. iFrame "H↦". iApply delay_sharing_nested; try done.
(* FIXME: "iApply lft_intersect_mono" only preserves the later on the last
goal, as does "iApply (lft_intersect_mono with ">")". *)
iNext. iApply lft_intersect_mono; first done. iApply lft_incl_refl.
Qed.
Next Obligation.
iIntros (??????) "#? H". iDestruct "H" as (l') "[#Hf #H]".
iExists _. iSplit.
- by iApply frac_bor_shorten.
- iIntros "!> * % Htok".
iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
{ iApply lft_intersect_mono; last done. iApply lft_incl_refl. }
iMod ("H" with "[] Htok") as "Hshr"; first done. iModIntro. iNext.
iMod "Hshr" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$".
iApply ty_shr_mono; try done. iApply lft_intersect_mono; last done. iApply lft_incl_refl.
Qed.
Global Instance rwlockwriteguard_wf α ty `{!TyWf ty} : TyWf (rwlockwriteguard α ty) :=
{ ty_lfts := [α]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty α }.
Global Instance rwlockwriteguard_type_contractive α : TypeContractive (rwlockwriteguard α).
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || (eapply rwlock_inv_type_ne; try reflexivity) ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance rwlockwriteguard_ne α : NonExpansive (rwlockwriteguard α).
Proof. apply type_contractive_ne, _. Qed.
Global Instance rwlockwriteguard_mono E L :
Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockwriteguard.
Proof.
intros α1 α2 ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold.
iIntros (Hty' qmax qL) "HL".
iDestruct (Hty' with "HL") as "#Hty". iDestruct ( with "HL") as "#Hα".
iDestruct (rwlock_inv_proper with "HL") as "#Hty1ty2"; first done.
iDestruct (rwlock_inv_proper with "HL") as "#Hty2ty1"; first by symmetry.
iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2.
iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro].
- done.
- iIntros (tid [|[[]|][]]) "H"; try done.
iDestruct "H" as (γ β tid_shr) "(Hb & #H⊑ & #Hinv & Hown)".
iExists γ, β, tid_shr. iFrame "∗#". iSplit; last iSplit.
+ iApply bor_iff; last done.
iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]";
iExists vl; iFrame; by iApply "Ho".
+ iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done.
+ iApply at_bor_iff; try done.
iIntros "!>!>"; iSplit; iIntros "H".
* by iApply "Hty1ty2".
* by iApply "Hty2ty1".
- iIntros (κ tid l) "H". iDestruct "H" as (l') "H". iExists l'.
iDestruct "H" as "[$ #H]". iIntros "!> * % Htok".
iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
{ iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. }
iMod ("H" with "[] Htok") as "Hshr"; first done. iModIntro. iNext.
iMod "Hshr" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$".
iApply ty_shr_mono; try done.
+ iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl.
+ by iApply "Hs".
Qed.
Global Instance rwlockwriteguard_mono_flip E L :
Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) rwlockwriteguard.
Proof. intros ??????. by apply rwlockwriteguard_mono. Qed.
Lemma rwlockwriteguard_mono' E L α1 α2 ty1 ty2 :
lctx_lft_incl E L α2 α1 eqtype E L ty1 ty2
subtype E L (rwlockwriteguard α1 ty1) (rwlockwriteguard α2 ty2) .
Proof. intros. by eapply rwlockwriteguard_mono. Qed.
Global Instance rwlockwriteguard_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) rwlockwriteguard.
Proof. intros ??[]???. split; by apply rwlockwriteguard_mono'. Qed.
Lemma rwlockwriteguard_proper' E L α1 α2 ty1 ty2 :
lctx_lft_eq E L α1 α2 eqtype E L ty1 ty2
eqtype E L (rwlockwriteguard α1 ty1) (rwlockwriteguard α2 ty2).
Proof. intros. by eapply rwlockwriteguard_proper. Qed.
(* Rust requires the type to also be Send. Not sure why. *)
Global Instance rwlockwriteguard_sync α ty :
Sync ty Sync (rwlockwriteguard α ty).
Proof.
move=>?????/=. apply bi.exist_mono=>?. do 7 f_equiv.
by rewrite sync_change_tid.
Qed.
(* POSIX requires the unlock to occur from the thread that acquired
the lock, so Rust does not implement Send for RwLockWriteGuard. We can
prove this for our spinlock implementation, however. *)
Global Instance rwlockwriteguard_send α ty :
Send ty Send (rwlockwriteguard α ty).
Proof.
iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H".
iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_entails.
repeat lazymatch goal with
| |- (ty_own _ _ _) (ty_own _ _ _) => by apply send_change_tid'
| |- (rwlock_inv _ _ _ _ _ _) _ => by apply rwlock_inv_change_tid_own
| |- _ => f_equiv
end.
Qed.
End rwlockwriteguard.
Global Hint Resolve rwlockwriteguard_mono' rwlockwriteguard_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From iris.bi Require Import fractional.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Import typing.
From lrust.typing.lib.rwlock Require Import rwlock rwlockwriteguard.
From iris.prelude Require Import options.
Section rwlockwriteguard_functions.
Context `{!typeGS Σ, !rwlockG Σ}.
(* Turning a rwlockwriteguard into a shared borrow. *)
Definition rwlockwriteguard_deref : val :=
fn: ["x"] :=
let: "x'" := !"x" in
let: "r'" := !"x'" + #1 in
letalloc: "r" <- "r'" in
delete [ #1; "x"];; return: ["r"].
Lemma rwlockwriteguard_deref_type ty `{!TyWf ty} :
typed_val rwlockwriteguard_deref
(fn( '(α, β), ; &shr{α}(rwlockwriteguard β ty)) &shr{α} ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done.
iDestruct "Hx'" as (l') "#[Hfrac Hshr]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "([Hα1 Hα2] & HL & Hclose)";
[solve_typing..|].
iMod (frac_bor_acc with "LFT Hfrac Hα1") as (qlx') "[H↦ Hcloseα1]"; first done.
rewrite heap_pointsto_vec_singleton.
iMod (lctx_lft_alive_tok_noend β with "HE HL") as () "(Hβ & HL & Hclose')";
[solve_typing..|].
iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hcloseβα2]".
wp_bind (!(LitV lx'))%E. iApply (wp_step_fupd with "[Hα2β]");
[|by iApply ("Hshr" with "[] Hα2β")|]; first done.
iMod "H↦" as "[H↦1 H↦2]". wp_read. iIntros "[#Hshr' Hα2β]!>". wp_op. wp_let.
iDestruct ("Hcloseβα2" with "Hα2β") as "[Hβ Hα2]".
iMod ("Hcloseα1" with "[$H↦1 $H↦2]") as "Hα1". iMod ("Hclose'" with "Hβ HL") as "HL".
iMod ("Hclose" with "[$] HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|].
iApply (type_type _ _ _ [ x box (&shr{α}(rwlockwriteguard β ty));
#(l' + 1) &shr{α}ty]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
iFrame. iApply (ty_shr_mono with "[] Hshr'").
iApply lft_incl_glb; first by iApply lft_incl_syn_sem.
by iApply lft_incl_refl. }
iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|].
iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Turning a rwlockwriteguard into a unique borrow. *)
Definition rwlockwriteguard_derefmut : val :=
fn: ["x"] :=
let: "x'" := !"x" in
let: "r'" := !"x'" + #1 in
letalloc: "r" <- "r'" in
delete [ #1; "x"];; return: ["r"].
Lemma rwlockwriteguard_derefmut_type ty `{!TyWf ty} :
typed_val rwlockwriteguard_derefmut
(fn( '(α, β), ; &uniq{α}(rwlockwriteguard β ty)) &uniq{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x').
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst.
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done.
iMod (bor_exists with "LFT Hx'") as (vl) "H"; first done.
iMod (bor_sep with "LFT H") as "[H↦ H]"; first done.
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)"; [solve_typing..|].
destruct vl as [|[[|l|]|][]];
try by iMod (bor_persistent with "LFT H Hα") as "[>[] _]".
rewrite heap_pointsto_vec_singleton.
iMod (bor_exists with "LFT H") as (γ) "H"; first done.
iMod (bor_exists with "LFT H") as (δ) "H"; first done.
iMod (bor_exists with "LFT H") as (tid_shr) "H"; first done.
iMod (bor_sep with "LFT H") as "[Hb H]"; first done.
iMod (bor_sep with "LFT H") as "[Hβδ _]"; first done.
iMod (bor_persistent with "LFT Hβδ Hα") as "[#Hβδ Hα]"; first done.
iMod (bor_acc with "LFT H↦ Hα") as "[H↦ Hcloseα]"; first done.
wp_bind (!(LitV lx'))%E. iMod (bor_unnest with "LFT Hb") as "Hb"; first done.
wp_read. wp_op. wp_let. iMod "Hb".
iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]". iMod ("Hclose" with "Hα HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|].
iApply (type_type _ _ _ [ x box (uninit 1); #(l + 1) &uniq{α}ty]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
iFrame. iApply (bor_shorten with "[] Hb"). iApply lft_incl_glb.
- iApply lft_incl_trans; last done. by iApply lft_incl_syn_sem.
- by iApply lft_incl_refl. }
iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|].
iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Drop a rwlockwriteguard and release the lock. *)
Definition rwlockwriteguard_drop : val :=
fn: ["x"] :=
let: "x'" := !"x" in
"x'" <-ˢᶜ #0;;
delete [ #1; "x"];;
let: "r" := new [ #0] in return: ["r"].
Lemma rwlockwriteguard_drop_type ty `{!TyWf ty} :
typed_val rwlockwriteguard_drop
(fn( α, ; rwlockwriteguard α ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk HT".
rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iDestruct "HT" as "[Hx Hx']".
destruct x' as [[|lx'|]|]; try done. simpl.
iDestruct "Hx'" as (γ β tid_own) "(Hx' & #Hαβ & #Hinv & H◯)".
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose)";
[solve_typing..|].
iMod (lft_incl_acc with "Hαβ Hα") as () "[Hβ Hcloseα]"; first done.
wp_bind (#lx' <-ˢᶜ #0)%E.
iMod (at_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|].
iDestruct "INV" as (st) "(H↦ & H● & INV)". wp_write.
iMod ("Hcloseβ" with "[> H↦ H● H◯ INV Hx']") as "Hβ".
{ iCombine "H● H◯" gives %[[[=]| (? & st0 & [=<-] & -> &
[Heq|Hle])]%option_included Hv]%auth_both_valid_discrete;
last by destruct (exclusive_included _ _ Hle).
destruct st0 as [[[]|]| |]; try by inversion Heq.
iExists None. iFrame. iMod (own_update_2 with "H● H◯") as "$"; last done.
apply auth_update_dealloc. rewrite -(right_id None op (Some _)).
apply cancel_local_update_unit, _. }
iModIntro. wp_seq. iMod ("Hcloseα" with "Hβ") as "Hα".
iMod ("Hclose" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ x box (uninit 1)]
with "[] LFT HE Hna HL Hk"); first last.
{ rewrite tctx_interp_singleton tctx_hasty_val //. }
iApply type_delete; [solve_typing..|].
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_jump; solve_typing.
Qed.
End rwlockwriteguard_functions.
From iris.proofmode Require Import proofmode.
From lrust.lang Require Import spawn.
From lrust.typing Require Export type.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Definition spawnN := lrustN .@ "spawn".
Section join_handle.
Context `{!typeGS Σ, !spawnG Σ}.
Definition join_inv (ty : type) (v : val) :=
( tid, (box ty).(ty_own) tid [v])%I.
Program Definition join_handle (ty : type) :=
{| ty_size := 1;
ty_own _ vl :=
match vl return _ with
| [ #(LitLoc l) ] => lang.lib.spawn.join_handle spawnN l (join_inv ty)
| _ => False
end%I;
ty_shr κ _ l := True%I |}.
Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". Qed.
Next Obligation. iIntros "* _ _ _ $". auto. Qed.
Next Obligation. iIntros (?) "**"; auto. Qed.
Global Instance join_handle_wf ty `{!TyWf ty} : TyWf (join_handle ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Lemma join_handle_subtype ty1 ty2 :
type_incl ty1 ty2 -∗ type_incl (join_handle ty1) (join_handle ty2).
Proof.
iIntros "#Hincl". iSplit; first done. iSplit; iModIntro.
- iIntros "%tid %vl Hvl". destruct vl as [|[[|vl|]|] [|]]; try done.
simpl. iApply (join_handle_impl with "[] Hvl"). clear tid.
iIntros "!> * Hown" (tid).
iDestruct (box_type_incl with "Hincl") as "{Hincl} (_ & Hincl & _)".
iApply "Hincl". done.
- iIntros "* _". auto.
Qed.
Global Instance join_handle_mono E L :
Proper (subtype E L ==> subtype E L) join_handle.
Proof.
iIntros (ty1 ty2 Hsub ??) "HL". iDestruct (Hsub with "HL") as "#Hsub".
iIntros "!> #HE". iApply join_handle_subtype. iApply "Hsub"; done.
Qed.
Global Instance join_handle_proper E L :
Proper (eqtype E L ==> eqtype E L) join_handle.
Proof. intros ??[]. by split; apply join_handle_mono. Qed.
Global Instance join_handle_type_contractive : TypeContractive join_handle.
Proof.
constructor;
solve_proper_core ltac:(fun _ => progress unfold join_inv || exact: type_dist2_dist || f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance join_handle_ne : NonExpansive join_handle.
Proof. apply type_contractive_ne, _. Qed.
Global Instance join_handle_send ty :
Send (join_handle ty).
Proof. iIntros (???) "$ //". Qed.
Global Instance join_handle_sync ty : Sync (join_handle ty).
Proof. iIntros (????) "_ //". Qed.
End join_handle.
Section spawn.
Context `{!typeGS Σ, !spawnG Σ}.
Definition spawn (call_once : val) : val :=
fn: ["f"] :=
let: "call_once" := call_once in
let: "join" := spawn [λ: ["c"],
letcall: "r" := "call_once" ["f":expr] in
finish ["c"; "r"]] in
letalloc: "r" <- "join" in
return: ["r"].
Lemma spawn_type fty retty call_once `{!TyWf fty, !TyWf retty}
`(!Send fty, !Send retty) :
typed_val call_once (fn(; fty) retty) (* fty : FnOnce() -> retty, as witnessed by the impl call_once *)
let E ϝ := ty_outlives_E fty static ++ ty_outlives_E retty static in
typed_val (spawn call_once) (fn(E; fty) join_handle retty).
Proof.
intros Hf ? E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>env. simpl_subst.
iApply type_let; [apply Hf|solve_typing|]. iIntros (f'). simpl_subst.
iApply (type_let _ _ _ _ ([f' _; env _])
(λ j, [j join_handle retty])); try solve_typing; [|].
{ (* The core of the proof: showing that spawn is safe. *)
iIntros (tid qmax) "#LFT #HE $ $ [Hf' [Henv _]]". rewrite !tctx_hasty_val [fn _]lock.
iApply (spawn_spec _ (join_inv retty) with "[-]"); last first.
{ iIntros "!> *". rewrite tctx_interp_singleton tctx_hasty_val.
iIntros "?". by iFrame. }
simpl_subst. iIntros (c) "Hfin". iMod na_alloc as (tid') "Htl". wp_let. wp_let.
unlock. iApply (type_call_iris _ [] () [_] with "LFT HE Htl [] Hf' [Henv]");
(* The `solve_typing` here shows that, because we assume that `fty` and `retty`
outlive `static`, the implicit requirmeents made by `call_once` are satisifed. *)
[solve_typing|iApply (lft_tok_static 1%Qp)| |].
- by rewrite big_sepL_singleton tctx_hasty_val send_change_tid.
- iIntros (r) "Htl _ Hret".
wp_rec. iApply (finish_spec with "[$Hfin Hret]"); last auto.
iIntros (?). by iApply @send_change_tid. }
iIntros (v). simpl_subst.
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition join : val :=
fn: ["c"] :=
let: "c'" := !"c" in
let: "r" := spawn.join ["c'"] in
delete [ #1; "c"];; return: ["r"].
Lemma join_type retty `{!TyWf retty} :
typed_val join (fn(; join_handle retty) retty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>c. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (c'); simpl_subst.
iApply (type_let _ _ _ _ ([c' _])
(λ r, [r box retty])); try solve_typing; [|].
{ iIntros (tid qmax) "#LFT _ $ $".
rewrite tctx_interp_singleton tctx_hasty_val. iIntros "Hc".
destruct c' as [[|c'|]|]; try done.
iApply (join_spec with "Hc"). iNext. iIntros "* Hret".
rewrite tctx_interp_singleton tctx_hasty_val. done. }
iIntros (r); simpl_subst. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End spawn.
From iris.proofmode Require Import proofmode.
From lrust.lang.lib Require Import swap.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Section swap.
Context `{!typeGS Σ}.
Definition swap ty : val :=
fn: ["p1"; "p2"] :=
let: "p1'" := !"p1" in
let: "p2'" := !"p2" in
swap ["p1'"; "p2'"; #ty.(ty_size)];;
delete [ #1; "p1"];; delete [ #1; "p2"];;
let: "r" := new [ #0] in return: ["r"].
Lemma swap_type τ `{!TyWf τ} :
typed_val (swap τ) (fn( α, ; &uniq{α} τ, &uniq{α} τ) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α ϝ ret p).
inv_vec p=>p1 p2. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (p1'). simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (p2'). simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk (H2 & H2' & H1 & H1' & _)".
rewrite !tctx_hasty_val.
iMod (lctx_lft_alive_tok α with "HE HL") as () "([Hα1 Hα2] & HL & Hclose)";
[solve_typing..|].
iDestruct (uniq_is_ptr with "H1'") as (l1) "#H1eq". iDestruct "H1eq" as %[=->].
iMod (bor_acc with "LFT H1' Hα1") as "[H1' Hclose1]"=>//.
iDestruct "H1'" as (vl1) "[>H↦1 H1']".
iDestruct (τ.(ty_size_eq) with "H1'") as "#>%".
iDestruct (uniq_is_ptr with "H2'") as (l2) "#H2eq". iDestruct "H2eq" as %[=->].
iMod (bor_acc with "LFT H2' Hα2") as "[H2' Hclose2]"=>//.
iDestruct "H2'" as (vl2) "[>H↦2 H2']".
iDestruct (τ.(ty_size_eq) with "H2'") as "#>%".
wp_apply (wp_swap with "[$H↦1 $H↦2]"); try lia. iIntros "[H↦1 H↦2]". wp_seq.
iMod ("Hclose" with "[>-Hna HL H1 H2 Hk] HL") as "HL".
{ iMod ("Hclose1" with "[H2' H↦1]") as "[_ $]"; first by iExists _; iFrame.
by iMod ("Hclose2" with "[H1' H↦2]") as "[_ $]"; first by iExists _; iFrame. }
(* Finish up the proof. *)
iApply (type_type _ _ _ [ p1 box (uninit 1); p2 box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply (type_new _); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_jump; solve_typing.
Qed.
End swap.
From iris.proofmode Require Import proofmode.
From lrust.lang.lib Require Import memcpy.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Export type.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Section code.
Context `{!typeGS Σ}.
Definition take ty (call_once : val) : val :=
fn: ["x"; "f"] :=
let: "x'" := !"x" in
let: "call_once" := call_once in
letalloc: "t" <-{ty.(ty_size)} !"x'" in
letcall: "r" := "call_once" ["f"; "t"]%E in
"x'" <-{ty.(ty_size)} !"r";;
delete [ #1; "x"];; delete [ #ty.(ty_size); "r"];;
let: "r" := new [ #0] in return: ["r"].
Lemma take_type ty fty call_once `{!TyWf ty, !TyWf fty} :
(* fty : FnOnce(ty) -> ty, as witnessed by the impl call_once *)
typed_val call_once (fn(; fty, ty) ty)
typed_val (take ty call_once) (fn( α, ; &uniq{α} ty, fty) unit).
Proof.
intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x env. simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (x'); simpl_subst.
iApply type_let; [apply Hf|solve_typing|]; iIntros (f'); simpl_subst.
iApply (type_new ty.(ty_size)); [solve_typing..|]; iIntros (t); simpl_subst.
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk (Ht & Hf' & Hx & Hx' & Henv & _)".
rewrite !tctx_hasty_val [[x]]lock [[env]]lock [fn _]lock.
iDestruct (ownptr_uninit_own with "Ht") as (tl tvl) "(% & >Htl & Htl†)". subst t. simpl.
destruct x' as [[|lx'|]|]; try done. simpl.
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose1)"; [solve_typing..|].
iMod (lctx_lft_alive_tok_noend ϝ with "HE HL") as () "(Hϝ & HL & Hclose2)"; [solve_typing..|].
iMod (bor_acc with "LFT Hx' Hα") as "[Hx' Hclose3]"; first done.
iDestruct (heap_pointsto_ty_own with "Hx'") as (x'vl) "[>Hx'↦ Hx'vl]".
wp_apply (wp_memcpy with "[$Htl $Hx'↦]"); [by auto using length_vec_to_list..|].
iIntros "[Htl Hx'↦]". wp_seq.
(* Call the function. *)
wp_let. unlock.
iApply (type_call_iris _ [ϝ] () [_; _]
with "LFT HE Hna [Hϝ] Hf' [Henv Htl Htl† Hx'vl]"); [solve_typing| | |].
{ by rewrite /= (right_id static). }
{ rewrite /= !tctx_hasty_val tctx_hasty_val' //. iFrame. }
(* Prove the continuation of the function call. *)
iIntros (r) "Hna Hϝ Hr". simpl.
iDestruct (ownptr_own with "Hr") as (lr rvl) "(% & Hlr & Hrvl & Hlr†)". subst r.
wp_rec.
rewrite (right_id static).
wp_apply (wp_memcpy with "[$Hx'↦ $Hlr]"); [by auto using length_vec_to_list..|].
iIntros "[Hlx' Hlr]". wp_seq.
iMod ("Hclose3" with "[Hlx' Hrvl]") as "[Hlx' Hα]".
{ iExists _. iFrame. }
iMod ("Hclose2" with "Hϝ HL") as "HL".
iMod ("Hclose1" with "Hα HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x _; #lr box (uninit ty.(ty_size)) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. auto using length_vec_to_list. }
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_new; [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_jump; solve_typing.
Qed.
End code.
From Coq Require Import Qcanon. From iris.proofmode Require Import proofmode.
From iris.proofmode Require Import tactics. From lrust.lang.lib Require Import memcpy.
From iris.base_logic Require Import big_op.
From lrust.lang Require Export new_delete.
From lrust.lang Require Import memcpy.
From lrust.typing Require Export type. From lrust.typing Require Export type.
From lrust.typing Require Import uninit type_context programs. From lrust.typing Require Import util uninit type_context programs.
Set Default Proof Using "Type". From iris.prelude Require Import options.
Section own. Section own.
Context `{typeG Σ}. Context `{!typeGS Σ}.
Program Definition freeable_sz (n : nat) (sz : nat) (l : loc) : iProp Σ := Definition freeable_sz (n : nat) (sz : nat) (l : loc) : iProp Σ :=
match sz, n return _ with match sz, n return _ with
| 0%nat, _ => True | 0%nat, _ => True
| _, 0%nat => False | _, 0%nat => False
| sz, n => {mk_Qp (sz / n) _}lsz | sz, n => {pos_to_Qp (Pos.of_nat sz) / pos_to_Qp (Pos.of_nat n)}lsz
end%I. end%I.
Next Obligation. intros _ _ _ sz0 ? n ?. by apply Qcmult_pos_pos. Qed. Global Arguments freeable_sz : simpl never.
Arguments freeable_sz : simpl never.
Global Instance freable_sz_timeless n sz l : TimelessP (freeable_sz n sz l). Global Instance freeable_sz_timeless n sz l : Timeless (freeable_sz n sz l).
Proof. destruct sz, n; apply _. Qed. Proof. destruct sz, n; apply _. Qed.
Lemma freeable_sz_full n l : freeable_sz n n l ⊣⊢ ({1}ln Z.of_nat n = 0)%I. Lemma freeable_sz_full n l : freeable_sz n n l ⊣⊢ {1}ln Z.of_nat n = 0⌝.
Proof. Proof.
destruct n. destruct n as [|n].
- iSplit; iIntros "H /="; auto. - iSplit; iIntros "H /="; auto.
- assert (Z.of_nat (S n) = 0 False) as -> by done. rewrite right_id. - assert (Z.of_nat (S n) = 0 False) as -> by done. rewrite right_id.
rewrite /freeable_sz (proj2 (Qp_eq (mk_Qp _ _) 1)) //= Qcmult_inv_r //. by rewrite /freeable_sz Qp.div_diag.
Qed. Qed.
Lemma freeable_sz_full_S n l : freeable_sz (S n) (S n) l ⊣⊢ {1}l(S n).
Proof. rewrite freeable_sz_full. iSplit; auto. iIntros "[$|%]"; done. Qed.
Lemma freeable_sz_split n sz1 sz2 l : Lemma freeable_sz_split n sz1 sz2 l :
freeable_sz n sz1 l freeable_sz n sz2 (shift_loc l sz1) ⊣⊢ freeable_sz n sz1 l freeable_sz n sz2 (l + sz1) ⊣⊢
freeable_sz n (sz1 + sz2) l. freeable_sz n (sz1 + sz2) l.
Proof. Proof.
destruct sz1; [|destruct sz2;[|rewrite /freeable_sz plus_Sn_m; destruct n]]. destruct sz1; [|destruct sz2;[|rewrite /freeable_sz plus_Sn_m; destruct n]].
- by rewrite left_id shift_loc_0. - by rewrite left_id shift_loc_0.
- by rewrite right_id Nat.add_0_r. - by rewrite right_id Nat.add_0_r.
- iSplit. by iIntros "[[]?]". by iIntros "[]". - iSplit.
- rewrite heap_freeable_op_eq. f_equiv. apply Qp_eq. + by iIntros "[[]?]".
rewrite /= -Qcmult_plus_distr_l -Z2Qc_inj_add /Z.add. do 3 f_equal. lia. + by iIntros "[]".
- rewrite heap_freeable_op_eq. f_equiv; [|done..].
by rewrite -Qp.div_add_distr pos_to_Qp_add -Nat2Pos.inj_add.
Qed. Qed.
(* Make sure 'simpl' doesn't unfold. *)
Global Opaque freeable_sz.
Program Definition own_ptr (n : nat) (ty : type) := Program Definition own_ptr (n : nat) (ty : type) :=
{| ty_size := 1; {| ty_size := 1;
ty_own tid vl := ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] =>
(* We put a later in front of the †{q}, because we cannot use (* We put a later in front of the †{q}, because we cannot use
[ty_size_eq] on [ty] at step index 0, which would in turn [ty_size_eq] on [ty] at step index 0, which would in turn
prevent us to prove [subtype_own]. prevent us to prove [subtype_own].
Since this assertion is timeless, this should not cause Since this assertion is timeless, this should not cause
problems. *) problems. *)
( l:loc, vl = [ #l ] l ↦∗: ty.(ty_own) tid (l ↦∗: ty.(ty_own) tid freeable_sz n ty.(ty_size) l)
freeable_sz n ty.(ty_size) l)%I; | _ => False
end%I;
ty_shr κ tid l := ty_shr κ tid l :=
( l':loc, &frac{κ}(λ q', l {q'} #l') ( l':loc, &frac{κ}(λ q', l {q'} #l')
( F q, ⌜↑shrN lftE F -∗ q.[κ] ={F,F∖↑shrN∖↑lftN}▷=∗ ty.(ty_shr) κ tid l' q.[κ]))%I ( F q, ⌜↑shrN lftN F -∗ q.[κ] ={F}[F∖↑shrN]▷=∗
|}. ty.(ty_shr) κ tid l' q.[κ]))%I |}.
Next Obligation. Next Obligation. by iIntros (q ty tid [|[[]|][]]) "H". Qed.
iIntros (q ty tid vl) "H". iDestruct "H" as (l) "[% _]". by subst.
Qed.
Next Obligation. Next Obligation.
move=>n ty N κ l tid ?? /=. iIntros "#LFT Hshr Htok". move=>n ty N κ l tid ?? /=. iIntros "#LFT Hshr Htok".
iMod (bor_exists with "LFT Hshr") as (vl) "Hb". set_solver. iMod (bor_exists with "LFT Hshr") as (vl) "Hb"; first solve_ndisj.
iMod (bor_sep with "LFT Hb") as "[Hb1 Hb2]". set_solver. iMod (bor_sep with "LFT Hb") as "[Hb1 Hb2]"; first solve_ndisj.
iMod (bor_exists with "LFT Hb2") as (l') "Hb2". set_solver. destruct vl as [|[[|l'|]|][]];
iMod (bor_sep with "LFT Hb2") as "[EQ Hb2]". set_solver. try (iMod (bor_persistent with "LFT Hb2 Htok") as "[>[]_]"; solve_ndisj).
iMod (bor_persistent_tok with "LFT EQ Htok") as "[>% $]". set_solver. iFrame. iExists l'. rewrite heap_pointsto_vec_singleton.
iExists l'. subst. rewrite heap_mapsto_vec_singleton. rewrite bi.later_sep.
iMod (bor_sep with "LFT Hb2") as "[Hb2 _]". set_solver. iMod (bor_sep with "LFT Hb2") as "[Hb2 _]"; first solve_ndisj.
iMod (bor_fracture (λ q, l {q} #l')%I with "LFT Hb1") as "$". set_solver. iMod (bor_fracture (λ q, l {q} #l')%I with "LFT Hb1") as "$"; first solve_ndisj.
rewrite bor_unfold_idx. iDestruct "Hb2" as (i) "(#Hpb&Hpbown)". iApply delay_sharing_later; done.
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_open with "Hinv") as "[INV Hclose]". set_solver.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iMod (bor_later with "LFT [Hbtok]") as "Hb".
{ apply ndisj_subseteq_difference. solve_ndisj. set_solver. } (* FIXME: some tactic should solve this in one go. *)
{ rewrite bor_unfold_idx. eauto. }
iModIntro. iNext. iMod "Hb".
iMod (ty.(ty_share) with "LFT Hb Htok") as "[#$ $]".
{ apply ndisj_subseteq_difference. solve_ndisj. set_solver. } (* FIXME: some tactic should solve this in one go. *)
iApply "Hclose". auto.
- iMod fupd_intro_mask' as "Hclose'"; last iModIntro. set_solver.
iNext. iMod "Hclose'" as "_". iMod ("Hclose" with "[]") as "_"; by eauto.
Qed. Qed.
Next Obligation. Next Obligation.
intros _ ty κ κ' tid l. iIntros "#LFT #Hκ #H". intros _ ty κ κ' tid l. iIntros "#Hκ #H".
iDestruct "H" as (l') "[Hfb #Hvs]". iDestruct "H" as (l') "[Hfb #Hvs]".
iExists l'. iSplit. by iApply (frac_bor_shorten with "[]"). iIntros "!# *% Htok". iExists l'. iSplit; first by iApply (frac_bor_shorten with "[]"). iIntros "!> %F %q % Htok".
iApply (step_fupd_mask_mono F _ _ (F∖↑shrN∖↑lftN)). set_solver. set_solver. iApply (step_fupd_mask_mono F _ (F∖↑shrN)); [solve_ndisj..|].
iMod (lft_incl_acc with "Hκ Htok") as (q') "[Htok Hclose]"; first set_solver. iMod (lft_incl_acc with "Hκ Htok") as (q') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hvs" with "* [%] Htok") as "Hvs'". set_solver. iModIntro. iNext. iMod ("Hvs" with "[%] Htok") as "Hvs'"; first solve_ndisj. iModIntro. iNext.
iMod "Hvs'" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$". iMod "Hvs'" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$".
by iApply (ty.(ty_shr_mono) with "LFT Hκ"). by iApply (ty.(ty_shr_mono) with "Hκ").
Qed. Qed.
Global Instance own_mono E L n : Global Instance own_ptr_wf n ty `{!TyWf ty} : TyWf (own_ptr n ty) :=
Proper (subtype E L ==> subtype E L) (own_ptr n). { ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Lemma own_type_incl n m ty1 ty2 :
n = m -∗ type_incl ty1 ty2 -∗ type_incl (own_ptr n ty1) (own_ptr m ty2).
Proof. Proof.
intros ty1 ty2 Hincl. iIntros. iSplit; first done. iIntros "#Heq (#Hsz & #Ho & #Hs)". iSplit; first done. iSplit; iModIntro.
iDestruct (Hincl with "* [] [] []") as "(_ & #Ho & #Hs)"; [done..|clear Hincl]. - iIntros (?[|[[| |]|][]]) "H"; try done. simpl.
iSplit; iAlways. iDestruct "H" as "[Hmt H†]". iNext. iDestruct ("Hsz") as %<-.
- iIntros (??) "H". iDestruct "H" as (l) "(%&Hmt&H†)". subst. iDestruct "Heq" as %->. iFrame. iApply (heap_pointsto_pred_wand with "Hmt").
iExists _. iSplit. done. iDestruct "Hmt" as (vl') "[Hmt Hown]". iNext. iApply "Ho".
iDestruct (ty_size_eq with "Hown") as %<-.
iDestruct ("Ho" with "* Hown") as "Hown".
iDestruct (ty_size_eq with "Hown") as %<-. iFrame.
iExists _. by iFrame.
- iIntros (???) "H". iDestruct "H" as (l') "[Hfb #Hvs]". - iIntros (???) "H". iDestruct "H" as (l') "[Hfb #Hvs]".
iExists l'. iFrame. iIntros "!#". iIntros (F' q) "% Htok". iExists l'. iFrame. iIntros "!>". iIntros (F' q) "% Htok".
iMod ("Hvs" with "* [%] Htok") as "Hvs'". done. iModIntro. iNext. iMod ("Hvs" with "[%] Htok") as "Hvs'"; first done. iModIntro. iNext.
iMod "Hvs'" as "[Hshr $]". iApply ("Hs" with "Hshr"). iMod "Hvs'" as "[Hshr $]". iApply ("Hs" with "Hshr").
Qed. Qed.
Lemma own_mono' E L n ty1 ty2 :
subtype E L ty1 ty2 subtype E L (own_ptr n ty1) (own_ptr n ty2).
Proof. apply own_mono. Qed.
Global Instance own_proper E L n :
Proper (eqtype E L ==> eqtype E L) (own_ptr n).
Proof. intros ?? Heq. split; f_equiv; apply Heq. Qed.
Lemma own_proper' E L n ty1 ty2 :
eqtype E L ty1 ty2 eqtype E L (own_ptr n ty1) (own_ptr n ty2).
Proof. apply own_proper. Qed.
Global Instance own_contractive n : Contractive (own_ptr n). Global Instance own_mono E L n :
Proper (subtype E L ==> subtype E L) (own_ptr n).
Proof. Proof.
intros m ?? EQ. split; [split|]; simpl. intros ty1 ty2 Hincl. iIntros (qmax qL) "HL".
- done. iDestruct (Hincl with "HL") as "#Hincl".
- destruct m=>// tid vl /=. repeat (apply EQ || f_contractive || f_equiv). iClear "∗". iIntros "!> #HE".
- intros κ tid l. repeat (apply EQ || f_contractive || f_equiv). iApply own_type_incl; first by auto. iApply "Hincl"; auto.
Qed. Qed.
Global Instance own_ne n m : Proper (dist m ==> dist m) (own_ptr n). Lemma own_mono' E L n1 n2 ty1 ty2 :
Proof. apply contractive_ne, _. Qed. n1 = n2 subtype E L ty1 ty2 subtype E L (own_ptr n1 ty1) (own_ptr n2 ty2).
Proof. intros -> *. by apply own_mono. Qed.
Global Instance own_proper E L n :
Proper (eqtype E L ==> eqtype E L) (own_ptr n).
Proof. intros ??[]; split; by apply own_mono. Qed.
Lemma own_proper' E L n1 n2 ty1 ty2 :
n1 = n2 eqtype E L ty1 ty2 eqtype E L (own_ptr n1 ty1) (own_ptr n2 ty2).
Proof. intros -> *. by apply own_proper. Qed.
Global Instance own_type_contractive n : TypeContractive (own_ptr n).
Proof. solve_type_proper. Qed.
Global Instance own_ne n : NonExpansive (own_ptr n).
Proof. apply type_contractive_ne, _. Qed.
Global Instance own_send n ty : Global Instance own_send n ty :
Send ty Send (own_ptr n ty). Send ty Send (own_ptr n ty).
Proof. Proof.
iIntros (Hsend tid1 tid2 vl) "H". iDestruct "H" as (l) "[% [Hm H†]]". subst vl. iIntros (Hsend tid1 tid2 [|[[| |]|][]]) "H"; try done.
iExists _. iSplit; first done. iFrame "H†". iNext. iDestruct "H" as "[Hm $]". iNext. iApply (heap_pointsto_pred_wand with "Hm").
iApply (heap_mapsto_pred_wand with "Hm"). iIntros (vl) "?". by iApply Hsend. iIntros (vl) "?". by iApply Hsend.
Qed. Qed.
Global Instance own_sync n ty : Global Instance own_sync n ty :
Sync ty Sync (own_ptr n ty). Sync ty Sync (own_ptr n ty).
Proof. Proof.
iIntros (Hsync κ tid1 tid2 l) "H". iDestruct "H" as (l') "[Hm #Hshr]". iIntros (Hsync κ tid1 tid2 l) "H". iDestruct "H" as (l') "[Hm #Hshr]".
iExists _. iFrame "Hm". iAlways. iIntros (F q) "% Htok". iExists _. iFrame "Hm". iModIntro. iIntros (F q) "% Htok".
iMod ("Hshr" with "* [] Htok") as "Hfin"; first done. iModIntro. iNext. iMod ("Hshr" with "[] Htok") as "Hfin"; first done. iModIntro. iNext.
iClear "Hshr". (* FIXME: Using "{HShr} [HShr $]" for the intro pattern in the following line doesn't work. *) iMod "Hfin" as "{Hshr} [Hshr $]". by iApply Hsync.
iMod "Hfin" as "[Hshr $]". by iApply Hsync.
Qed. Qed.
End own. End own.
Notation box ty := (own_ptr ty.(ty_size) ty). Section box.
Context `{!typeGS Σ}.
Definition box ty := own_ptr ty.(ty_size) ty.
Lemma box_type_incl ty1 ty2 :
type_incl ty1 ty2 -∗ type_incl (box ty1) (box ty2).
Proof.
iIntros "#Hincl". iApply own_type_incl; last done.
iDestruct "Hincl" as "(? & _ & _)". done.
Qed.
Global Instance box_mono E L :
Proper (subtype E L ==> subtype E L) box.
Proof.
intros ty1 ty2 Hincl. iIntros (qmax qL) "HL".
iDestruct (Hincl with "HL") as "#Hincl".
iClear "∗". iIntros "!> #HE".
iApply box_type_incl. iApply "Hincl"; auto.
Qed.
Lemma box_mono' E L ty1 ty2 :
subtype E L ty1 ty2 subtype E L (box ty1) (box ty2).
Proof. intros. by apply box_mono. Qed.
Global Instance box_proper E L :
Proper (eqtype E L ==> eqtype E L) box.
Proof. intros ??[]; split; by apply box_mono. Qed.
Lemma box_proper' E L ty1 ty2 :
eqtype E L ty1 ty2 eqtype E L (box ty1) (box ty2).
Proof. intros. by apply box_proper. Qed.
Global Instance box_type_contractive : TypeContractive box.
Proof. solve_type_proper. Qed.
Global Instance box_ne : NonExpansive box.
Proof. apply type_contractive_ne, _. Qed.
End box.
Section util.
Context `{!typeGS Σ}.
Lemma ownptr_own n ty tid v :
(own_ptr n ty).(ty_own) tid [v] ⊣⊢
(l : loc) (vl : vec val ty.(ty_size)),
v = #l l ↦∗ vl ty.(ty_own) tid vl freeable_sz n ty.(ty_size) l.
Proof.
iSplit.
- iIntros "Hown". destruct v as [[|l|]|]; try done.
iExists l. iDestruct "Hown" as "[Hown $]". rewrite heap_pointsto_ty_own.
iDestruct "Hown" as (vl) "[??]". eauto with iFrame.
- iIntros "Hown". iDestruct "Hown" as (l vl) "(% & ? & ? & ?)". subst v. iFrame.
Qed.
Lemma ownptr_uninit_own n m tid v :
(own_ptr n (uninit m)).(ty_own) tid [v] ⊣⊢
(l : loc) (vl' : vec val m), v = #l l ↦∗ vl' freeable_sz n m l.
Proof.
rewrite ownptr_own. apply bi.exist_proper=>l. iSplit.
(* FIXME: The goals here look rather confusing: One cannot tell that we are looking at
a statement in Iris; the top-level → could just as well be a Coq implication. *)
- iIntros "H". iDestruct "H" as (vl) "(% & Hl & _ & $)". subst v.
iExists vl. iSplit; done.
- iIntros "H". iDestruct "H" as (vl) "(% & Hl & $)". subst v.
iExists vl. rewrite /= length_vec_to_list.
eauto with iFrame.
Qed.
End util.
Section typing. Section typing.
Context `{typeG Σ}. Context `{!typeGS Σ}.
(** Typing *) (** Typing *)
Lemma write_own {E L} ty ty' n : Lemma write_own {E L} ty ty' n :
ty.(ty_size) = ty'.(ty_size) typed_write E L (own_ptr n ty') ty (own_ptr n ty). ty.(ty_size) = ty'.(ty_size) typed_write E L (own_ptr n ty') ty (own_ptr n ty).
Proof. Proof.
iIntros (Hsz) "!#". iIntros (p tid F qE qL ?) "_ $ $ Hown". iDestruct "Hown" as (l) "(Heq & H↦ & H†)". rewrite typed_write_eq. iIntros (Hsz) "!>".
iDestruct "Heq" as %[= ->]. iDestruct "H↦" as (vl) "[>H↦ Hown]". iIntros ([[]|] tid F qmax qL ?) "_ _ $ Hown"; try done.
iDestruct (ty_size_eq with "Hown") as "#>%". iModIntro. rewrite /= Hsz. iDestruct "Hown" as "[H↦ $]". iDestruct "H↦" as (vl) "[>H↦ Hown]".
iExists _, _. iFrame "H↦". iDestruct (ty_size_eq with "Hown") as "#>%". iFrame. auto.
iSplit; first by rewrite Hsz. iIntros "Hown' !>".
iExists _. iSplit; first done. rewrite Hsz. iFrame.
Qed. Qed.
Lemma read_own_copy E L ty n : Lemma read_own_copy E L ty n :
Copy ty typed_read E L (own_ptr n ty) ty (own_ptr n ty). Copy ty typed_read E L (own_ptr n ty) ty (own_ptr n ty).
Proof. Proof.
iIntros (Hsz) "!#". iIntros (p tid F qE qL ?) "_ $ $ $ Hown". iDestruct "Hown" as (l) "(Heq & H↦ & H†)". rewrite typed_read_eq. iIntros (Hsz) "!>".
iDestruct "Heq" as %[= ->]. iDestruct "H↦" as (vl) "[>H↦ #Hown]". iModIntro. iIntros ([[|l|]|] tid F qmax qL ?) "_ _ $ $ Hown"; try done.
iExists l, _, _. iSplit; first done. iFrame "∗#". iIntros "Hl !>". iDestruct "Hown" as "[H↦ $]". iDestruct "H↦" as (vl) "[>H↦ #Hown]".
iExists _. iSplit; first done. iFrame "H†". iExists _. by iFrame. iExists l, _, _. iSplitR; first done.
iFrame "∗#". by iIntros "!> $".
Qed. Qed.
Lemma read_own_move E L ty n : Lemma read_own_move E L ty n :
typed_read E L (own_ptr n ty) ty (own_ptr n $ uninit ty.(ty_size)). typed_read E L (own_ptr n ty) ty (own_ptr n $ uninit ty.(ty_size)).
Proof. Proof.
iAlways. iIntros (p tid F qE qL ?) "_ $ $ $ Hown". iDestruct "Hown" as (l) "(Heq & H↦ & H†)". rewrite typed_read_eq. iModIntro.
iDestruct "Heq" as %[= ->]. iDestruct "H↦" as (vl) "[>H↦ Hown]". iIntros ([[|l|]|] tid F qmax qL ?) "_ _ $ $ Hown"; try done.
iDestruct "Hown" as "[H↦ $]". iDestruct "H↦" as (vl) "[>H↦ Hown]".
iDestruct (ty_size_eq with "Hown") as "#>%". iDestruct (ty_size_eq with "Hown") as "#>%".
iModIntro. iExists l, vl, _. iSplit; first done. iFrame "∗#". iIntros "Hl !>". iExists l, vl, _. iSplitR; first done.
iExists _. iSplit; first done. iFrame "H†". iExists _. iFrame. iFrame; auto.
iApply uninit_own. auto. by iIntros "!> $".
Qed. Qed.
Lemma type_new_instr {E L} (n : Z) : Lemma type_new_instr {E L} (n : Z) :
0 n 0 n
let n' := Z.to_nat n in let n' := Z.to_nat n in
typed_instruction_ty E L [] (new [ #n ]%E) (own_ptr n' (uninit n')). typed_instruction_ty E L [] (new [ #n ]%E) (own_ptr n' (uninit n')).
Proof. Proof.
intros. iAlways. iIntros (tid q) "#HEAP #LFT $ $ $ _". iIntros (? tid qmax) "#LFT #HE $ $ _".
iApply (wp_new with "HEAP"); try done. iModIntro. iApply wp_new; try done. iModIntro.
iIntros (l vl) "(% & H† & Hlft)". subst. rewrite tctx_interp_singleton tctx_hasty_val. iIntros (l) "(H† & Hlft)". rewrite tctx_interp_singleton tctx_hasty_val.
iExists _. iSplit; first done. iNext. rewrite freeable_sz_full Z2Nat.id //. iFrame. iNext. rewrite freeable_sz_full Z2Nat.id //. iFrame. by rewrite /= repeat_length.
iExists vl. iFrame. by rewrite Nat2Z.id uninit_own.
Qed. Qed.
Lemma type_new E L C T x (n : Z) e : Lemma type_new {E L C T} (n' : nat) x (n : Z) e :
Closed (x :b: []) e Closed (x :b: []) e
0 n 0 n
( (v : val) (n' := Z.to_nat n), n' = Z.to_nat n
typed_body E L C ((v box (uninit n')) :: T) (subst' x v e)) ( (v : val),
typed_body E L C ((v own_ptr n' (uninit n')) :: T) (subst' x v e)) -∗
typed_body E L C T (let: x := new [ #n ] in e). typed_body E L C T (let: x := new [ #n ] in e).
Proof. intros. eapply type_let. done. by apply type_new_instr. solve_typing. done. Qed. Proof. iIntros. subst. iApply type_let; [by apply type_new_instr|solve_typing..]. Qed.
Lemma type_new_subtype ty E L C T x (n : Z) e : Lemma type_new_subtype ty E L C T x (n : Z) e :
Closed (x :b: []) e Closed (x :b: []) e
0 n 0 n
let n' := Z.to_nat n in let n' := Z.to_nat n in
subtype E L (uninit n') ty subtype E L (uninit n') ty
( (v : val), typed_body E L C ((v own_ptr n' ty) :: T) (subst' x v e)) ( (v : val), typed_body E L C ((v own_ptr n' ty) :: T) (subst' x v e)) -∗
typed_body E L C T (let: x := new [ #n ] in e). typed_body E L C T (let: x := new [ #n ] in e).
Proof. Proof.
intros ???? Htyp. eapply type_let. done. by apply type_new_instr. solve_typing. iIntros (????) "Htyp". iApply type_let; [by apply type_new_instr|solve_typing|].
iIntros (v). iApply typed_body_mono; [done| |done|]. iIntros (v). iApply typed_body_mono; last iApply "Htyp"; try done.
(* FIXME : why can't we iApply Htyp? *)
2:by iDestruct (Htyp v) as "$".
by apply (tctx_incl_frame_r _ [_] [_]), subtype_tctx_incl, own_mono. by apply (tctx_incl_frame_r _ [_] [_]), subtype_tctx_incl, own_mono.
Qed. Qed.
Lemma type_delete_instr {E L} ty (n : Z) p : Lemma type_delete_instr {E L} ty (n : Z) p :
Z.of_nat (ty.(ty_size)) = n Z.of_nat (ty.(ty_size)) = n
typed_instruction E L [p box ty] (delete [ #n; p])%E (λ _, []). typed_instruction E L [p own_ptr ty.(ty_size) ty] (delete [ #n; p])%E (λ _, []).
Proof. Proof.
iIntros (<-) "!#". iIntros (tid eq) "#HEAP #LFT $ $ $ Hp". rewrite tctx_interp_singleton. iIntros (<- tid qmax) "#LFT #HE $ $ Hp". rewrite tctx_interp_singleton.
wp_bind p. iApply (wp_hasty with "Hp"). iIntros (v) "_ Hown". wp_bind p. iApply (wp_hasty with "Hp"). iIntros ([[]|]) "_ Hown"; try done.
iDestruct "Hown" as (l) "(Hv & H↦: & >H†)". iDestruct "Hv" as %[=->]. iDestruct "Hown" as "[H↦: >H†]". iDestruct "H↦:" as (vl) "[>H↦ Hown]".
iDestruct "H↦∗:" as (vl) "[>H↦ Hown]". rewrite tctx_interp_nil. iDestruct (ty_size_eq with "Hown") as "#>EQ".
iDestruct (ty_size_eq with "Hown") as "#>EQ". iDestruct "EQ" as %<-. iDestruct "EQ" as %<-. iApply (wp_delete with "[-]"); auto.
iApply (wp_delete with "[-]"); try (by auto); []. - iFrame "H↦". by iApply freeable_sz_full.
rewrite freeable_sz_full. by iFrame. - rewrite /tctx_interp /=; auto.
Qed. Qed.
Lemma type_delete {E L} ty C T T' (n' : nat) (n : Z) p e : Lemma type_delete {E L} ty C T T' (n' : nat) (n : Z) p e :
Closed [] e Closed [] e
tctx_extract_hasty E L p (own_ptr n' ty) T T' tctx_extract_hasty E L p (own_ptr n' ty) T T'
n = n' Z.of_nat (ty.(ty_size)) = n n = n' Z.of_nat (ty.(ty_size)) = n
typed_body E L C T' e typed_body E L C T' e -∗
typed_body E L C T (delete [ #n; p ] ;; e). typed_body E L C T (delete [ #n; p ] ;; e).
Proof. Proof.
intros ?? -> Hlen ?. eapply type_seq; [done|by apply type_delete_instr| |done]. iIntros (?? -> Hlen) "?". iApply type_seq; [by apply type_delete_instr| |done].
by rewrite (inj _ _ _ Hlen). by rewrite (inj _ _ _ Hlen).
Qed. Qed.
...@@ -251,44 +310,46 @@ Section typing. ...@@ -251,44 +310,46 @@ Section typing.
Closed [] p Closed (x :b: []) e Closed [] p Closed (x :b: []) e
tctx_extract_hasty E L p ty T T' tctx_extract_hasty E L p ty T T'
ty.(ty_size) = 1%nat ty.(ty_size) = 1%nat
( (v : val), typed_body E L C ((v own_ptr 1 ty)::T') (subst x v e)) ( (v : val), typed_body E L C ((v own_ptr 1 ty)::T') (subst x v e)) -∗
typed_body E L C T (letalloc: x <- p in e). typed_body E L C T (letalloc: x <- p in e).
Proof. Proof.
intros. eapply type_new. iIntros (??? Hsz) "**". iApply type_new.
- rewrite /Closed /=. rewrite !andb_True. - rewrite /Closed /=. rewrite !andb_True.
eauto 10 using is_closed_weaken with set_solver. eauto 10 using is_closed_weaken with set_solver.
- done. - done.
- move=>xv /=. - solve_typing.
- iIntros (xv) "/=". rewrite -Hsz.
assert (subst x xv (x <- p ;; e)%E = (xv <- p ;; subst x xv e)%E) as ->. assert (subst x xv (x <- p ;; e)%E = (xv <- p ;; subst x xv e)%E) as ->.
{ (* TODO : simpl_subst should be able to do this. *) { (* TODO : simpl_subst should be able to do this. *)
unfold subst=>/=. repeat f_equal. unfold subst=>/=. repeat f_equal.
- by rewrite bool_decide_true. - by rewrite bool_decide_true.
- eapply is_closed_subst. done. set_solver. } - eapply is_closed_subst; first done. set_solver. }
eapply type_assign; [|solve_typing|by eapply write_own|done]. iApply type_assign; [|solve_typing|by eapply write_own|solve_typing].
apply subst_is_closed; last done. apply is_closed_of_val. apply subst_is_closed; last done. apply is_closed_of_val.
Qed. Qed.
Lemma type_letalloc_n {E L} ty ty1 ty2 C T T' (x : string) p e : Lemma type_letalloc_n {E L} ty ty1 ty2 C T T' (x : string) p e :
Closed [] p Closed (x :b: []) e Closed [] p Closed (x :b: []) e
typed_read E L ty1 ty ty2
tctx_extract_hasty E L p ty1 T T' tctx_extract_hasty E L p ty1 T T'
( typed_read E L ty1 ty ty2)
( (v : val), ( (v : val),
typed_body E L C ((v own_ptr (ty.(ty_size)) ty)::(p ty2)::T') (subst x v e)) typed_body E L C ((v own_ptr (ty.(ty_size)) ty)::(p ty2)::T') (subst x v e)) -∗
typed_body E L C T (letalloc: x < !{ty.(ty_size)}p in e). typed_body E L C T (letalloc: x <-{ty.(ty_size)} !p in e).
Proof. Proof.
intros. eapply type_new. iIntros. iApply type_new.
- rewrite /Closed /=. rewrite !andb_True. - rewrite /Closed /=. rewrite !andb_True.
eauto 100 using is_closed_weaken with set_solver. eauto 10 using is_closed_of_val, is_closed_weaken with set_solver.
- lia. - lia.
- move=>xv /=. - done.
assert (subst x xv (x <⋯ !{ty.(ty_size)}p ;; e)%E = - iIntros (xv) "/=".
(xv <⋯ !{ty.(ty_size)}p ;; subst x xv e)%E) as ->. assert (subst x xv (x <-{ty.(ty_size)} !p ;; e)%E =
(xv <-{ty.(ty_size)} !p ;; subst x xv e)%E) as ->.
{ (* TODO : simpl_subst should be able to do this. *) { (* TODO : simpl_subst should be able to do this. *)
unfold subst=>/=. repeat f_equal. unfold subst=>/=. repeat f_equal.
- eapply (is_closed_subst []). done. set_solver. - eapply (is_closed_subst []); last set_solver. apply is_closed_of_val.
- by rewrite bool_decide_true. - by rewrite bool_decide_true.
- eapply is_closed_subst. done. set_solver. } - eapply is_closed_subst; first done. set_solver. }
rewrite Nat2Z.id. eapply type_memcpy. rewrite Nat2Z.id. iApply type_memcpy.
+ apply subst_is_closed; last done. apply is_closed_of_val. + apply subst_is_closed; last done. apply is_closed_of_val.
+ solve_typing. + solve_typing.
+ (* TODO: Doing "eassumption" here shows that unification takes *forever* to fail. + (* TODO: Doing "eassumption" here shows that unification takes *forever* to fail.
...@@ -302,8 +363,8 @@ Section typing. ...@@ -302,8 +363,8 @@ Section typing.
Qed. Qed.
End typing. End typing.
Hint Resolve own_mono' own_proper' : lrust_typing. Global Hint Resolve own_mono' own_proper' box_mono' box_proper'
write_own read_own_copy : lrust_typing.
Hint Extern 100 (_ _) => simpl; lia : lrust_typing. (* By setting the priority high, we make sure copying is tried before
Hint Extern 100 (@eq Z _ _) => simpl; lia : lrust_typing. moving. *)
Hint Extern 100 (@eq nat _ _) => simpl; lia : lrust_typing. Global Hint Resolve read_own_move | 100 : lrust_typing.
\ No newline at end of file
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.algebra Require Import list. From iris.algebra Require Import list numbers.
From lrust.typing Require Import lft_contexts.
From lrust.typing Require Export type. From lrust.typing Require Export type.
Set Default Proof Using "Type". From iris.prelude Require Import options.
Section product. Section product.
Context `{typeG Σ}. Context `{!typeGS Σ}.
(* "Pre"-unit. We later define the full unit as the empty product. That's
convertible, but products are opaque in some hint DBs, so this does make a
difference. *)
Program Definition unit0 : type := Program Definition unit0 : type :=
{| ty_size := 0; ty_own tid vl := vl = []⌝%I; ty_shr κ tid l := True%I |}. {| ty_size := 0; ty_own tid vl := vl = []⌝%I; ty_shr κ tid l := True%I |}.
Next Obligation. iIntros (tid vl) "%". by subst. Qed. Next Obligation. iIntros (tid vl) "%". by subst. Qed.
Next Obligation. by iIntros (??????) "_ _ $". Qed. Next Obligation. by iIntros (??????) "_ _ $". Qed.
Next Obligation. by iIntros (????) "_ _ $". Qed. Next Obligation. by iIntros (????) "_ $". Qed.
Global Instance unit0_copy : Copy unit0. Global Instance unit0_copy : Copy unit0.
Proof. Proof.
split. by apply _. iIntros (????????) "_ _ Htok $". split; first by apply _. iIntros (????????) "_ _ Htok $".
iDestruct (na_own_acc with "Htok") as "[$ Htok]"; first set_solver+. iDestruct (na_own_acc with "Htok") as "[$ Htok]"; first solve_ndisj.
iExists 1%Qp. iModIntro. iSplitR. iExists 1%Qp. iModIntro. iSplitR.
{ iExists []. iSplitL; last by auto. rewrite heap_mapsto_vec_nil. auto. } { iExists []. iSplitL; last by auto. rewrite heap_pointsto_vec_nil. auto. }
iIntros "Htok2 _". iApply "Htok". done. iIntros "Htok2 _". iApply "Htok". done.
Qed. Qed.
...@@ -30,17 +34,17 @@ Section product. ...@@ -30,17 +34,17 @@ Section product.
Lemma split_prod_mt tid ty1 ty2 q l : Lemma split_prod_mt tid ty1 ty2 q l :
(l ↦∗{q}: λ vl, (l ↦∗{q}: λ vl,
vl1 vl2, vl = vl1 ++ vl2 ty1.(ty_own) tid vl1 ty2.(ty_own) tid vl2)%I vl1 vl2, vl = vl1 ++ vl2 ty1.(ty_own) tid vl1 ty2.(ty_own) tid vl2)%I
⊣⊢ l ↦∗{q}: ty1.(ty_own) tid shift_loc l ty1.(ty_size) ↦∗{q}: ty2.(ty_own) tid. ⊣⊢ l ↦∗{q}: ty1.(ty_own) tid (l + ty1.(ty_size)) ↦∗{q}: ty2.(ty_own) tid.
Proof. Proof.
iSplit; iIntros "H". iSplit; iIntros "H".
- iDestruct "H" as (vl) "[H↦ H]". iDestruct "H" as (vl1 vl2) "(% & H1 & H2)". - iDestruct "H" as (vl) "[H↦ H]". iDestruct "H" as (vl1 vl2) "(% & H1 & H2)".
subst. rewrite heap_mapsto_vec_app. iDestruct "H↦" as "[H↦1 H↦2]". subst. rewrite heap_pointsto_vec_app. iDestruct "H↦" as "[H↦1 H↦2]".
iDestruct (ty_size_eq with "H1") as %->. iDestruct (ty_size_eq with "H1") as %->.
iSplitL "H1 H↦1"; iExists _; iFrame. iSplitL "H1 H↦1"; iExists _; iFrame.
- iDestruct "H" as "[H1 H2]". iDestruct "H1" as (vl1) "[H↦1 H1]". - iDestruct "H" as "[H1 H2]". iDestruct "H1" as (vl1) "[H↦1 H1]".
iDestruct "H2" as (vl2) "[H↦2 H2]". iExists (vl1 ++ vl2). iDestruct "H2" as (vl2) "[H↦2 H2]". iExists (vl1 ++ vl2).
rewrite heap_mapsto_vec_app. iDestruct (ty_size_eq with "H1") as %->. rewrite heap_pointsto_vec_app. iDestruct (ty_size_eq with "H1") as %->.
iFrame. iExists _, _. by iFrame. by iFrame.
Qed. Qed.
Program Definition product2 (ty1 ty2 : type) := Program Definition product2 (ty1 ty2 : type) :=
...@@ -49,42 +53,46 @@ Section product. ...@@ -49,42 +53,46 @@ Section product.
vl = vl1 ++ vl2 ty1.(ty_own) tid vl1 ty2.(ty_own) tid vl2)%I; vl = vl1 ++ vl2 ty1.(ty_own) tid vl1 ty2.(ty_own) tid vl2)%I;
ty_shr κ tid l := ty_shr κ tid l :=
(ty1.(ty_shr) κ tid l (ty1.(ty_shr) κ tid l
ty2.(ty_shr) κ tid (shift_loc l ty1.(ty_size)))%I |}. ty2.(ty_shr) κ tid (l + ty1.(ty_size)))%I |}.
Next Obligation. Next Obligation.
iIntros (ty1 ty2 tid vl) "H". iDestruct "H" as (vl1 vl2) "(% & H1 & H2)". iIntros (ty1 ty2 tid vl) "H". iDestruct "H" as (vl1 vl2) "(% & H1 & H2)".
subst. rewrite app_length !ty_size_eq. subst. rewrite length_app !ty_size_eq.
iDestruct "H1" as %->. iDestruct "H2" as %->. done. iDestruct "H1" as %->. iDestruct "H2" as %->. done.
Qed. Qed.
Next Obligation. Next Obligation.
intros ty1 ty2 E κ l tid q ?. iIntros "#LFT /=H Htok". rewrite split_prod_mt. intros ty1 ty2 E κ l tid q ?. iIntros "#LFT /=H Htok". rewrite split_prod_mt.
iMod (bor_sep with "LFT H") as "[H1 H2]". set_solver. iMod (bor_sep with "LFT H") as "[H1 H2]"; first solve_ndisj.
iMod (ty1.(ty_share) with "LFT H1 Htok") as "[? Htok]". solve_ndisj. iMod (ty1.(ty_share) with "LFT H1 Htok") as "[? Htok]"; first solve_ndisj.
iMod (ty2.(ty_share) with "LFT H2 Htok") as "[? Htok]". solve_ndisj. iMod (ty2.(ty_share) with "LFT H2 Htok") as "[? Htok]"; first solve_ndisj.
iModIntro. iFrame "Htok". iFrame. iModIntro. iFrame "Htok". iFrame.
Qed. Qed.
Next Obligation. Next Obligation.
intros ty1 ty2 κ κ' tid l. iIntros "#LFT /= #H⊑ [H1 H2]". intros ty1 ty2 κ κ' tid l. iIntros "/= #H⊑ [H1 H2]".
iSplitL "H1"; by iApply (ty_shr_mono with "LFT H⊑"). iSplitL "H1"; by iApply (ty_shr_mono with "H⊑").
Qed. Qed.
Global Instance product2_ne n: Global Instance product2_type_ne n:
Proper (dist n ==> dist n ==> dist n) product2. Proper (type_dist2 n ==> type_dist2 n ==> type_dist2 n) product2.
Proof. solve_type_proper. Qed.
Global Instance product2_ne :
NonExpansive2 product2.
Proof. Proof.
intros ?? EQ1 ?? EQ2. split; [split|]; simpl. constructor;
- by rewrite EQ1 EQ2. solve_proper_core ltac:(fun _ => (eapply ty_size_ne; try reflexivity) || f_equiv).
- f_contractive=>tid vl. by setoid_rewrite EQ1; setoid_rewrite EQ2.
- intros ???. by rewrite EQ1 EQ2.
Qed. Qed.
Global Instance product2_mono E L: Global Instance product2_mono E L:
Proper (subtype E L ==> subtype E L ==> subtype E L) product2. Proper (subtype E L ==> subtype E L ==> subtype E L) product2.
Proof. Proof.
iIntros (ty11 ty12 H1 ty21 ty22 H2). iIntros. iIntros (ty11 ty12 H1 ty21 ty22 H2). iIntros (qmax qL) "HL".
iDestruct (H1 with "* [] [] []") as "(% & #Ho1 & #Hs1)"; [done..|]. clear H1. iDestruct (H1 with "HL") as "#H1". iDestruct (H2 with "HL") as "#H2".
iDestruct (H2 with "* [] [] []") as "(% & #Ho2 & #Hs2)"; [done..|]. clear H2. iClear "∗". iIntros "!> #HE".
iSplit; first by (iPureIntro; simpl; f_equal). iSplit; iAlways. iDestruct ("H1" with "HE") as "#(% & #Ho1 & #Hs1)". clear H1.
iDestruct ("H2" with "HE") as "#(% & #Ho2 & #Hs2)". clear H2.
iSplit; first by (iPureIntro; simpl; f_equal). iSplit; iModIntro.
- iIntros (??) "H". iDestruct "H" as (vl1 vl2) "(% & Hown1 & Hown2)". - iIntros (??) "H". iDestruct "H" as (vl1 vl2) "(% & Hown1 & Hown2)".
iExists _, _. iSplit. done. iSplitL "Hown1". iExists _, _. iSplit; first done. iSplitL "Hown1".
+ by iApply "Ho1". + by iApply "Ho1".
+ by iApply "Ho2". + by iApply "Ho2".
- iIntros (???) "#[Hshr1 Hshr2]". iSplit. - iIntros (???) "#[Hshr1 Hshr2]". iSplit.
...@@ -100,32 +108,34 @@ Section product. ...@@ -100,32 +108,34 @@ Section product.
Proof. Proof.
split; first (intros; apply _). split; first (intros; apply _).
intros κ tid E F l q ? HF. iIntros "#LFT [H1 H2] Htl [Htok1 Htok2]". intros κ tid E F l q ? HF. iIntros "#LFT [H1 H2] Htl [Htok1 Htok2]".
iMod (@copy_shr_acc with "LFT H1 Htl Htok1") as (q1) "(Htl & H1 & Hclose1)"; first set_solver. iMod (@copy_shr_acc with "LFT H1 Htl Htok1") as (q1) "(Htl & H1 & Hclose1)"; first solve_ndisj.
{ rewrite <-HF. apply shr_locsE_subseteq. simpl. clear. omega. } { rewrite <-HF. apply shr_locsE_subseteq. simpl. clear. lia. }
iMod (@copy_shr_acc with "LFT H2 Htl Htok2") as (q2) "(Htl & H2 & Hclose2)"; first set_solver. iMod (@copy_shr_acc with "LFT H2 Htl Htok2") as (q2) "(Htl & H2 & Hclose2)"; first solve_ndisj.
{ move: HF. rewrite /= -plus_assoc shr_locsE_shift. { move: HF. rewrite /= -Nat.add_assoc shr_locsE_shift.
assert (shr_locsE l (ty_size ty1) shr_locsE (shift_loc l (ty_size ty1)) (ty_size ty2 + 1)) assert (shr_locsE l (ty_size ty1) ## shr_locsE (l + (ty_size ty1)) (ty_size ty2 + 1))
by exact: shr_locsE_disj. by exact: shr_locsE_disj.
set_solver. } set_solver. }
iDestruct (na_own_acc with "Htl") as "[$ Htlclose]". iDestruct (na_own_acc with "Htl") as "[$ Htlclose]".
{ generalize (shr_locsE_shift l ty1.(ty_size) ty2.(ty_size)). simpl. set_solver+. } { generalize (shr_locsE_shift l ty1.(ty_size) ty2.(ty_size)). simpl. set_solver+. }
destruct (Qp_lower_bound q1 q2) as (qq & q'1 & q'2 & -> & ->). iExists qq. destruct (Qp.lower_bound q1 q2) as (qq & q'1 & q'2 & -> & ->). iExists qq.
iDestruct "H1" as (vl1) "[H↦1 H1]". iDestruct "H2" as (vl2) "[H↦2 H2]". iDestruct "H1" as (vl1) "[H↦1 H1]". iDestruct "H2" as (vl2) "[H↦2 H2]".
rewrite !split_prod_mt. rewrite !split_prod_mt.
iDestruct (ty_size_eq with "H1") as "#>%". iDestruct (ty_size_eq with "H1") as "#>%".
iDestruct (ty_size_eq with "H2") as "#>%". iDestruct (ty_size_eq with "H2") as "#>%".
iDestruct "H↦1" as "[H↦1 H↦1f]". iDestruct "H↦2" as "[H↦2 H↦2f]". iDestruct "H↦1" as "[H↦1 H↦1f]". iDestruct "H↦2" as "[H↦2 H↦2f]".
iIntros "!>". iSplitL "H↦1 H1 H↦2 H2". iIntros "!>". iSplitL "H↦1 H1 H↦2 H2".
- iNext. iSplitL "H↦1 H1". iExists vl1. by iFrame. iExists vl2. by iFrame. - iNext. iSplitL "H↦1 H1".
+ iExists vl1. by iFrame.
+ iExists vl2. by iFrame.
- iIntros "Htl [H1 H2]". iDestruct ("Htlclose" with "Htl") as "Htl". - iIntros "Htl [H1 H2]". iDestruct ("Htlclose" with "Htl") as "Htl".
iDestruct "H1" as (vl1') "[H↦1 H1]". iDestruct "H2" as (vl2') "[H↦2 H2]". iDestruct "H1" as (vl1') "[H↦1 H1]". iDestruct "H2" as (vl2') "[H↦2 H2]".
iDestruct (ty_size_eq with "H1") as "#>%". iDestruct (ty_size_eq with "H1") as "#>%".
iDestruct (ty_size_eq with "H2") as "#>%". iDestruct (ty_size_eq with "H2") as "#>%".
iCombine "H↦1" "H↦1f" as "H↦1". iCombine "H↦2" "H↦2f" as "H↦2". iCombine "H↦1" "H↦1f" as "H↦1". iCombine "H↦2" "H↦2f" as "H↦2".
rewrite !heap_mapsto_vec_op; [|congruence..]. rewrite !heap_pointsto_vec_op; [|congruence..].
iDestruct "H↦1" as "[_ H↦1]". iDestruct "H↦2" as "[_ H↦2]". iDestruct "H↦1" as "[_ H↦1]". iDestruct "H↦2" as "[_ H↦2]".
iMod ("Hclose2" with "Htl [H2 H↦2]") as "[Htl $]". by iExists _; iFrame. iMod ("Hclose2" with "Htl [H2 H↦2]") as "[Htl $]"; first by iExists _; iFrame.
iMod ("Hclose1" with "Htl [H1 H↦1]") as "[$$]". by iExists _; iFrame. done. iMod ("Hclose1" with "Htl [H1 H↦1]") as "[$$]"; last done. by iExists _; iFrame.
Qed. Qed.
Global Instance product2_send `{!Send ty1} `{!Send ty2} : Global Instance product2_send `{!Send ty1} `{!Send ty2} :
...@@ -144,7 +154,16 @@ Section product. ...@@ -144,7 +154,16 @@ Section product.
Definition product := foldr product2 unit0. Definition product := foldr product2 unit0.
Definition unit := product []. Definition unit := product [].
Global Instance product_ne n: Proper (dist n ==> dist n) product. Global Instance product_wf tyl `{!ListTyWf tyl} : TyWf (product tyl) :=
{ ty_lfts := tyl_lfts tyl; ty_wf_E := tyl_wf_E tyl }.
Lemma outlives_product ty1 ty2 ϝ `{!TyWf ty1, !TyWf ty2} :
ty_outlives_E (product [ty1; ty2]) ϝ = ty_outlives_E ty1 ϝ ++ ty_outlives_E ty2 ϝ.
Proof. rewrite /product /ty_outlives_E /= fmap_app //. Qed.
Global Instance product_type_ne n: Proper (Forall2 (type_dist2 n) ==> type_dist2 n) product.
Proof. intros ??. induction 1=>//=. by f_equiv. Qed.
Global Instance product_ne : NonExpansive product.
Proof. intros ??. induction 1=>//=. by f_equiv. Qed. Proof. intros ??. induction 1=>//=. by f_equiv. Qed.
Global Instance product_mono E L: Global Instance product_mono E L:
Proper (Forall2 (subtype E L) ==> subtype E L) product. Proper (Forall2 (subtype E L) ==> subtype E L) product.
...@@ -158,11 +177,11 @@ Section product. ...@@ -158,11 +177,11 @@ Section product.
Lemma product_proper' E L tyl1 tyl2 : Lemma product_proper' E L tyl1 tyl2 :
Forall2 (eqtype E L) tyl1 tyl2 eqtype E L (product tyl1) (product tyl2). Forall2 (eqtype E L) tyl1 tyl2 eqtype E L (product tyl1) (product tyl2).
Proof. apply product_proper. Qed. Proof. apply product_proper. Qed.
Global Instance product_copy tys : LstCopy tys Copy (product tys). Global Instance product_copy tys : ListCopy tys Copy (product tys).
Proof. induction 1; apply _. Qed. Proof. induction 1; apply _. Qed.
Global Instance product_send tys : LstSend tys Send (product tys). Global Instance product_send tys : ListSend tys Send (product tys).
Proof. induction 1; apply _. Qed. Proof. induction 1; apply _. Qed.
Global Instance product_sync tys : LstSync tys Sync (product tys). Global Instance product_sync tys : ListSync tys Sync (product tys).
Proof. induction 1; apply _. Qed. Proof. induction 1; apply _. Qed.
Definition product_cons ty tyl : Definition product_cons ty tyl :
...@@ -174,43 +193,43 @@ End product. ...@@ -174,43 +193,43 @@ End product.
Notation Π := product. Notation Π := product.
Section typing. Section typing.
Context `{typeG Σ}. Context `{!typeGS Σ}.
Global Instance prod2_assoc E L : Assoc (eqtype E L) product2. Global Instance prod2_assoc E L : Assoc (eqtype E L) product2.
Proof. Proof.
split; (iIntros; (iSplit; first by rewrite /= assoc); iSplit; iAlways; intros ???. apply eqtype_unfold. iIntros (??) "_ !> _".
last iIntros (?); iIntros (??) "H"). iSplit; first by rewrite /= assoc. iSplit; iIntros "!> *"; iSplit; iIntros "H".
- iDestruct "H" as (vl1 vl') "(% & Ho1 & H)". - iDestruct "H" as (vl1 vl') "(% & Ho1 & H)".
iDestruct "H" as (vl2 vl3) "(% & Ho2 & Ho3)". subst. iDestruct "H" as (vl2 vl3) "(% & Ho2 & Ho3)". subst.
iExists _, _. iSplit. by rewrite assoc. iFrame. iExists _, _. by iFrame. iExists _, _. iSplit; first by rewrite assoc. by iFrame.
- iDestruct "H" as "(Hs1 & Hs2 & Hs3)". rewrite shift_loc_assoc_nat.
by iFrame.
- iDestruct "H" as (vl1 vl') "(% & H & Ho3)". - iDestruct "H" as (vl1 vl') "(% & H & Ho3)".
iDestruct "H" as (vl2 vl3) "(% & Ho1 & Ho2)". subst. iDestruct "H" as (vl2 vl3) "(% & Ho1 & Ho2)". subst.
iExists _, _. iSplit. by rewrite -assoc. iFrame. iExists _, _. by iFrame. iExists _, _. iSplit; first by rewrite -assoc. by iFrame.
- iDestruct "H" as "(Hs1 & Hs2 & Hs3)". rewrite shift_loc_assoc_nat.
by iFrame.
- iDestruct "H" as "[[Hs1 Hs2] Hs3]". rewrite /= shift_loc_assoc_nat. - iDestruct "H" as "[[Hs1 Hs2] Hs3]". rewrite /= shift_loc_assoc_nat.
by iFrame. by iFrame.
Qed. Qed.
Global Instance prod2_unit_leftid E L : LeftId (eqtype E L) unit product2. Global Instance prod2_unit_leftid E L : LeftId (eqtype E L) unit product2.
Proof. Proof.
intros ty. split; (iIntros; (iSplit; first done); iSplit; iAlways; intros ty. apply eqtype_unfold. iIntros (??) "_ !> _". iSplit; first done.
last iIntros (?); iIntros (??) "H"). iSplit; iIntros "!> *"; iSplit; iIntros "H".
- iDestruct "H" as (? ?) "(% & % & ?)". by subst. - iDestruct "H" as (? ?) "(% & % & ?)". by subst.
- iExists [], _. eauto.
- iDestruct "H" as "(? & ?)". rewrite shift_loc_0. - iDestruct "H" as "(? & ?)". rewrite shift_loc_0.
iApply (ty_shr_mono with "[] []"); [done| | done]. iApply ty_shr_mono; last done.
iApply lft_incl_refl. iApply lft_incl_refl.
- iExists [], _. eauto.
- simpl. rewrite shift_loc_0. by iFrame. - simpl. rewrite shift_loc_0. by iFrame.
Qed. Qed.
Global Instance prod2_unit_rightid E L : RightId (eqtype E L) unit product2. Global Instance prod2_unit_rightid E L : RightId (eqtype E L) unit product2.
Proof. Proof.
intros ty. split; (iIntros; (iSplit; first by rewrite /= -plus_n_O); iSplit; intros ty. apply eqtype_unfold. iIntros (??) "_ !> _".
iAlways; last iIntros (?); iIntros (??) "H"). iSplit; first by rewrite /= -plus_n_O. iSplit; iIntros "!> *"; iSplit; iIntros "H".
- iDestruct "H" as (? ?) "(% & ? & %)". subst. by rewrite app_nil_r. - iDestruct "H" as (? ?) "(% & ? & %)". subst. by rewrite app_nil_r.
- iDestruct "H" as "(? & _)". done.
- iExists _, []. rewrite app_nil_r. eauto. - iExists _, []. rewrite app_nil_r. eauto.
- iDestruct "H" as "(? & _)". done.
- simpl. iFrame. - simpl. iFrame.
Qed. Qed.
...@@ -218,7 +237,7 @@ Section typing. ...@@ -218,7 +237,7 @@ Section typing.
eqtype E L (Π(tyl1 ++ Π tyl2 :: tyl3)) (Π(tyl1 ++ tyl2 ++ tyl3)). eqtype E L (Π(tyl1 ++ Π tyl2 :: tyl3)) (Π(tyl1 ++ tyl2 ++ tyl3)).
Proof. Proof.
unfold product. induction tyl1; simpl; last by f_equiv. unfold product. induction tyl1; simpl; last by f_equiv.
induction tyl2. by rewrite left_id. by rewrite /= -assoc; f_equiv. induction tyl2; first by rewrite left_id. by rewrite /= -assoc; f_equiv.
Qed. Qed.
Lemma prod_flatten_l E L tyl1 tyl2 : Lemma prod_flatten_l E L tyl1 tyl2 :
...@@ -230,8 +249,17 @@ Section typing. ...@@ -230,8 +249,17 @@ Section typing.
Lemma prod_app E L tyl1 tyl2 : Lemma prod_app E L tyl1 tyl2 :
eqtype E L (Π[Π tyl1; Π tyl2]) (Π(tyl1 ++ tyl2)). eqtype E L (Π[Π tyl1; Π tyl2]) (Π(tyl1 ++ tyl2)).
Proof. by rewrite -prod_flatten_r -prod_flatten_l. Qed. Proof. by rewrite -prod_flatten_r -prod_flatten_l. Qed.
Lemma ty_outlives_E_elctx_sat_product E L tyl {Wf : ListTyWf tyl} α:
elctx_sat E L (tyl_outlives_E tyl α)
elctx_sat E L (ty_outlives_E (Π 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 typing. End typing.
Arguments product : simpl never. Global Arguments product : simpl never.
Hint Opaque product : lrust_typing lrust_typing_merge. Global Hint Opaque product : lrust_typing lrust_typing_merge.
Hint Resolve product_mono' product_proper' : lrust_typing. Global Hint Resolve product_mono' product_proper' ty_outlives_E_elctx_sat_product
: lrust_typing.
From Coq Require Import Qcanon. From iris.algebra Require Import numbers.
From iris.proofmode Require Import tactics. From iris.proofmode Require Import proofmode.
From iris.base_logic Require Import big_op.
From lrust.typing Require Export type. From lrust.typing Require Export type.
From lrust.typing Require Import type_context lft_contexts product own uniq_bor shr_bor. From lrust.typing Require Import type_context lft_contexts product own uniq_bor.
Set Default Proof Using "Type". From lrust.typing Require Import shr_bor.
From iris.prelude Require Import options.
Section product_split. Section product_split.
Context `{typeG Σ}. Context `{!typeGS Σ}.
(** General splitting / merging for pointer types *) (** General splitting / merging for pointer types *)
Fixpoint hasty_ptr_offsets (p : path) (ptr: type type) tyl (off : nat) : tctx := Fixpoint hasty_ptr_offsets (p : path) (ptr: type type) tyl (off : nat) : tctx :=
match tyl with match tyl with
| [] => [] | [] => []
| ty :: tyl => | ty :: tyl =>
(p + #off ptr ty)%TC :: hasty_ptr_offsets p ptr tyl (off + ty.(ty_size)) (p + #off ptr ty) :: hasty_ptr_offsets p ptr tyl (off + ty.(ty_size))
end. end.
Lemma hasty_ptr_offsets_offset (l : loc) p (off1 off2 : nat) ptr tyl tid : Lemma hasty_ptr_offsets_offset (l : loc) p (off1 off2 : nat) ptr tyl tid :
...@@ -22,11 +22,11 @@ Section product_split. ...@@ -22,11 +22,11 @@ Section product_split.
tctx_interp tid $ hasty_ptr_offsets p ptr tyl (off1 + off2)%nat. tctx_interp tid $ hasty_ptr_offsets p ptr tyl (off1 + off2)%nat.
Proof. Proof.
intros Hp. intros Hp.
revert off1 off2; induction tyl; intros off1 off2; simpl; first done. revert off1 off2; induction tyl as [|ty tyl IH]; intros off1 off2; simpl; first done.
rewrite !tctx_interp_cons. f_equiv; last first. rewrite !tctx_interp_cons. f_equiv; last first.
{ rewrite IHtyl plus_assoc. done. } (* FIXME RJ: Using assoc, even with a pattern, is pretty slow here. *) { by rewrite IH assoc_L. }
apply tctx_elt_interp_hasty_path. clear Hp. simpl. apply tctx_elt_interp_hasty_path. clear Hp. simpl.
clear. destruct (eval_path p); last done. destruct v; try done. clear. destruct (eval_path p) as [v|]; last done. destruct v as [l|]; try done.
destruct l; try done. rewrite shift_loc_assoc Nat2Z.inj_add //. destruct l; try done. rewrite shift_loc_assoc Nat2Z.inj_add //.
Qed. Qed.
...@@ -37,17 +37,17 @@ Section product_split. ...@@ -37,17 +37,17 @@ Section product_split.
( tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]) ( tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val])
p, tctx_incl E L [p ptr $ product tyl] (hasty_ptr_offsets p ptr tyl 0). p, tctx_incl E L [p ptr $ product tyl] (hasty_ptr_offsets p ptr tyl 0).
Proof. Proof.
iIntros (Hsplit Hloc p tid q1 q2) "#LFT HE HL H". iIntros (Hsplit Hloc p tid qmax qL) "#LFT #HE HL H".
iInduction tyl as [|ty tyl IH] "IH" forall (p). iInduction tyl as [|ty tyl IH] "IH" forall (p).
{ iFrame "HE HL". rewrite tctx_interp_nil. done. } { rewrite tctx_interp_nil. auto. }
rewrite product_cons. iMod (Hsplit with "LFT HE HL H") as "(HE & HL & H)". rewrite product_cons. iMod (Hsplit with "LFT HE HL H") as "(HL & H)".
cbn -[tctx_elt_interp]. rewrite tctx_interp_cons tctx_interp_singleton tctx_interp_cons. cbn -[tctx_elt_interp].
iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". iDestruct "Hp" as %Hp. iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". iDestruct "Hp" as %Hp.
iDestruct (Hloc with "Hty") as %[l [=->]]. iDestruct (Hloc with "Hty") as %[l [=->]].
iAssert (tctx_elt_interp tid (p + #0 ptr ty)) with "[Hty]" as "$". iAssert (tctx_elt_interp tid (p + #0 ptr ty)) with "[Hty]" as "$".
{ iExists #l. iSplit; last done. simpl; by rewrite Hp shift_loc_0. } { iExists #l. iSplit; last done. simpl; by rewrite Hp shift_loc_0. }
iMod ("IH" with "* HE HL [Htyl]") as "($ & $ & Htyl)". iMod ("IH" with "HL [Htyl]") as "($ & Htyl)".
{ rewrite tctx_interp_singleton //. } { auto. }
iClear "IH". rewrite (hasty_ptr_offsets_offset l) // -plus_n_O //. iClear "IH". rewrite (hasty_ptr_offsets_offset l) // -plus_n_O //.
Qed. Qed.
...@@ -59,28 +59,26 @@ Section product_split. ...@@ -59,28 +59,26 @@ Section product_split.
( tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]) ( tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val])
p, tctx_incl E L (hasty_ptr_offsets p ptr tyl 0) [p ptr $ product tyl]. p, tctx_incl E L (hasty_ptr_offsets p ptr tyl 0) [p ptr $ product tyl].
Proof. Proof.
iIntros (Hptr Htyl Hmerge Hloc p tid q1 q2) "#LFT HE HL H". iIntros (Hptr Htyl Hmerge Hloc p tid qmax qL) "#LFT #HE HL H".
iInduction tyl as [|ty tyl IH] "IH" forall (p Htyl); first done. iInduction tyl as [|ty tyl IH] "IH" forall (p Htyl); first done.
rewrite product_cons. rewrite /= tctx_interp_singleton tctx_interp_cons. rewrite product_cons. rewrite /= tctx_interp_singleton tctx_interp_cons.
iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]".
iDestruct "Hp" as %Hp. iDestruct (Hloc with "Hty") as %[l [=->]]. iDestruct "Hp" as %Hp. iDestruct (Hloc with "Hty") as %[l [=->]].
assert (eval_path p = Some #l) as Hp'. assert (eval_path p = Some #l) as Hp'.
{ move:Hp. simpl. clear. destruct (eval_path p); last done. { move:Hp. simpl. clear. destruct (eval_path p) as [v|]; last done.
destruct v; try done. destruct l0; try done. rewrite shift_loc_0. done. } destruct v as [l'|]; try done. destruct l'; try done. rewrite shift_loc_0. done. }
clear Hp. destruct tyl. clear Hp. destruct tyl.
{ iDestruct (elctx_interp_persist with "HE") as "#HE'". { assert (eqtype E L (ptr ty) (ptr (product2 ty unit))) as [Hincl _].
iDestruct (llctx_interp_persist with "HL") as "#HL'". iFrame "HE HL".
iClear "IH Htyl". iExists #l. rewrite product_nil. iSplitR; first done.
assert (eqtype E L (ptr ty) (ptr (product2 ty unit))) as [Hincl _].
{ rewrite right_id. done. } { rewrite right_id. done. }
iDestruct (Hincl with "LFT HE' HL'") as "#(_ & #Heq & _)". by iApply "Heq". } iDestruct (Hincl with "HL HE") as "#(_ & #Heq & _)".
iMod ("IH" with "* [] HE HL [Htyl]") as "(HE & HL & Htyl)"; first done. iFrame. iClear "IH Htyl". iExists #l. rewrite product_nil. iSplitR; first done.
by iApply "Heq". }
iMod ("IH" with "[] HL [Htyl]") as "(HL & Htyl)"; first done.
{ change (ty_size ty) with (0+ty_size ty)%nat at 1. { change (ty_size ty) with (0+ty_size ty)%nat at 1.
rewrite plus_comm -hasty_ptr_offsets_offset //. } rewrite Nat.add_comm -hasty_ptr_offsets_offset //. }
iClear "IH". iMod (Hmerge with "LFT HE HL [Hty Htyl]") as "($ & $ & ?)"; iClear "IH". iMod (Hmerge with "LFT HE HL [Hty Htyl]") as "($ & ?)";
last by rewrite tctx_interp_singleton. last by rewrite tctx_interp_singleton.
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iFrame. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. by iFrame.
iExists #l. iSplit; done.
Qed. Qed.
(** Owned pointers *) (** Owned pointers *)
...@@ -88,43 +86,34 @@ Section product_split. ...@@ -88,43 +86,34 @@ Section product_split.
tctx_incl E L [p own_ptr n $ product2 ty1 ty2] tctx_incl E L [p own_ptr n $ product2 ty1 ty2]
[p own_ptr n ty1; p + #ty1.(ty_size) own_ptr n ty2]. [p own_ptr n ty1; p + #ty1.(ty_size) own_ptr n ty2].
Proof. Proof.
iIntros (tid q1 q2) "#LFT $ $ H". iIntros (tid qmax qL) "#LFT _ $ H".
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as (v) "[Hp H]". iDestruct "Hp" as %Hp. iDestruct "H" as (l) "(EQ & H & >H†)". iDestruct "H" as ([[]|]) "[#Hp H]"; try done.
iDestruct "EQ" as %[=->]. iDestruct "H" as (vl) "[>H↦ H]". iDestruct "H" as "[H >H†]". iDestruct "H" as (vl) "[>H↦ H]".
iDestruct "H" as (vl1 vl2) "(>% & H1 & H2)". subst. iDestruct "H" as (vl1 vl2) "(>% & H1 & H2)". subst.
rewrite heap_mapsto_vec_app -freeable_sz_split. rewrite heap_pointsto_vec_app -freeable_sz_split.
iDestruct "H†" as "[H†1 H†2]". iDestruct "H↦" as "[H↦1 H↦2]". iDestruct "H†" as "[H†1 H†2]". iDestruct "H↦" as "[H↦1 H↦2]".
iDestruct (ty_size_eq with "H1") as "#>EQ". iDestruct (ty_size_eq with "H1") as "#>EQ".
iDestruct "EQ" as %->. iSplitL "H↦1 H†1 H1". iDestruct "EQ" as %->. iSplitL "H↦1 H†1 H1".
+ iExists _. iSplitR. done. iExists _. iFrame. iSplitL ""; first done. + iExists _. by iFrame "#∗".
iExists _. iFrame. done. + iExists _. iSplitR; first (by simpl; iDestruct "Hp" as %->).
+ iExists _. iSplitR. simpl. by rewrite Hp. iExists _. iFrame. iSplitR; first done. by iFrame.
iExists _. iFrame. done.
Qed. Qed.
Lemma tctx_merge_own_prod2 E L p n ty1 ty2 : Lemma tctx_merge_own_prod2 E L p n ty1 ty2 :
tctx_incl E L [p own_ptr n ty1; p + #ty1.(ty_size) own_ptr n ty2] tctx_incl E L [p own_ptr n ty1; p + #ty1.(ty_size) own_ptr n ty2]
[p own_ptr n $ product2 ty1 ty2]. [p own_ptr n $ product2 ty1 ty2].
Proof. Proof.
iIntros (tid q1 q2) "#LFT $ $ H". iIntros (tid qmax qL) "#LFT _ $ H".
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as "[H1 H2]". iDestruct "H1" as (v1) "(Hp1 & H1)". iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[|l|]|]) "(Hp1 & H1)"; try done.
iDestruct "Hp1" as %Hp1. iDestruct "H1" as (l) "(EQ & H↦1 & H†1)". iDestruct "H1" as "(H↦1 & H†1)".
iDestruct "EQ" as %[=->]. iDestruct "H2" as (v2) "(Hp2 & H2)". iDestruct "H2" as (v2) "(Hp2 & H2)". simpl. iDestruct "Hp1" as %Hρ1.
rewrite /= Hp1. iDestruct "Hp2" as %[=<-]. iDestruct "H2" as (l') "(EQ & H↦2 & H†2)". rewrite Hρ1. iDestruct "Hp2" as %[=<-]. iDestruct "H2" as "[H↦2 H†2]".
iDestruct "EQ" as %[=<-]. iExists #l. iSplitR; first done. iExists l. iExists #l. iSplitR; first done. rewrite /= -freeable_sz_split. iFrame.
iSplitR; first done. rewrite -freeable_sz_split. iFrame.
iDestruct "H↦1" as (vl1) "[H↦1 H1]". iDestruct "H↦2" as (vl2) "[H↦2 H2]". iDestruct "H↦1" as (vl1) "[H↦1 H1]". iDestruct "H↦2" as (vl2) "[H↦2 H2]".
iExists (vl1 ++ vl2). rewrite heap_mapsto_vec_app. iFrame. iExists (vl1 ++ vl2). rewrite heap_pointsto_vec_app.
iDestruct (ty_size_eq with "H1") as "#>EQ". iDestruct (ty_size_eq with "H1") as "#>->". by iFrame.
iDestruct "EQ" as %->. iFrame. iExists vl1, vl2. iFrame. auto.
Qed.
Lemma own_is_ptr n ty tid (vl : list val) :
ty_own (own_ptr n ty) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝.
Proof.
iIntros "H". iDestruct "H" as (l) "[% _]". iExists l. done.
Qed. Qed.
Lemma tctx_split_own_prod E L n tyl p : Lemma tctx_split_own_prod E L n tyl p :
...@@ -132,7 +121,7 @@ Section product_split. ...@@ -132,7 +121,7 @@ Section product_split.
Proof. Proof.
apply tctx_split_ptr_prod. apply tctx_split_ptr_prod.
- intros. apply tctx_split_own_prod2. - intros. apply tctx_split_own_prod2.
- intros. apply own_is_ptr. - iIntros (??[|[[]|][]]) "?"; eauto.
Qed. Qed.
Lemma tctx_merge_own_prod E L n tyl : Lemma tctx_merge_own_prod E L n tyl :
...@@ -143,51 +132,40 @@ Section product_split. ...@@ -143,51 +132,40 @@ Section product_split.
intros. apply tctx_merge_ptr_prod; try done. intros. apply tctx_merge_ptr_prod; try done.
- apply _. - apply _.
- intros. apply tctx_merge_own_prod2. - intros. apply tctx_merge_own_prod2.
- intros. apply own_is_ptr. - iIntros (??[|[[]|][]]) "?"; eauto.
Qed. Qed.
(** Unique borrows *) (** Unique borrows *)
Lemma tctx_split_uniq_prod2 E L p κ ty1 ty2 : Lemma tctx_split_uniq_prod2 E L p κ ty1 ty2 :
tctx_incl E L [p &uniq{κ} product2 ty1 ty2] tctx_incl E L [p &uniq{κ}(product2 ty1 ty2)]
[p &uniq{κ} ty1; p + #ty1.(ty_size) &uniq{κ} ty2]. [p &uniq{κ} ty1; p + #ty1.(ty_size) &uniq{κ} ty2].
Proof. Proof.
iIntros (tid q1 q2) "#LFT $ $ H". iIntros (tid qmax qL) "#LFT _ $ H".
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as (v) "[Hp H]". iDestruct "Hp" as %Hp. iDestruct "H" as ([[]|]) "[Hp H]"; try done. iDestruct "Hp" as %Hp.
iDestruct "H" as (l P) "[[EQ #HPiff] HP]". iDestruct "EQ" as %[=->]. rewrite /= split_prod_mt. iMod (bor_sep with "LFT H") as "[H1 H2]"; first solve_ndisj.
iMod (bor_iff with "LFT [] HP") as "Hown". set_solver. by eauto. rewrite /tctx_elt_interp /=.
rewrite /= split_prod_mt. iMod (bor_sep with "LFT Hown") as "[H1 H2]". iSplitL "H1"; iExists _; (iSplitR; first by rewrite Hp); auto.
set_solver. rewrite /tctx_elt_interp /=.
iSplitL "H1"; iExists _; (iSplitR; first by rewrite Hp);
iExists _, _; erewrite <-uPred.iff_refl; auto.
Qed. Qed.
Lemma tctx_merge_uniq_prod2 E L p κ ty1 ty2 : Lemma tctx_merge_uniq_prod2 E L p κ ty1 ty2 :
tctx_incl E L [p &uniq{κ} ty1; p + #ty1.(ty_size) &uniq{κ} ty2] tctx_incl E L [p &uniq{κ} ty1; p + #ty1.(ty_size) &uniq{κ} ty2]
[p &uniq{κ} product2 ty1 ty2]. [p &uniq{κ}(product2 ty1 ty2)].
Proof. Proof.
iIntros (tid q1 q2) "#LFT $ $ H". iIntros (tid qmax qL) "#LFT _ $ H".
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as "[H1 H2]". iDestruct "H1" as (v1) "(Hp1 & H1)". iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[|l|]|]) "[Hp1 H1]"; try done.
iDestruct "Hp1" as %Hp1. iDestruct "H1" as (l P) "[[EQ #HPiff] HP]". iDestruct "Hp1" as %Hp1. iDestruct "H2" as (v2) "(Hp2 & H2)". rewrite /= Hp1.
iDestruct "EQ" as %[=->]. iDestruct "Hp2" as %[=<-]. iExists #l. iFrame "%".
iMod (bor_iff with "LFT [] HP") as "Hown1". set_solver. by eauto. iMod (bor_combine with "LFT H1 H2") as "H"; first solve_ndisj. by rewrite /= split_prod_mt.
iDestruct "H2" as (v2) "(Hp2 & H2)". rewrite /= Hp1. iDestruct "Hp2" as %[=<-].
iDestruct "H2" as (l' Q) "[[EQ #HQiff] HQ]". iDestruct "EQ" as %[=<-].
iMod (bor_iff with "LFT [] HQ") as "Hown2". set_solver. by eauto.
iExists #l. iSplitR; first done. iExists l, _. iSplitR.
{ iSplitR; first done. erewrite <-uPred.iff_refl; auto. }
rewrite split_prod_mt. iApply (bor_combine with "LFT Hown1 Hown2"). set_solver.
Qed. Qed.
Lemma uniq_is_ptr κ ty tid (vl : list val) : Lemma uniq_is_ptr κ ty tid (vl : list val) :
ty_own (&uniq{κ} ty) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝. ty_own (&uniq{κ}ty) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝.
Proof. Proof. iIntros "H". destruct vl as [|[[]|][]]; eauto. Qed.
iIntros "H". iDestruct "H" as (l P) "[[% _] _]". iExists l. done.
Qed.
Lemma tctx_split_uniq_prod E L κ tyl p : Lemma tctx_split_uniq_prod E L κ tyl p :
tctx_incl E L [p &uniq{κ} product tyl] tctx_incl E L [p &uniq{κ}(product tyl)]
(hasty_ptr_offsets p (uniq_bor κ) tyl 0). (hasty_ptr_offsets p (uniq_bor κ) tyl 0).
Proof. Proof.
apply tctx_split_ptr_prod. apply tctx_split_ptr_prod.
...@@ -198,7 +176,7 @@ Section product_split. ...@@ -198,7 +176,7 @@ Section product_split.
Lemma tctx_merge_uniq_prod E L κ tyl : Lemma tctx_merge_uniq_prod E L κ tyl :
tyl [] tyl []
p, tctx_incl E L (hasty_ptr_offsets p (uniq_bor κ) tyl 0) p, tctx_incl E L (hasty_ptr_offsets p (uniq_bor κ) tyl 0)
[p &uniq{κ} product tyl]. [p &uniq{κ}(product tyl)].
Proof. Proof.
intros. apply tctx_merge_ptr_prod; try done. intros. apply tctx_merge_ptr_prod; try done.
- apply _. - apply _.
...@@ -208,40 +186,33 @@ Section product_split. ...@@ -208,40 +186,33 @@ Section product_split.
(** Shared borrows *) (** Shared borrows *)
Lemma tctx_split_shr_prod2 E L p κ ty1 ty2 : Lemma tctx_split_shr_prod2 E L p κ ty1 ty2 :
tctx_incl E L [p &shr{κ} product2 ty1 ty2] tctx_incl E L [p &shr{κ}(product2 ty1 ty2)]
[p &shr{κ} ty1; p + #ty1.(ty_size) &shr{κ} ty2]. [p &shr{κ} ty1; p + #ty1.(ty_size) &shr{κ} ty2].
Proof. Proof.
iIntros (tid q1 q2) "#LFT $ $ H". iIntros (tid qmax qL) "#LFT _ $ H".
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as (v) "[Hp H]". iDestruct "Hp" as %Hp. iDestruct "H" as ([[]|]) "[Hp H]"; try iDestruct "H" as "[]".
iDestruct "H" as (l) "[EQ [H1 H2]]". iDestruct "EQ" as %[=->]. iDestruct "H" as "[H1 H2]". iDestruct "Hp" as %Hp.
iSplitL "H1"; iExists _; (iSplitR; first by rewrite /= Hp); by iSplitL "H1"; iExists _; (iSplitR; first by rewrite /= Hp).
iExists _; iSplitR; done.
Qed. Qed.
Lemma tctx_merge_shr_prod2 E L p κ ty1 ty2 : Lemma tctx_merge_shr_prod2 E L p κ ty1 ty2 :
tctx_incl E L [p &shr{κ} ty1; p + #ty1.(ty_size) &shr{κ} ty2] tctx_incl E L [p &shr{κ} ty1; p + #ty1.(ty_size) &shr{κ} ty2]
[p &shr{κ} product2 ty1 ty2]. [p &shr{κ}(product2 ty1 ty2)].
Proof. Proof.
iIntros (tid q1 q2) "#LFT $ $ H". iIntros (tid qmax qL) "#LFT _ $ H".
rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
iDestruct "H" as "[H1 H2]". iDestruct "H1" as (v1) "(Hp1 & H1)". iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[|l|]|]) "[Hp1 Hown1]"; try done.
iDestruct "Hp1" as %Hp1. iDestruct "H1" as (l) "[EQ Hown1]". iDestruct "Hp1" as %Hp1. iDestruct "H2" as ([[]|]) "[Hp2 Hown2]"; try done.
iDestruct "EQ" as %[=->]. iDestruct "H2" as (v2) "(Hp2 & H2)". rewrite /= Hp1. iDestruct "Hp2" as %[=<-]. iExists #l. by iFrame.
rewrite /= Hp1. iDestruct "Hp2" as %[=<-].
iDestruct "H2" as (l') "[EQ Hown2]". iDestruct "EQ" as %[=<-].
iExists #l. iSplitR; first done. iExists l. iSplitR; first done.
by iFrame.
Qed. Qed.
Lemma shr_is_ptr κ ty tid (vl : list val) : Lemma shr_is_ptr κ ty tid (vl : list val) :
ty_own (&shr{κ} ty) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝. ty_own (&shr{κ} ty) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝.
Proof. Proof. iIntros "H". destruct vl as [|[[]|][]]; eauto. Qed.
iIntros "H". iDestruct "H" as (l) "[% _]". iExists l. done.
Qed.
Lemma tctx_split_shr_prod E L κ tyl p : Lemma tctx_split_shr_prod E L κ tyl p :
tctx_incl E L [p &shr{κ} product tyl] tctx_incl E L [p &shr{κ}(product tyl)]
(hasty_ptr_offsets p (shr_bor κ) tyl 0). (hasty_ptr_offsets p (shr_bor κ) tyl 0).
Proof. Proof.
apply tctx_split_ptr_prod. apply tctx_split_ptr_prod.
...@@ -252,7 +223,7 @@ Section product_split. ...@@ -252,7 +223,7 @@ Section product_split.
Lemma tctx_merge_shr_prod E L κ tyl : Lemma tctx_merge_shr_prod E L κ tyl :
tyl [] tyl []
p, tctx_incl E L (hasty_ptr_offsets p (shr_bor κ) tyl 0) p, tctx_incl E L (hasty_ptr_offsets p (shr_bor κ) tyl 0)
[p &shr{κ} product tyl]. [p &shr{κ}(product tyl)].
Proof. Proof.
intros. apply tctx_merge_ptr_prod; try done. intros. apply tctx_merge_ptr_prod; try done.
- apply _. - apply _.
...@@ -274,18 +245,18 @@ Section product_split. ...@@ -274,18 +245,18 @@ Section product_split.
Lemma tctx_extract_split_uniq_prod E L p p' κ ty tyl T T' : Lemma tctx_extract_split_uniq_prod E L p p' κ ty tyl T T' :
tctx_extract_hasty E L p' ty (hasty_ptr_offsets p (uniq_bor κ) tyl 0) T' tctx_extract_hasty E L p' ty (hasty_ptr_offsets p (uniq_bor κ) tyl 0) T'
tctx_extract_hasty E L p' ty ((p &uniq{κ} Π tyl) :: T) (T' ++ T). tctx_extract_hasty E L p' ty ((p &uniq{κ}(Π tyl)) :: T) (T' ++ T).
Proof. Proof.
intros. apply (tctx_incl_frame_r T [_] (_::_)). by rewrite tctx_split_uniq_prod. intros. apply (tctx_incl_frame_r T [_] (_::_)). by rewrite tctx_split_uniq_prod.
Qed. Qed.
Lemma tctx_extract_split_shr_prod E L p p' κ ty tyl T T' : Lemma tctx_extract_split_shr_prod E L p p' κ ty tyl T T' :
tctx_extract_hasty E L p' ty (hasty_ptr_offsets p (shr_bor κ) tyl 0) T' tctx_extract_hasty E L p' ty (hasty_ptr_offsets p (shr_bor κ) tyl 0) T'
tctx_extract_hasty E L p' ty ((p &shr{κ} Π tyl) :: T) ((p &shr{κ} Π tyl) :: T). tctx_extract_hasty E L p' ty ((p &shr{κ}(Π tyl)) :: T) ((p &shr{κ}(Π tyl)) :: T).
Proof. Proof.
intros. apply (tctx_incl_frame_r _ [_] [_;_]). intros. apply (tctx_incl_frame_r _ [_] [_;_]).
rewrite {1}copy_tctx_incl. apply (tctx_incl_frame_r _ [_] [_]). rewrite {1}copy_tctx_incl. apply (tctx_incl_frame_r _ [_] [_]).
rewrite tctx_split_shr_prod -(contains_tctx_incl _ _ [p' ty]%TC) //. rewrite tctx_split_shr_prod -(contains_tctx_incl _ _ [p' ty]) //.
apply submseteq_skip, submseteq_nil_l. apply submseteq_skip, submseteq_nil_l.
Qed. Qed.
...@@ -318,23 +289,23 @@ Section product_split. ...@@ -318,23 +289,23 @@ Section product_split.
Lemma tctx_extract_merge_uniq_prod E L p κ tyl T T' : Lemma tctx_extract_merge_uniq_prod E L p κ tyl T T' :
tyl [] tyl []
extract_tyl E L p (uniq_bor κ) tyl 0 T T' extract_tyl E L p (uniq_bor κ) tyl 0 T T'
tctx_extract_hasty E L p (&uniq{κ}Π tyl) T T'. tctx_extract_hasty E L p (&uniq{κ}(Π tyl)) T T'.
Proof. auto using tctx_extract_merge_ptr_prod, tctx_merge_uniq_prod. Qed. Proof. auto using tctx_extract_merge_ptr_prod, tctx_merge_uniq_prod. Qed.
Lemma tctx_extract_merge_shr_prod E L p κ tyl T T' : Lemma tctx_extract_merge_shr_prod E L p κ tyl T T' :
tyl [] tyl []
extract_tyl E L p (shr_bor κ) tyl 0 T T' extract_tyl E L p (shr_bor κ) tyl 0 T T'
tctx_extract_hasty E L p (&shr{κ}Π tyl) T T'. tctx_extract_hasty E L p (&shr{κ}(Π tyl)) T T'.
Proof. auto using tctx_extract_merge_ptr_prod, tctx_merge_shr_prod. Qed. Proof. auto using tctx_extract_merge_ptr_prod, tctx_merge_shr_prod. Qed.
End product_split. End product_split.
(* We do not want unification to try to unify the definition of these (* We do not want unification to try to unify the definition of these
types with anything in order to try splitting or merging. *) types with anything in order to try splitting or merging. *)
Hint Opaque own_ptr uniq_bor shr_bor tctx_extract_hasty : lrust_typing lrust_typing_merge. Global Hint Opaque own_ptr uniq_bor shr_bor tctx_extract_hasty : lrust_typing lrust_typing_merge.
(* We make sure that splitting is tried before borrowing, so that not (* We make sure that splitting is tried before borrowing, so that not
the entire product is borrowed when only a part is needed. *) the entire product is borrowed when only a part is needed. *)
Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extract_split_shr_prod Global Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extract_split_shr_prod
| 5 : lrust_typing. | 5 : lrust_typing.
(* Merging is also tried after everything, except (* Merging is also tried after everything, except
...@@ -347,6 +318,6 @@ Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extra ...@@ -347,6 +318,6 @@ Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extra
solve_typing get slow because of that. See: solve_typing get slow because of that. See:
https://coq.inria.fr/bugs/show_bug.cgi?id=5304 https://coq.inria.fr/bugs/show_bug.cgi?id=5304
*) *)
Hint Resolve tctx_extract_merge_own_prod tctx_extract_merge_uniq_prod tctx_extract_merge_shr_prod Global Hint Resolve tctx_extract_merge_own_prod tctx_extract_merge_uniq_prod tctx_extract_merge_shr_prod
| 40 : lrust_typing_merge. | 40 : lrust_typing_merge.
Hint Unfold extract_tyl : lrust_typing. Global Hint Unfold extract_tyl : lrust_typing.
\ No newline at end of file
From iris.base_logic Require Import big_op.
From lrust.lang Require Import proofmode memcpy. From lrust.lang Require Import proofmode memcpy.
From lrust.typing Require Export type lft_contexts type_context cont_context. From lrust.typing Require Export type lft_contexts type_context cont_context.
Set Default Proof Using "Type". From iris.prelude Require Import options.
Section typing. Section typing.
Context `{typeG Σ}. Context `{!typeGS Σ}.
(** Function Body *) (** Function Body *)
(* This is an iProp because it is also used by the function type. *) (* This is an iProp because it is also used by the function type. *)
Definition typed_body (E : elctx) (L : llctx) (C : cctx) (T : tctx) Definition typed_body (E : elctx) (L : llctx) (C : cctx) (T : tctx)
(e : expr) : iProp Σ := (e : expr) : iProp Σ :=
( tid qE, heap_ctx -∗ lft_ctx -∗ na_own tid -∗ ( tid (qmax : Qp), lft_ctx -∗ elctx_interp E -∗ na_own tid -∗ llctx_interp qmax L -∗
elctx_interp E qE -∗ llctx_interp L 1 -∗ cctx_interp tid qmax C -∗ tctx_interp tid T -∗
(elctx_interp E qE -∗ cctx_interp tid C) -∗ tctx_interp tid T -∗
WP e {{ _, cont_postcondition }})%I. WP e {{ _, cont_postcondition }})%I.
Global Arguments typed_body _%EL _%LL _%CC _%TC _%E. Global Arguments typed_body _ _ _ _ _%_E.
Global Instance typed_body_llctx_permut E : Global Instance typed_body_llctx_permut E :
Proper (() ==> eq ==> eq ==> eq ==> ()) (typed_body E). Proper (() ==> eq ==> eq ==> eq ==> ()) (typed_body E).
...@@ -34,11 +32,21 @@ Section typing. ...@@ -34,11 +32,21 @@ Section typing.
Proper (flip (cctx_incl E) ==> flip (tctx_incl E L) ==> eq ==> ()) Proper (flip (cctx_incl E) ==> flip (tctx_incl E L) ==> eq ==> ())
(typed_body E L). (typed_body E L).
Proof. Proof.
intros C1 C2 HC T1 T2 HT e ? <-. iIntros "#H !#". intros C1 C2 HC T1 T2 HT e ? <-. iIntros "H".
iIntros (tid qE) "#HEAP #LFT Htl HE HL HC HT". iIntros (tid qmax) "#LFT #HE Htl HL HC HT".
iMod (HT with "LFT HE HL HT") as "(HE & HL & HT)". iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iApply ("H" with "HEAP LFT Htl HE HL [HC] HT"). iMod (HT with "LFT HE HL HT") as "(HL & HT)".
iIntros "HE". by iApply (HC with "LFT HC"). iDestruct ("HLclose" with "HL") as "HL".
iApply ("H" with "LFT HE Htl HL [HC] HT").
by iApply (HC with "LFT HE HC").
Qed.
Lemma typed_body_tctx_incl E L T2 T1 C e :
tctx_incl E L T1 T2
( typed_body E L C T2 e)
typed_body E L C T1 e.
Proof.
intros Hincl He2. iApply typed_body_mono; last done; done.
Qed. Qed.
Global Instance typed_body_mono_flip E L: Global Instance typed_body_mono_flip E L:
...@@ -49,171 +57,234 @@ Section typing. ...@@ -49,171 +57,234 @@ Section typing.
(** Instruction *) (** Instruction *)
Definition typed_instruction (E : elctx) (L : llctx) Definition typed_instruction (E : elctx) (L : llctx)
(T1 : tctx) (e : expr) (T2 : val tctx) : iProp Σ := (T1 : tctx) (e : expr) (T2 : val tctx) : iProp Σ :=
( tid qE, heap_ctx -∗ lft_ctx -∗ na_own tid -∗ ( tid qmax, lft_ctx -∗ elctx_interp E -∗ na_own tid -∗
elctx_interp E qE -∗ llctx_interp L 1 -∗ tctx_interp tid T1 -∗ llctx_interp qmax L -∗ tctx_interp tid T1 -∗
WP e {{ v, na_own tid elctx_interp E qE WP e {{ v, na_own tid
llctx_interp L 1 tctx_interp tid (T2 v) }})%I. llctx_interp qmax L tctx_interp tid (T2 v) }})%I.
Global Arguments typed_instruction _%EL _%LL _%TC _%E _%TC. Global Arguments typed_instruction _ _ _ _%_E _.
(** Writing and Reading **) (** Writing and Reading **)
Definition typed_write (E : elctx) (L : llctx) (ty1 ty ty2 : type) : iProp Σ := Definition typed_write_def (E : elctx) (L : llctx) (ty1 ty ty2 : type) : iProp Σ :=
( v tid F qE qL, lftE (lrustN) F ( v tid F qmax qL, lftN (lrustN) F
lft_ctx -∗ elctx_interp E qE -∗ llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗ lft_ctx -∗ elctx_interp E -∗ llctx_interp_noend qmax L qL -∗ ty1.(ty_own) tid [v] ={F}=∗
(l : loc) vl, length vl = ty.(ty_size) v = #l l ↦∗ vl (l : loc) vl, length vl = ty.(ty_size) v = #l l ↦∗ vl
( l ↦∗: ty.(ty_own) tid ={F}=∗ ( l ↦∗: ty.(ty_own) tid ={F}=∗
elctx_interp E qE llctx_interp L qL ty2.(ty_own) tid [v]))%I. llctx_interp_noend qmax L qL ty2.(ty_own) tid [v]))%I.
Global Arguments typed_write _%EL _%LL _%T _%T _%T. Definition typed_write_aux : seal (@typed_write_def). Proof. by eexists. Qed.
Definition typed_write := typed_write_aux.(unseal).
Definition typed_write_eq : @typed_write = @typed_write_def := typed_write_aux.(seal_eq).
Global Arguments typed_write _ _ _%_T _%_T _%_T.
Global Instance typed_write_persistent (E : elctx) (L : llctx) (ty1 ty ty2 : type) :
Persistent (typed_write E L ty1 ty ty2).
Proof. rewrite typed_write_eq. apply _. Qed.
(* Technically speaking, we could remvoe the vl quantifiaction here and use (* Technically speaking, we could remvoe the vl quantifiaction here and use
mapsto_pred instead (i.e., l ↦∗: ty.(ty_own) tid). However, that would pointsto_pred instead (i.e., l ↦∗: ty.(ty_own) tid). However, that would
make work for some of the provers way harder, since they'd have to show make work for some of the provers way harder, since they'd have to show
that nobody could possibly have changed the vl (because only half the that nobody could possibly have changed the vl (because only half the
fraction was given). So we go with the definition that is easier to prove. *) fraction was given). So we go with the definition that is easier to prove. *)
Definition typed_read (E : elctx) (L : llctx) (ty1 ty ty2 : type) : iProp Σ := Definition typed_read_def (E : elctx) (L : llctx) (ty1 ty ty2 : type) : iProp Σ :=
( v tid F qE qL, lftE (lrustN) F ( v tid F qmax qL, lftN lrustN F
lft_ctx -∗ na_own tid F -∗ lft_ctx -∗ elctx_interp E -∗ na_own tid F -∗
elctx_interp E qE -∗ llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗ llctx_interp_noend qmax L qL -∗ ty1.(ty_own) tid [v] ={F}=∗
(l : loc) vl q, v = #l l ↦∗{q} vl ty.(ty_own) tid vl (l : loc) vl q, v = #l l ↦∗{q} vl ty.(ty_own) tid vl
(l ↦∗{q} vl ={F}=∗ na_own tid F elctx_interp E qE (l ↦∗{q} vl ={F}=∗ na_own tid F
llctx_interp L qL ty2.(ty_own) tid [v]))%I. llctx_interp_noend qmax L qL ty2.(ty_own) tid [v]))%I.
Global Arguments typed_read _%EL _%LL _%T _%T _%T. Definition typed_read_aux : seal (@typed_read_def). Proof. by eexists. Qed.
Definition typed_read := typed_read_aux.(unseal).
Definition typed_read_eq : @typed_read = @typed_read_def := typed_read_aux.(seal_eq).
Global Arguments typed_read _ _ _%_T _%_T _%_T.
Global Instance typed_read_persistent (E : elctx) (L : llctx) (ty1 ty ty2 : type) :
Persistent (typed_read E L ty1 ty ty2).
Proof. rewrite typed_read_eq. apply _. Qed.
End typing. End typing.
Notation typed_instruction_ty E L T1 e ty := Definition typed_instruction_ty `{!typeGS Σ} (E : elctx) (L : llctx) (T : tctx)
(typed_instruction E L T1 e (λ v : val, [v ty]%TC)). (e : expr) (ty : type) : iProp Σ :=
typed_instruction E L T e (λ v, [v ty]).
Global Arguments typed_instruction_ty {_ _} _ _ _ _%_E _%_T.
Definition typed_val `{!typeGS Σ} (v : val) (ty : type) : Prop :=
E L, typed_instruction_ty E L [] (of_val v) ty.
Global Arguments typed_val _ _ _%_V _%_T.
Section typing_rules. Section typing_rules.
Context `{typeG Σ}. Context `{!typeGS Σ}.
(* TODO: notation scopes need tuing to avoid the %list here. *) (* This lemma is helpful when switching from proving unsafe code in Iris
(* TODO: Proof a version of this that substitutes into a compatible context... back to proving it in the type system. *)
if we really want to do that. *) Lemma type_type E L C T e :
Lemma type_equivalize_lft E L C T κ1 κ2 e : typed_body E L C T e typed_body E L C T e.
typed_body ((κ1 κ2) :: (κ2 κ1) :: E) L C T e Proof. done. Qed.
typed_body E ((κ1 [κ2]%list) :: L) C T e.
(** This lemma can replace [κ1] by [κ2] and vice versa in positions that
respect "semantic lifetime equivalence"; in particular, lifetimes of
references can be adjusted this way. However, it cannot replace lifetimes in
other type constructors, as those might only respect *syntactic* lifetime
equivalence. This lemma is *weaker* than what was in the original paper where
lifetimes could be replaced everywhere; it had to be adjusted for GhostCell.
See [typing.lib.diverging_static] for an example of how
[type_equivalize_lft_static] without this restriction ciuld be used to subvert
branding.
This is technically not a proper typing rule since the type system has no way
to express "subtyping wrt semantic lifetime inclusion". However, there is no
fundamental reason that we could not also reflect all these semantic facts on
the syntactic side, it would just be very clunky (and note that in Coq we do
not reflect this syntactic side anway). *)
Lemma type_equivalize_lft E L C T1 T2 κ1 κ2 e :
( tid, lft_ctx -∗ κ1 κ2 -∗ κ2 κ1 -∗ tctx_interp tid T1 -∗ tctx_interp tid T2)
( typed_body E L C T2 e)
typed_body E ((κ1 [κ2]) :: L) C T1 e.
Proof. Proof.
iIntros (He) "!#". iIntros (tid qE) "#HEAP #LFT Htl HE HL HC HT". iIntros (Hswitch He tid qmax) "#LFT #HE Htl [Hκ HL] HC HT".
rewrite /llctx_interp big_sepL_cons. iDestruct "HL" as "[Hκ HL]". iMod (lctx_equalize_lft_sem with "LFT Hκ") as "[Hκ1 Hκ2]".
iMod (lctx_equalize_lft with "LFT Hκ") as "[Hκ1 Hκ2]". iApply (He with "LFT HE Htl HL HC [-]").
iApply (He with "HEAP LFT Htl [Hκ1 Hκ2 HE] HL [HC] HT"). iApply (Hswitch with "LFT Hκ1 Hκ2"). done.
- rewrite /elctx_interp !big_sepL_cons. by iFrame. Qed.
- rewrite /elctx_interp !big_sepL_cons. iIntros "(_ & _ & HE)". Lemma type_equivalize_lft_static E L C T1 T2 κ e :
iApply "HC". done. ( tid, lft_ctx -∗ static κ -∗ tctx_interp tid T1 -∗ tctx_interp tid T2)
( typed_body E L C T2 e)
typed_body E ((κ []) :: L) C T1 e.
Proof.
iIntros (Hswitch He tid qmax) "#LFT #HE Htl [Hκ HL] HC HT".
iMod (lctx_equalize_lft_sem_static with "LFT Hκ") as "Hκ".
iApply (He with "LFT HE Htl HL HC [-]").
iApply (Hswitch with "LFT Hκ"). done.
Qed. Qed.
Lemma type_let' E L T1 T2 (T : tctx) C xb e e' : Lemma type_let' E L T1 T2 (T : tctx) C xb e e' :
Closed (xb :b: []) e' Closed (xb :b: []) e'
typed_instruction E L T1 e T2 typed_instruction E L T1 e T2 -∗
( v : val, typed_body E L C (T2 v ++ T) (subst' xb v e')) ( v : val, typed_body E L C (T2 v ++ T) (subst' xb v e')) -∗
typed_body E L C (T1 ++ T) (let: xb := e in e'). typed_body E L C (T1 ++ T) (let: xb := e in e').
Proof. Proof.
iIntros (Hc He He') "!#". iIntros (tid qE) "#HEAP #LFT Htl HE HL HC HT". rewrite tctx_interp_app. iIntros (Hc) "He He'". iIntros (tid qmax) "#LFT #HE Htl HL HC HT". rewrite tctx_interp_app.
iDestruct "HT" as "[HT1 HT]". wp_bind e. iApply (wp_wand with "[HE HL HT1 Htl]"). iDestruct "HT" as "[HT1 HT]". wp_bind e. iApply (wp_wand with "[He HL HT1 Htl]").
{ iApply (He with "HEAP LFT Htl HE HL HT1"). } { iApply ("He" with "LFT HE Htl HL HT1"). }
iIntros (v) "/= (Htl & HE & HL & HT2)". iApply wp_let; first wp_done. iIntros (v) "/= (Htl & HL & HT2)". wp_let.
iModIntro. iApply (He' with "HEAP LFT Htl HE HL HC [HT2 HT]"). iApply ("He'" with "LFT HE Htl HL HC [HT2 HT]").
rewrite tctx_interp_app. by iFrame. rewrite tctx_interp_app. by iFrame.
Qed. Qed.
(* We do not make the [typed_instruction] hypothesis part of the
Iris hypotheses, because we want to preserve the order of the
hypotheses. The is important, since proving [typed_instruction]
will instantiate [T1] and [T2], and hence we know what to search
for the following hypothesis. *)
Lemma type_let E L T T' T1 T2 C xb e e' : Lemma type_let E L T T' T1 T2 C xb e e' :
Closed (xb :b: []) e' Closed (xb :b: []) e'
typed_instruction E L T1 e T2 ( typed_instruction E L T1 e T2)
tctx_extract_ctx E L T1 T T' tctx_extract_ctx E L T1 T T'
( v : val, typed_body E L C (T2 v ++ T') (subst' xb v e')) ( v : val, typed_body E L C (T2 v ++ T') (subst' xb v e')) -∗
typed_body E L C T (let: xb := e in e'). typed_body E L C T (let: xb := e in e').
Proof. unfold tctx_extract_ctx. intros ?? -> ?. by eapply type_let'. Qed. Proof.
unfold tctx_extract_ctx. iIntros (? He ->) "?". iApply type_let'; last done.
iApply He.
Qed.
Lemma type_seq E L T T' T1 T2 C e e' : Lemma type_seq E L T T' T1 T2 C e e' :
Closed [] e' Closed [] e'
typed_instruction E L T1 e (λ _, T2) ( typed_instruction E L T1 e (λ _, T2))
tctx_extract_ctx E L T1 T T' tctx_extract_ctx E L T1 T T'
(typed_body E L C (T2 ++ T') e') typed_body E L C (T2 ++ T') e' -∗
typed_body E L C T (e ;; e'). typed_body E L C T (e ;; e').
Proof. intros. by eapply (type_let E L T T' T1 (λ _, T2)). Qed. Proof. iIntros. iApply (type_let E L T T' T1 (λ _, T2)); auto. Qed.
Lemma type_newlft {E L C T} κs e : Lemma type_newlft {E L C T} κs e :
Closed [] e Closed [] e
( κ, typed_body E ((κ κs) :: L) C T e) ( κ, typed_body E ((κ κs) :: L) C T e) -∗
typed_body E L C T (Newlft ;; e). typed_body E L C T (Newlft ;; e).
Proof. Proof.
iIntros (Hc He) "!#". iIntros (tid qE) "#HEAP #LFT Htl HE HL HC HT". iIntros (Hc) "He". iIntros (tid qmax) "#LFT #HE Htl HL HC HT".
iMod (lft_create with "LFT") as (Λ) "[Htk #Hinh]"; first done. iMod (lft_create with "LFT") as (Λ) "[Htk #Hinh]"; first done.
set (κ' := foldr () static κs). wp_seq. set (κ' := lft_intersect_list κs). wp_seq.
iApply (He (κ' Λ) with "HEAP LFT Htl HE [HL Htk] HC HT"). iApply ("He" $! (κ' Λ) with "LFT HE Htl [HL Htk] HC HT").
rewrite /llctx_interp big_sepL_cons. iFrame "HL". rewrite /llctx_interp /=. iFrame "HL".
iExists Λ. iSplit; first done. iFrame. done. iExists Λ. iSplit; first done.
destruct (decide (1 qmax)%Qp) as [_|Hlt%Qp.lt_nge].
- by iFrame "#∗".
- apply Qp.lt_sum in Hlt as [q' ->]. iDestruct "Htk" as "[$ Htk]".
iIntros "Htk'". iApply "Hinh". iFrame.
Qed. Qed.
(* TODO: It should be possible to show this while taking only one step. (* TODO: It should be possible to show this while taking only one step.
Right now, we could take two. *) Right now, we could take two. *)
Lemma type_endlft E L C T1 T2 κ κs e : Lemma type_endlft E L C T1 T2 κ κs e :
Closed [] e UnblockTctx κ T1 T2 Closed [] e UnblockTctx κ T1 T2
typed_body E L C T2 e typed_body E ((κ κs) :: L) C T1 (Endlft ;; e). typed_body E L C T2 e -∗ typed_body E ((κ κs) :: L) C T1 (Endlft ;; e).
Proof. Proof.
iIntros (Hc Hub He) "!#". iIntros (tid qE) "#HEAP #LFT Htl HE". rewrite /llctx_interp big_sepL_cons. iIntros (Hc Hub) "He". iIntros (tid qmax) "#LFT #HE Htl [Hκ HL] HC HT".
iIntros "[Hκ HL] HC HT". iDestruct "Hκ" as (Λ) "(% & Htok & #Hend)". iDestruct "Hκ" as (Λ) "(% & Htok & Hend)".
iSpecialize ("Hend" with "Htok"). wp_bind Endlft. iSpecialize ("Hend" with "Htok"). wp_bind Endlft.
iApply (wp_fupd_step with "Hend"); try done. wp_seq. iApply (wp_mask_mono _ (lftN lft_userE)); first done.
iIntros "!> #Hdead !>". wp_seq. iApply (He with "HEAP LFT Htl HE HL HC >"). iApply (wp_step_fupd with "Hend"); first set_solver+. wp_seq.
iIntros "#Hdead !>". wp_seq. iApply ("He" with "LFT HE Htl HL HC [> -]").
iApply (Hub with "[] HT"). simpl in *. subst κ. rewrite -lft_dead_or. auto. iApply (Hub with "[] HT"). simpl in *. subst κ. rewrite -lft_dead_or. auto.
Qed. Qed.
Lemma type_path_instr {E L} p ty : Lemma type_path_instr {E L} p ty :
typed_instruction_ty E L [p ty] p ty. typed_instruction_ty E L [p ty] p ty.
Proof. Proof.
iIntros "!# * _ _ $$$ ?". rewrite tctx_interp_singleton. iIntros (??) "_ _ $$ [? _]".
iApply (wp_hasty with "[-]"). done. iIntros (v) "_ Hv". iApply (wp_hasty with "[-]"); first done. iIntros (v) "_ Hv".
rewrite tctx_interp_singleton. iExists v. iFrame. by rewrite eval_path_of_val. rewrite tctx_interp_singleton. iExists v. iFrame. by rewrite eval_path_of_val.
Qed. Qed.
Lemma type_letpath {E L} ty C T T' x p e : Lemma type_letpath {E L} ty C T T' x p e :
Closed (x :b: []) e Closed (x :b: []) e
tctx_extract_hasty E L p ty T T' tctx_extract_hasty E L p ty T T'
( (v : val), typed_body E L C ((v ty) :: T') (subst' x v e)) ( (v : val), typed_body E L C ((v ty) :: T') (subst' x v e)) -∗
typed_body E L C T (let: x := p in e). typed_body E L C T (let: x := p in e).
Proof. Proof. iIntros. iApply type_let; [by apply type_path_instr|solve_typing|done]. Qed.
intros. eapply type_let; [done|by apply type_path_instr|solve_typing|by simpl].
Qed.
Lemma type_assign_instr {E L} ty ty1 ty1' p1 p2 : Lemma type_assign_instr {E L} ty ty1 ty1' p1 p2 :
typed_write E L ty1 ty ty1' ( typed_write E L ty1 ty ty1')
typed_instruction E L [p1 ty1; p2 ty] (p1 <- p2) (λ _, [p1 ty1']%TC). ( typed_instruction E L [p1 ty1; p2 ty] (p1 <- p2) (λ _, [p1 ty1'])).
Proof. Proof.
iIntros (Hwrt) "!#". iIntros (tid qE) "#HEAP #LFT $ HE HL". iIntros (Hwrt tid ?) "#LFT #HE $ HL".
rewrite tctx_interp_cons tctx_interp_singleton. iIntros "[Hp1 Hp2]". rewrite tctx_interp_cons tctx_interp_singleton. iIntros "[Hp1 Hp2]".
wp_bind p1. iApply (wp_hasty with "Hp1"). iIntros (v1) "% Hown1". wp_bind p1. iApply (wp_hasty with "Hp1"). iIntros (v1) "% Hown1".
wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2) "_ Hown2". wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2) "_ Hown2".
iMod (Hwrt with "* [] LFT HE HL Hown1") as (l vl) "([% %] & Hl & Hclose)"; first done. rewrite typed_write_eq in Hwrt.
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (Hwrt with "[] LFT HE HL Hown1") as (l vl) "([% %] & Hl & Hclose)"; first done.
subst v1. iDestruct (ty_size_eq with "Hown2") as "#Hsz". iDestruct "Hsz" as %Hsz. subst v1. iDestruct (ty_size_eq with "Hown2") as "#Hsz". iDestruct "Hsz" as %Hsz.
rewrite <-Hsz in *. destruct vl as [|v[|]]; try done. rewrite <-Hsz in *. destruct vl as [|v[|]]; try done.
rewrite heap_mapsto_vec_singleton. wp_write. rewrite heap_pointsto_vec_singleton. iApply wp_fupd. wp_write.
rewrite -heap_mapsto_vec_singleton. rewrite -heap_pointsto_vec_singleton.
iMod ("Hclose" with "* [Hl Hown2]") as "($ & $ & Hown)". iMod ("Hclose" with "[Hl Hown2]") as "(HL & Hown)".
{ iExists _. iFrame. } { iExists _. iFrame. }
iDestruct ("HLclose" with "HL") as "$".
rewrite tctx_interp_singleton tctx_hasty_val' //. rewrite tctx_interp_singleton tctx_hasty_val' //.
Qed. Qed.
Lemma type_assign {E L} ty1 ty ty1' C T T' p1 p2 e: Lemma type_assign {E L} ty1 ty ty1' C T T' p1 p2 e:
Closed [] e Closed [] e
tctx_extract_ctx E L [p1 ty1; p2 ty] T T' tctx_extract_ctx E L [p1 ty1; p2 ty] T T'
typed_write E L ty1 ty ty1' ( typed_write E L ty1 ty ty1')
typed_body E L C ((p1 ty1') :: T') e typed_body E L C ((p1 ty1') :: T') e -∗
typed_body E L C T (p1 <- p2 ;; e). typed_body E L C T (p1 <- p2 ;; e).
Proof. intros. eapply type_seq; [done |by apply type_assign_instr|done|done]. Qed. Proof. iIntros. by iApply type_seq; first apply type_assign_instr. Qed.
Lemma type_deref_instr {E L} ty ty1 ty1' p : Lemma type_deref_instr {E L} ty ty1 ty1' p :
ty.(ty_size) = 1%nat typed_read E L ty1 ty ty1' ty.(ty_size) = 1%nat ( typed_read E L ty1 ty ty1')
typed_instruction E L [p ty1] (!p) (λ v, [p ty1'; v ty]%TC). ( typed_instruction E L [p ty1] (!p) (λ v, [p ty1'; v ty])).
Proof. Proof.
iIntros (Hsz Hread) "!#". iIntros (tid qE) "#HEAP #LFT Htl HE HL Hp". iIntros (Hsz Hread tid qmax) "#LFT #HE Htl HL Hp".
rewrite tctx_interp_singleton. wp_bind p. iApply (wp_hasty with "Hp"). rewrite tctx_interp_singleton. wp_bind p. iApply (wp_hasty with "Hp").
iIntros (v) "% Hown". iIntros (v) "% Hown".
iMod (Hread with "* [] LFT Htl HE HL Hown") as rewrite typed_read_eq in Hread.
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (Hread with "[] LFT HE Htl HL Hown") as
(l vl q) "(% & Hl & Hown & Hclose)"; first done. (l vl q) "(% & Hl & Hown & Hclose)"; first done.
subst v. iDestruct (ty_size_eq with "Hown") as "#>%". rewrite ->Hsz in *. subst v. iDestruct (ty_size_eq with "Hown") as "#>%". rewrite ->Hsz in *.
destruct vl as [|v [|]]; try done. destruct vl as [|v [|]]; try done.
rewrite heap_mapsto_vec_singleton. wp_read. rewrite heap_pointsto_vec_singleton. iApply wp_fupd. wp_read.
iMod ("Hclose" with "Hl") as "($ & $ & Hown2)". iMod ("Hclose" with "Hl") as "($ & HL & Hown2)".
iDestruct ("HLclose" with "HL") as "$".
rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
by iFrame. by iFrame.
Qed. Qed.
...@@ -221,63 +292,60 @@ Section typing_rules. ...@@ -221,63 +292,60 @@ Section typing_rules.
Lemma type_deref {E L} ty1 C T T' ty ty1' x p e: Lemma type_deref {E L} ty1 C T T' ty ty1' x p e:
Closed (x :b: []) e Closed (x :b: []) e
tctx_extract_hasty E L p ty1 T T' tctx_extract_hasty E L p ty1 T T'
typed_read E L ty1 ty ty1' ( typed_read E L ty1 ty ty1')
ty.(ty_size) = 1%nat ty.(ty_size) = 1%nat
( (v : val), typed_body E L C ((p ty1') :: (v ty) :: T') (subst' x v e)) ( (v : val), typed_body E L C ((p ty1') :: (v ty) :: T') (subst' x v e)) -∗
typed_body E L C T (let: x := !p in e). typed_body E L C T (let: x := !p in e).
Proof. Proof. iIntros. by iApply type_let; [apply type_deref_instr|solve_typing|]. Qed.
intros. eapply type_let; [done|by apply type_deref_instr|solve_typing|by simpl].
Qed.
Lemma type_memcpy_iris E qE L qL tid ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 : Lemma type_memcpy_iris E L qmax qL tid ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 :
Z.of_nat (ty.(ty_size)) = n Z.of_nat (ty.(ty_size)) = n
typed_write E L ty1 ty ty1' -∗ typed_read E L ty2 ty ty2' -∗ typed_write E L ty1 ty ty1' -∗ typed_read E L ty2 ty ty2' -∗
{{{ heap_ctx lft_ctx na_own tid elctx_interp E qE llctx_interp L qL {{{ lft_ctx elctx_interp E na_own tid llctx_interp_noend qmax L qL
tctx_elt_interp tid (p1 ty1) tctx_elt_interp tid (p2 ty2) }}} tctx_elt_interp tid (p1 ty1) tctx_elt_interp tid (p2 ty2) }}}
(p1 < !{n}p2) (p1 <-{n} !p2)
{{{ RET #(); na_own tid elctx_interp E qE llctx_interp L qL {{{ RET #; na_own tid llctx_interp_noend qmax L qL
tctx_elt_interp tid (p1 ty1') tctx_elt_interp tid (p2 ty2') }}}. tctx_elt_interp tid (p1 ty1') tctx_elt_interp tid (p2 ty2') }}}.
Proof. Proof.
iIntros (<-) "#Hwrt #Hread !#". iIntros (<-) "#Hwrt #Hread !>".
iIntros (Φ) "(#HEAP & #LFT & Htl & [HE1 HE2] & [HL1 HL2] & [Hp1 Hp2]) HΦ". iIntros (Φ) "(#LFT & #HE & Htl & [HL1 HL2] & [Hp1 Hp2]) HΦ".
wp_bind p1. iApply (wp_hasty with "Hp1"). iIntros (v1) "% Hown1". wp_bind p1. iApply (wp_hasty with "Hp1"). iIntros (v1) "% Hown1".
wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2) "% Hown2". wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2) "% Hown2".
iMod ("Hwrt" with "* [] LFT HE1 HL1 Hown1") rewrite typed_write_eq typed_read_eq.
iMod ("Hwrt" with "[] LFT HE HL1 Hown1")
as (l1 vl1) "([% %] & Hl1 & Hcl1)"; first done. as (l1 vl1) "([% %] & Hl1 & Hcl1)"; first done.
iMod ("Hread" with "* [] LFT Htl HE2 HL2 Hown2") iMod ("Hread" with "[] LFT HE Htl HL2 Hown2")
as (l2 vl2 q2) "(% & Hl2 & Hown2 & Hcl2)"; first done. as (l2 vl2 q2) "(% & Hl2 & Hown2 & Hcl2)"; first done.
iDestruct (ty_size_eq with "Hown2") as "#>%". subst v1 v2. iApply wp_fupd. iDestruct (ty_size_eq with "Hown2") as "#>%". subst v1 v2. iApply wp_fupd.
iApply (wp_memcpy with "[$HEAP $Hl1 $Hl2]"); first done; try congruence; []. iApply (wp_memcpy with "[$Hl1 $Hl2]"); try congruence; [].
iNext. iIntros "[Hl1 Hl2]". iApply ("HΦ" with ">"). rewrite !tctx_hasty_val' //. iNext. iIntros "[Hl1 Hl2]". iApply ("HΦ" with "[> -]"). rewrite !tctx_hasty_val' //.
iMod ("Hcl1" with "[Hl1 Hown2]") as "($ & $ & $)". iMod ("Hcl1" with "[Hl1 Hown2]") as "($ & $)".
{ iExists _. iFrame. } { iExists _. iFrame. }
iMod ("Hcl2" with "Hl2") as "($ & $ & $ & $)". done. iMod ("Hcl2" with "Hl2") as "($ & $ & $)". done.
Qed. Qed.
Lemma type_memcpy_instr {E L} ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 : Lemma type_memcpy_instr {E L} ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 :
Z.of_nat (ty.(ty_size)) = n Z.of_nat (ty.(ty_size)) = n
typed_write E L ty1 ty ty1' typed_read E L ty2 ty ty2' ( typed_write E L ty1 ty ty1')
typed_instruction E L [p1 ty1; p2 ty2] (p1 <⋯ !{n}p2) ( typed_read E L ty2 ty ty2')
(λ _, [p1 ty1'; p2 ty2']%TC). typed_instruction E L [p1 ty1; p2 ty2] (p1 <-{n} !p2)
(λ _, [p1 ty1'; p2 ty2']).
Proof. Proof.
iIntros (Hsz Hwrt Hread) "!#". iIntros (tid qE) "#HEAP #LFT Htl HE HL HT". iIntros (Hsz Hwrt Hread tid qmax) "#LFT #HE Htl HL HT".
iApply (type_memcpy_iris with "[] [] * [$HEAP $LFT $Htl $HE $HL HT]"); try done. iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
(* TODO: This is kind of silly, why can't I iApply the assumptions directly? *) iApply (type_memcpy_iris with "[] [] [$LFT $Htl $HE $HL HT]"); try done.
{ iPoseProof Hwrt as "Hwrt". done. }
{ iPoseProof Hread as "Hread". done. }
{ by rewrite tctx_interp_cons tctx_interp_singleton. } { by rewrite tctx_interp_cons tctx_interp_singleton. }
rewrite tctx_interp_cons tctx_interp_singleton. auto. rewrite tctx_interp_cons tctx_interp_singleton.
iIntros "!> ($ & HL & $ & $)". by iApply "HLclose".
Qed. Qed.
Lemma type_memcpy {E L} ty1 ty2 (n : Z) C T T' ty ty1' ty2' p1 p2 e: Lemma type_memcpy {E L} ty ty1 ty2 (n : Z) C T T' ty1' ty2' p1 p2 e:
Closed [] e Closed [] e
tctx_extract_ctx E L [p1 ty1; p2 ty2] T T' tctx_extract_ctx E L [p1 ty1; p2 ty2] T T'
typed_write E L ty1 ty ty1' ( typed_write E L ty1 ty ty1')
typed_read E L ty2 ty ty2' ( typed_read E L ty2 ty ty2')
Z.of_nat (ty.(ty_size)) = n Z.of_nat (ty.(ty_size)) = n
typed_body E L C ((p1 ty1') :: (p2 ty2') :: T') e typed_body E L C ((p1 ty1') :: (p2 ty2') :: T') e -∗
typed_body E L C T (p1 <⋯ !{n}p2;; e). typed_body E L C T (p1 <-{n} !p2;; e).
Proof. Proof. iIntros. by iApply type_seq; first eapply (type_memcpy_instr ty ty1 ty1'). Qed.
intros. by eapply type_seq; [|by eapply (type_memcpy_instr ty ty1 ty1')|..].
Qed.
End typing_rules. End typing_rules.
From iris.base_logic Require Import big_op. From iris.proofmode Require Import proofmode.
From iris.proofmode Require Import tactics.
From lrust.typing Require Export type. From lrust.typing Require Export type.
From lrust.typing Require Import lft_contexts type_context programs. From lrust.typing Require Import lft_contexts type_context programs.
Set Default Proof Using "Type". From iris.prelude Require Import options.
Section shr_bor. Section shr_bor.
Context `{typeG Σ}. Context `{!typeGS Σ}.
Program Definition shr_bor (κ : lft) (ty : type) : type := Program Definition shr_bor (κ : lft) (ty : type) : type :=
{| st_own tid vl := ( (l:loc), vl = [ #l ] ty.(ty_shr) κ tid l)%I |}. {| st_own tid vl :=
Next Obligation. match vl return _ with
iIntros (κ ty tid vl) "H". iDestruct "H" as (l) "[% _]". by subst. | [ #(LitLoc l) ] => ty.(ty_shr) κ tid l
| _ => False
end%I |}.
Next Obligation. by iIntros (κ ty tid [|[[]|][]]) "H". Qed.
Next Obligation. intros κ ty tid [|[[]|][]]; apply _. Qed.
Global Instance shr_bor_wf κ ty `{!TyWf ty} : TyWf (shr_bor κ ty) :=
{ ty_lfts := [κ]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty κ }.
Lemma shr_type_incl κ1 κ2 ty1 ty2 :
κ2 κ1 -∗ type_incl ty1 ty2 -∗ type_incl (shr_bor κ1 ty1) (shr_bor κ2 ty2).
Proof.
iIntros "#Hκ #[_ [_ Hty]]". iApply type_incl_simple_type. simpl.
iIntros "!>" (tid [|[[]|][]]) "H"; try done. iApply "Hty".
iApply (ty1.(ty_shr_mono) with "Hκ"). done.
Qed. Qed.
Global Instance shr_mono E L : Global Instance shr_mono E L :
Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) shr_bor. Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) shr_bor.
Proof. Proof.
intros κ1 κ2 ty1 ty2 Hty. apply subtype_simple_type. intros κ1 κ2 ty1 ty2 Hty.
iIntros (??) "#LFT #HE #HL H". iDestruct ( with "HE HL") as "#Hκ". iIntros (??) "/= HL". iDestruct ( with "HL") as "#Hincl".
iDestruct "H" as (l) "(% & H)". subst. iExists _. iSplit. done. iDestruct (Hty with "HL") as "#Hty". iIntros "!> #HE".
iApply (ty2.(ty_shr_mono) with "LFT Hκ"). iDestruct ("Hincl" with "HE") as "%".
iDestruct (Hty with "* [] [] []") as "(_ & _ & #Hs1)"; [done..|clear Hty]. iApply shr_type_incl.
by iApply "Hs1". - by iApply lft_incl_syn_sem.
- by iApply "Hty".
Qed. Qed.
Global Instance shr_mono_flip E L : Global Instance shr_mono_flip E L :
Proper (lctx_lft_incl E L ==> flip (subtype E L) ==> flip (subtype E L)) shr_bor. Proper (lctx_lft_incl E L ==> flip (subtype E L) ==> flip (subtype E L)) shr_bor.
...@@ -30,68 +44,58 @@ Section shr_bor. ...@@ -30,68 +44,58 @@ Section shr_bor.
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) shr_bor. Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) shr_bor.
Proof. intros ??[] ??[]. by split; apply shr_mono. Qed. Proof. intros ??[] ??[]. by split; apply shr_mono. Qed.
Global Instance shr_contractive κ : Contractive (shr_bor κ). Global Instance shr_type_contractive κ : TypeContractive (shr_bor κ).
Proof. Proof.
intros n ?? EQ. unfold shr_bor. f_equiv. rewrite st_dist_unfold. intros n ???. apply: ty_of_st_type_ne. dist_later_fin_intro.
f_contractive=> /= tid vl. repeat f_equiv. apply EQ. solve_type_proper.
Qed. Qed.
Global Instance shr_ne κ n : Proper (dist n ==> dist n) (shr_bor κ).
Proof. apply contractive_ne, _. Qed. Global Instance shr_ne κ : NonExpansive (shr_bor κ).
Proof. apply type_contractive_ne, _. Qed.
Global Instance shr_send κ ty : Global Instance shr_send κ ty :
Sync ty Send (shr_bor κ ty). Sync ty Send (shr_bor κ ty).
Proof. Proof. by iIntros (Hsync tid1 tid2 [|[[]|][]]) "H"; try iApply Hsync. Qed.
iIntros (Hsync tid1 tid2 vl) "H". iDestruct "H" as (l) "[% Hshr]".
iExists _. iSplit; first done. by iApply Hsync.
Qed.
End shr_bor. End shr_bor.
Notation "&shr{ κ } ty" := (shr_bor κ ty) Notation "&shr{ κ }" := (shr_bor κ) (format "&shr{ κ }") : lrust_type_scope.
(format "&shr{ κ } ty", at level 20, right associativity) : lrust_type_scope.
Section typing. Section typing.
Context `{typeG Σ}. Context `{!typeGS Σ}.
Lemma shr_mono' E L κ1 κ2 ty1 ty2 : Lemma shr_mono' E L κ1 κ2 ty1 ty2 :
lctx_lft_incl E L κ2 κ1 subtype E L ty1 ty2 lctx_lft_incl E L κ2 κ1 subtype E L ty1 ty2
subtype E L (&shr{κ1} ty1) (&shr{κ2} ty2). subtype E L (&shr{κ1}ty1) (&shr{κ2}ty2).
Proof. by intros; apply shr_mono. Qed. Proof. by intros; apply shr_mono. Qed.
Lemma shr_proper' E L κ ty1 ty2 : Lemma shr_proper' E L κ1 κ2 ty1 ty2 :
eqtype E L ty1 ty2 eqtype E L (&shr{κ} ty1) (&shr{κ} ty2). lctx_lft_eq E L κ1 κ2 eqtype E L ty1 ty2 eqtype E L (&shr{κ1}ty1) (&shr{κ2}ty2).
Proof. by intros; apply shr_proper. Qed. Proof. by intros; apply shr_proper. Qed.
Lemma tctx_reborrow_shr E L p ty κ κ' : Lemma tctx_reborrow_shr E L p ty κ κ' :
lctx_lft_incl E L κ' κ lctx_lft_incl E L κ' κ
tctx_incl E L [p &shr{κ}ty] [p &shr{κ'}ty; p {κ} &shr{κ}ty]. tctx_incl E L [p &shr{κ}ty] [p &shr{κ'}ty; p {κ} &shr{κ}ty].
Proof. Proof.
iIntros (Hκκ' tid ??) "#LFT HE HL H". iIntros (Hκκ' tid ??) "#LFT #HE HL [H _]". iDestruct (Hκκ' with "HL HE") as "%".
iDestruct (elctx_interp_persist with "HE") as "#HE'". iFrame. rewrite /tctx_interp /=.
iDestruct (llctx_interp_persist with "HL") as "#HL'". iFrame "HE HL". iDestruct "H" as ([[]|]) "[% #Hshr]"; try done. iModIntro. iSplit.
iDestruct (Hκκ' with "HE' HL'") as "Hκκ'". - iExists _. iSplit; first done. iApply (ty_shr_mono with "[] Hshr").
rewrite /tctx_interp big_sepL_singleton big_sepL_cons big_sepL_singleton. by iApply lft_incl_syn_sem.
iDestruct "H" as (v) "[% #H]". iDestruct "H" as (l) "[EQ Hshr]". - iSplit=> //. iExists _. auto.
iDestruct "EQ" as %[=->]. simpl. iModIntro. iSplit.
- iExists _. iSplit. done. iExists _. iSplit. done.
by iApply (ty_shr_mono with "LFT Hκκ' Hshr").
- iExists _. iSplit. done. iIntros "_". iIntros "!>".
iExists _. auto.
Qed. Qed.
Lemma read_shr E L κ ty : Lemma read_shr E L κ ty :
Copy ty lctx_lft_alive E L κ typed_read E L (&shr{κ}ty) ty (&shr{κ}ty). Copy ty lctx_lft_alive E L κ typed_read E L (&shr{κ}ty) ty (&shr{κ}ty).
Proof. Proof.
iIntros (Hcopy Halive) "!#". iIntros (v tid F qE qL ?) "#LFT Htl HE HL Hown". rewrite typed_read_eq. iIntros (Hcopy Halive) "!>".
iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first set_solver. iIntros ([[]|] tid F qmax qL ?) "#LFT #HE Htl HL #Hshr"; try done.
iDestruct "Hown" as (l) "[EQ #Hshr]". iDestruct "EQ" as %[=->]. iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first solve_ndisj.
assert (shrN (lrustN : coPset)) by solve_ndisj. (* set_solver needs some help. *) iMod (copy_shr_acc with "LFT Hshr Htl Hκ") as (q') "(Htl & H↦ & Hcl)"; first solve_ndisj.
iMod (copy_shr_acc with "LFT Hshr Htl Hκ") as (q') "(Htl & H↦ & Hcl)". { rewrite ->shr_locsE_shrN. solve_ndisj. }
{ set_solver. } { rewrite ->shr_locsE_shrN. set_solver. }
iDestruct "H↦" as (vl) "[>Hmt #Hown]". iModIntro. iExists _, _, _. iDestruct "H↦" as (vl) "[>Hmt #Hown]". iModIntro. iExists _, _, _.
iSplit; first done. iFrame "∗#". iIntros "Hmt". iFrame "∗#". iSplit; first done. iIntros "Hmt".
iMod ("Hcl" with "Htl [Hmt]") as "[$ Hκ]". iMod ("Hcl" with "Htl [Hmt]") as "[$ Hκ]"; first by iExists _; iFrame.
{ iExists _. iFrame "∗#". } iApply ("Hclose" with "Hκ").
iMod ("Hclose" with "Hκ") as "[$ $]". iExists _. auto.
Qed. Qed.
End typing. End typing.
Hint Resolve shr_mono' shr_proper' : lrust_typing. Global Hint Resolve shr_mono' shr_proper' read_shr : lrust_typing.