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
  • HumamAlhusaini/lambda-rust
12 results
Show changes
Showing
with 6430 additions and 0 deletions
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import typing lib.option.
From iris.prelude Require Import options.
(* Typing "problem case #3" from:
http://smallcultfollowing.com/babysteps/blog/2016/04/27/non-lexical-lifetimes-introduction/
Differences with the original implementation:
We ignore dropping.
We have to add a Copy bound on the key type.
We do not support panic, hence we do not support option::unwrap. We use
unwrap_or as a replacement.
*)
Section non_lexical.
Context `{!typeGS Σ} (HashMap K V : type)
`{!TyWf hashmap, !TyWf K, !TyWf V, !Copy K}
(defaultv get_mut insert: val).
Hypothesis defaultv_type :
typed_val defaultv (fn() V).
Hypothesis get_mut_type :
typed_val get_mut (fn( '(α, β), ; &uniq{α} hashmap, &shr{β} K)
option (&uniq{α} V)).
Hypothesis insert_type :
typed_val insert (fn( α, ; &uniq{α} hashmap, K, V) option V).
Definition get_default : val :=
fn: ["map"; "key"] :=
let: "get_mut" := get_mut in
let: "map'" := !"map" in
Newlft;;
Newlft;;
letalloc: "map0" <- "map'" in
letalloc: "key0" <- "key" in
letcall: "o" := "get_mut" ["map0"; "key0"]%E in
Endlft;;
withcont: "k":
case: !"o" of
[ (* None *)
Endlft;;
let: "insert" := insert in
Newlft;;
letalloc: "map0" <- "map'" in
letalloc: "key0" <-{K.(ty_size)} !"key" in
let: "defaultv" := defaultv in
letcall: "default0" := "defaultv" [] in
letcall: "old" := "insert" ["map0"; "key0"; "default0"]%E in
Endlft;;
delete [ #(option V).(ty_size); "old"];; (* We should drop here. *)
Newlft;;
letalloc: "map0" <- "map'" in
letalloc: "key0" <- "key" in
letcall: "r'" := "get_mut" ["map0"; "key0"]%E in
Endlft;;
let: "unwrap" := option_unwrap (&uniq{static}V) in
letcall: "r" := "unwrap" ["r'"]%E in
"k" ["r"];
(* Some *)
letalloc: "r" <-{1} !("o" + #1) in
"k" ["r"]
]
cont: "k" ["r"] :=
delete [ #(option (&uniq{static}V)).(ty_size); "o"];;
delete [ #1; "map"];; delete [ #K.(ty_size); "key"];; (* We should drop key here *)
return: ["r"].
Lemma get_default_type :
typed_val get_default (fn( m, ; &uniq{m} hashmap, K) &uniq{m} V).
Proof using Type*.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (m ϝ ret p). inv_vec p=>map key. simpl_subst.
iApply type_let; [iApply get_mut_type|solve_typing|].
iIntros (get_mut'). simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (map'). simpl_subst.
iApply (type_newlft [m]). iIntros (κ).
iApply (type_newlft [ϝ]). iIntros (α).
iApply (type_letalloc_1 (&uniq{κ}_)); [solve_typing..|]. iIntros (map0). simpl_subst.
iApply (type_letalloc_1 (&shr{α}K)); [solve_typing..|]. iIntros (key0). simpl_subst.
iApply (type_letcall (κ, α)); [solve_typing..|]. iIntros (o). simpl_subst.
iApply type_endlft.
iApply (type_cont [_] [ϝ []]
(λ r, [o box (Π[uninit 1;uninit 1]); map box (uninit 1);
key box K; (r!!!0%fin:val) box (&uniq{m} V)])).
{ iIntros (k). simpl_subst.
iApply type_case_own;
[solve_typing| constructor; [|constructor; [|constructor]]; left].
- iApply type_endlft.
iApply type_let; [iApply insert_type|solve_typing|].
iIntros (insert'). simpl_subst.
iApply (type_newlft [ϝ]). iIntros (β). clear map0 key0.
iApply (type_letalloc_1 (&uniq{β}_)); [solve_typing..|].
iIntros (map0). simpl_subst.
iApply (type_letalloc_n _(box K)); [solve_typing..|]. iIntros (key0). simpl_subst.
iApply type_let; [iApply defaultv_type|solve_typing|].
iIntros (defaultv'). simpl_subst.
iApply (type_letcall ()); [solve_typing..|]. iIntros (default0). simpl_subst.
iApply (type_letcall β); [solve_typing..|]. iIntros (old). simpl_subst.
iApply type_endlft.
iApply type_delete; [solve_typing..|].
iApply (type_newlft [ϝ]). iIntros (γ). clear map0 key0.
iApply (type_letalloc_1 (&uniq{m}_)); [solve_typing..|].
iIntros (map0). simpl_subst.
iApply (type_letalloc_1 (&shr{γ}K)); [solve_typing..|].
iIntros (key0). simpl_subst.
iApply (type_letcall (m, γ)); [solve_typing..|]. iIntros (r'). simpl_subst.
iApply type_endlft.
iApply type_let; [iApply (option_unwrap_type (&uniq{m}V))|solve_typing|].
iIntros (unwrap'). simpl_subst.
iApply (type_letcall ()); [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_jump; solve_typing.
- (* Use lifetime equalization to replace one lifetime by another:
[&uniq{κ} V] becomes [&uniq{m} V] in the type of [o +ₗ #1]. *)
iApply (type_equivalize_lft _ _ _ _ [_; _; _; _; _; _; _; _]).
{ iIntros (tid) "#LFT #Hκ1 #Hκ2 ($ & Href & $ & $ & $ & $ & $ & $ & _)".
rewrite -tctx_interp_singleton.
iApply (type_incl_tctx_incl with "[] Href").
iApply own_type_incl; first done.
iNext. iApply uniq_type_incl.
- iApply "Hκ2".
- iApply type_equal_refl. }
iApply (type_letalloc_n (&uniq{m}V) (own_ptr _ (&uniq{m}V))); [solve_typing..|].
iIntros (r). simpl_subst.
iApply type_jump; solve_typing. }
iIntros "!> %k %args". inv_vec args=>r. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End non_lexical.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Section rebor.
Context `{!typeGS Σ}.
Definition rebor : val :=
fn: ["t1"; "t2"] :=
Newlft;;
letalloc: "x" <- "t1" in
let: "x'" := !"x" in let: "y" := "x'" + #0 in
"x" <- "t2";;
let: "y'" := !"y" in
letalloc: "r" <- "y'" in
Endlft ;; delete [ #2; "t1"] ;; delete [ #2; "t2"] ;;
delete [ #1; "x"] ;; return: ["r"].
Lemma rebor_type :
typed_val rebor (fn(; Π[int; int], Π[int; int]) int).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([] ϝ ret p). inv_vec p=>t1 t2. simpl_subst.
iApply (type_newlft []). iIntros (α).
iApply (type_letalloc_1 (&uniq{α}(Π[int; int]))); [solve_typing..|]. iIntros (x). simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iApply (type_letpath (&uniq{α}int)); [solve_typing|]. iIntros (y). simpl_subst.
iApply (type_assign _ (&uniq{α}(Π[int; int]))); [solve_typing..|].
iApply type_deref; [solve_typing..|]. iIntros (y'). simpl_subst.
iApply type_letalloc_1; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_endlft.
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End rebor.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Section unbox.
Context `{!typeGS Σ}.
Definition unbox : val :=
fn: ["b"] :=
let: "b'" := !"b" in let: "bx" := !"b'" in
letalloc: "r" <- "bx" + #0 in
delete [ #1; "b"] ;; return: ["r"].
Lemma ubox_type :
typed_val unbox (fn( α, ; &uniq{α}(box(Π[int; int]))) &uniq{α}int).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret b). inv_vec b=>b. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (b'); simpl_subst.
iApply type_deref_uniq_own; [solve_typing..|]. iIntros (bx); simpl_subst.
iApply type_letalloc_1; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End unbox.
From lrust.lang Require Import proofmode.
From lrust.typing Require Export lft_contexts type bool.
From iris.prelude Require Import options.
Import uPred.
Section fixpoint_def.
Context `{!typeGS Σ}.
Context (T : type type) {HT: TypeContractive T}.
Global Instance type_inhabited : Inhabited type := populate bool.
Local Instance type_2_contractive : Contractive (Nat.iter 2 T).
Proof using Type*.
intros n ? **. simpl.
by apply dist_later_S, type_dist2_dist_later, HT, HT, type_later_dist2_later.
Qed.
Definition type_fixpoint : type := fixpointK 2 T.
(* The procedure for computing ty_lfts and ty_wf_E for fixpoints is
the following:
- We do a first pass for computing [ty_lfts].
- In a second pass, we compute [ty_wf_E], by using the result of the
first pass.
I believe this gives the right sets for all types that we can define in
Rust, but I do not have any proof of this.
TODO : investigate this in more detail. *)
Global Instance type_fixpoint_wf `{!∀ `{!TyWf ty}, TyWf (T ty)} : TyWf type_fixpoint :=
let lfts :=
let _ : TyWf type_fixpoint := {| ty_lfts := []; ty_wf_E := [] |} in
ty_lfts (T type_fixpoint)
in
let wf_E :=
let _ : TyWf type_fixpoint := {| ty_lfts := lfts; ty_wf_E := [] |} in
ty_wf_E (T type_fixpoint)
in
{| ty_lfts := lfts; ty_wf_E := wf_E |}.
End fixpoint_def.
Lemma type_fixpoint_ne `{!typeGS Σ} (T1 T2 : type type)
`{!TypeContractive T1, !TypeContractive T2} n :
( t, T1 t {n} T2 t) type_fixpoint T1 {n} type_fixpoint T2.
Proof. eapply fixpointK_ne; apply type_contractive_ne, _. Qed.
Section fixpoint.
Context `{!typeGS Σ}.
Context (T : type type) {HT: TypeContractive T}.
Global Instance type_fixpoint_copy :
( `(!Copy ty), Copy (T ty)) Copy (type_fixpoint T).
Proof.
intros ?. unfold type_fixpoint. apply fixpointK_ind.
- apply type_contractive_ne, _.
- apply copy_equiv.
- exists bool. apply _.
- done.
- (* If Copy was an Iris assertion, this would be trivial -- we'd just
use limit_preserving_and directly. However, on the Coq side, it is
more convenient as a record -- so this is where we pay. *)
eapply (limit_preserving_ext (λ _, _ _)).
{ split; (intros [H1 H2]; split; [apply H1|apply H2]). }
apply limit_preserving_and; repeat (apply limit_preserving_forall=> ?).
+ apply bi.limit_preserving_Persistent; solve_proper.
+ apply limit_preserving_impl', bi.limit_preserving_emp_valid;
solve_proper_core ltac:(fun _ => eapply ty_size_ne || f_equiv).
Qed.
Global Instance type_fixpoint_send :
( `(!Send ty), Send (T ty)) Send (type_fixpoint T).
Proof.
intros ?. unfold type_fixpoint. apply fixpointK_ind.
- apply type_contractive_ne, _.
- apply send_equiv.
- exists bool. apply _.
- done.
- repeat (apply limit_preserving_forall=> ?).
apply bi.limit_preserving_entails; solve_proper.
Qed.
Global Instance type_fixpoint_sync :
( `(!Sync ty), Sync (T ty)) Sync (type_fixpoint T).
Proof.
intros ?. unfold type_fixpoint. apply fixpointK_ind.
- apply type_contractive_ne, _.
- apply sync_equiv.
- exists bool. apply _.
- done.
- repeat (apply limit_preserving_forall=> ?).
apply bi.limit_preserving_entails; solve_proper.
Qed.
Lemma type_fixpoint_unfold : type_fixpoint T T (type_fixpoint T).
Proof. apply fixpointK_unfold. by apply type_contractive_ne. Qed.
Lemma fixpoint_unfold_eqtype E L : eqtype E L (type_fixpoint T) (T (type_fixpoint T)).
Proof. apply type_equal_eqtype, type_equal_equiv, type_fixpoint_unfold. Qed.
End fixpoint.
Section subtyping.
Context `{!typeGS Σ} (E : elctx) (L : llctx).
(* TODO : is there a way to declare these as a [Proper] instances ? *)
Lemma fixpoint_mono T1 `{!TypeContractive T1} T2 `{!TypeContractive T2} :
( ty1 ty2, subtype E L ty1 ty2 subtype E L (T1 ty1) (T2 ty2))
subtype E L (type_fixpoint T1) (type_fixpoint T2).
Proof.
intros H12. rewrite /type_fixpoint. apply fixpointK_ind.
- apply type_contractive_ne, _.
- intros ?? EQ ?. etrans; last done. by apply equiv_subtype.
- by eexists _.
- intros. setoid_rewrite (fixpoint_unfold_eqtype T2). by apply H12.
- repeat (apply limit_preserving_forall=> ?).
apply bi.limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_proper T1 `{!TypeContractive T1} T2 `{!TypeContractive T2} :
( ty1 ty2, eqtype E L ty1 ty2 eqtype E L (T1 ty1) (T2 ty2))
eqtype E L (type_fixpoint T1) (type_fixpoint T2).
Proof.
intros H12. rewrite /type_fixpoint. apply fixpointK_ind.
- apply type_contractive_ne, _.
- intros ?? EQ ?. etrans; last done. by apply equiv_eqtype.
- by eexists _.
- intros. setoid_rewrite (fixpoint_unfold_eqtype T2). by apply H12.
- apply limit_preserving_and; repeat (apply limit_preserving_forall=> ?);
apply bi.limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_unfold_subtype_l ty T `{!TypeContractive T} :
subtype E L ty (T (type_fixpoint T)) subtype E L ty (type_fixpoint T).
Proof. intros. by rewrite fixpoint_unfold_eqtype. Qed.
Lemma fixpoint_unfold_subtype_r ty T `{!TypeContractive T} :
subtype E L (T (type_fixpoint T)) ty subtype E L (type_fixpoint T) ty.
Proof. intros. by rewrite fixpoint_unfold_eqtype. Qed.
Lemma fixpoint_unfold_eqtype_l ty T `{!TypeContractive T} :
eqtype E L ty (T (type_fixpoint T)) eqtype E L ty (type_fixpoint T).
Proof. intros. by rewrite fixpoint_unfold_eqtype. Qed.
Lemma fixpoint_unfold_eqtype_r ty T `{!TypeContractive T} :
eqtype E L (T (type_fixpoint T)) ty eqtype E L (type_fixpoint T) ty.
Proof. intros. by rewrite fixpoint_unfold_eqtype. Qed.
End subtyping.
Global Hint Resolve fixpoint_mono fixpoint_proper : lrust_typing.
(* These hints can loop if [fixpoint_mono] and [fixpoint_proper] have
not been tried before, so we give them a high cost *)
Global Hint Resolve fixpoint_unfold_subtype_l|100 : lrust_typing.
Global Hint Resolve fixpoint_unfold_subtype_r|100 : lrust_typing.
Global Hint Resolve fixpoint_unfold_eqtype_l|100 : lrust_typing.
Global Hint Resolve fixpoint_unfold_eqtype_r|100 : lrust_typing.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import vector list.
From lrust.typing Require Export type.
From lrust.typing Require Import own programs cont.
From iris.prelude Require Import options.
Section fn.
Context `{!typeGS Σ} {A : Type} {n : nat}.
Record fn_params := FP { fp_E : lft elctx; fp_tys : vec type n; fp_ty : type }.
Definition FP_wf E (tys : vec type n) `{!ListTyWf tys} ty `{!TyWf ty} :=
FP (λ ϝ, E ϝ ++ tyl_wf_E tys ++ tyl_outlives_E tys ϝ ++
ty_wf_E ty ++ ty_outlives_E ty ϝ)
tys ty.
(* The other alternative for defining the fn type would be to state
that the value applied to its parameters is a typed body whose type
is the return type.
That would be slightly simpler, but, unfortunately, we are no longer
able to prove that this is contractive. *)
Program Definition fn (fp : A fn_params) : type :=
{| st_own tid vl := tc_opaque ( fb kb xb e H,
vl = [@RecV fb (kb::xb) e H] length xb = n
(x : A) (ϝ : lft) (k : val) (xl : vec val (length xb)),
typed_body ((fp x).(fp_E) ϝ) [ϝ []]
[kcont([ϝ []], λ v : vec _ 1, [(v!!!0%fin:val) box (fp x).(fp_ty)])]
(zip_with (TCtx_hasty of_val) xl
(box <$> (vec_to_list (fp x).(fp_tys))))
(subst_v (fb::kb::xb) (RecV fb (kb::xb) e:::k:::xl) e))%I |}.
Next Obligation.
iIntros (fp tid vl) "H". iDestruct "H" as (fb kb xb e ?) "[% _]". by subst.
Qed.
Next Obligation.
unfold tc_opaque. apply _.
Qed.
(* FIXME : This definition is less restrictive than the one used in
Rust. In Rust, the type of parameters are taken into account for
well-formedness, and all the liftime constrains relating a
generalized liftime are ignored. For simplicity, we ignore all of
them, but this is not very faithful. *)
Global Instance fn_wf fp : TyWf (fn fp) :=
{ ty_lfts := []; ty_wf_E := [] }.
Global Instance fn_send fp : Send (fn fp).
Proof. iIntros (tid1 tid2 vl). done. Qed.
Definition fn_params_rel (ty_rel : relation type) : relation fn_params :=
λ fp1 fp2,
Forall2 ty_rel fp2.(fp_tys) fp1.(fp_tys) ty_rel fp1.(fp_ty) fp2.(fp_ty)
pointwise_relation lft eq fp1.(fp_E) fp2.(fp_E).
Global Instance fp_tys_proper R :
Proper (flip (fn_params_rel R) ==> (Forall2 R : relation (vec _ _))) fp_tys.
Proof. intros ?? HR. apply HR. Qed.
Global Instance fp_tys_proper_flip R :
Proper (fn_params_rel R ==> flip (Forall2 R : relation (vec _ _))) fp_tys.
Proof. intros ?? HR. apply HR. Qed.
Global Instance fp_ty_proper R :
Proper (fn_params_rel R ==> R) fp_ty.
Proof. intros ?? HR. apply HR. Qed.
Global Instance fp_E_proper R :
Proper (fn_params_rel R ==> eq ==> eq) fp_E.
Proof. intros ?? HR ??->. apply HR. Qed.
Global Instance FP_proper R :
Proper (pointwise_relation lft eq ==>
flip (Forall2 R : relation (vec _ _)) ==> R ==>
fn_params_rel R) FP.
Proof. by split; [|split]. Qed.
Global Instance fn_type_contractive n' :
Proper (pointwise_relation A (fn_params_rel (type_dist2_later n')) ==>
type_dist2 n') fn.
Proof.
intros fp1 fp2 Hfp. apply ty_of_st_type_ne. dist_later_fin_intro.
constructor; unfold ty_own; simpl.
(* TODO: 'f_equiv' is slow here because reflexivity is slow. *)
(* The clean way to do this would be to have a metric on type contexts. Oh well. *)
intros tid vl. unfold typed_body.
do 12 f_equiv. f_contractive_fin.
do 18 ((eapply fp_E_proper; try reflexivity) || exact: Hfp || f_equiv).
- rewrite !cctx_interp_singleton /=. do 5 f_equiv.
rewrite !tctx_interp_singleton /tctx_elt_interp /=.
do 5 f_equiv. apply type_dist2_dist. apply Hfp.
- rewrite /tctx_interp !big_sepL_zip_with /=. do 4 f_equiv.
cut ( n tid p i, Proper (dist n ==> dist n)
(λ (l : list type),
match l !! i with
| Some ty => tctx_elt_interp tid (p ty) | None => emp
end)%I).
{ intros Hprop. apply Hprop, list_fmap_ne; last first.
- symmetry. eapply Forall2_impl; first apply Hfp. intros.
apply dist_later_S, type_dist2_dist_later. done.
- solve_proper. }
clear. intros n tid p i x y. rewrite list_dist_lookup=>/(_ i).
case _ : (x !! i)=>[tyx|]; case _ : (y !! i)=>[tyy|];
inversion_clear 1; [solve_proper|done].
Qed.
Global Instance fn_ne n' :
Proper (pointwise_relation A (fn_params_rel (dist n')) ==> dist n') fn.
Proof.
intros ?? Hfp. apply dist_later_S, type_dist2_dist_later.
apply fn_type_contractive=>u. split; last split.
- eapply Forall2_impl; first apply Hfp. intros. simpl.
apply type_dist_dist2. done.
- apply type_dist_dist2. apply Hfp.
- apply Hfp.
Qed.
End fn.
Global Arguments fn_params {_ _} _.
(* We use recursive notation for binders as well, to allow patterns
like '(a, b) to be used. In practice, only one binder is ever used,
but using recursive binders is the only way to make Coq accept
patterns. *)
(* FIXME : because of a bug in Coq, such patterns only work for
printing. Once on 8.6pl1, this should work. *)
Notation "'fn(∀' x .. x' ',' E ';' T1 ',' .. ',' TN ')' '→' R" :=
(fn (λ x, (.. (λ x',
FP_wf E%EL (Vector.cons T1%T .. (Vector.cons TN%T Vector.nil) ..) R%T)..)))
(at level 99, R at level 200, x binder, x' binder,
format "'fn(∀' x .. x' ',' E ';' T1 ',' .. ',' TN ')' '→' R") : lrust_type_scope.
Notation "'fn(∀' x .. x' ',' E ')' '→' R" :=
(fn (λ x, (.. (λ x', FP_wf E%EL Vector.nil R%T)..)))
(at level 99, R at level 200, x binder, x' binder,
format "'fn(∀' x .. x' ',' E ')' '→' R") : lrust_type_scope.
Notation "'fn(' E ';' T1 ',' .. ',' TN ')' '→' R" :=
(fn (λ _:(), FP_wf E%EL (Vector.cons T1%T .. (Vector.cons TN%T Vector.nil) ..) R%T))
(at level 99, R at level 200,
format "'fn(' E ';' T1 ',' .. ',' TN ')' '→' R") : lrust_type_scope.
Notation "'fn(' E ')' '→' R" :=
(fn (λ _:(), FP_wf E%EL Vector.nil R%T))
(at level 99, R at level 200,
format "'fn(' E ')' '→' R") : lrust_type_scope.
Global Instance elctx_empty : Empty (lft elctx) := λ ϝ, [].
Section typing.
Context `{!typeGS Σ}.
Lemma fn_subtype {A n} E0 L0 (fp fp' : A fn_params n) :
( x ϝ, let EE := E0 ++ (fp' x).(fp_E) ϝ in
elctx_sat EE L0 ((fp x).(fp_E) ϝ)
Forall2 (subtype EE L0) (fp' x).(fp_tys) (fp x).(fp_tys)
subtype EE L0 (fp x).(fp_ty) (fp' x).(fp_ty))
subtype E0 L0 (fn fp) (fn fp').
Proof.
intros Hcons. apply subtype_simple_type=>//= qmax qL. iIntros "HL0".
(* We massage things so that we can throw away HL0 before going under the box. *)
iAssert ( x ϝ, let EE := E0 ++ (fp' x).(fp_E) ϝ in (elctx_interp EE -∗
elctx_interp ((fp x).(fp_E) ϝ)
([ list] tys (zip (fp' x).(fp_tys) (fp x).(fp_tys)), type_incl (tys.1) (tys.2))
type_incl (fp x).(fp_ty) (fp' x).(fp_ty)))%I as "#Hcons".
{ iIntros (x ϝ). destruct (Hcons x ϝ) as (HE &Htys &Hty). clear Hcons.
iDestruct (HE with "HL0") as "#HE".
iDestruct (subtype_Forall2_llctx_noend with "HL0") as "#Htys"; first done.
iDestruct (Hty with "HL0") as "#Hty".
iClear "∗". iIntros "!> #HEE".
iSplit; last iSplit.
- by iApply "HE".
- by iApply "Htys".
- by iApply "Hty". }
iClear "∗". clear Hcons. iIntros "!> #HE0 * Hf".
iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst.
iExists fb, kb, xb, e, _. iSplit; first done. iSplit; first done. iNext.
rewrite /typed_body. iIntros (x ϝ k xl) "!> * #LFT #HE' Htl HL HC HT".
iDestruct ("Hcons" with "[$]") as "#(HE & Htys & Hty)".
iApply ("Hf" with "LFT HE Htl HL [HC] [HT]").
- unfold cctx_interp. iIntros (elt) "Helt".
iDestruct "Helt" as %->%elem_of_list_singleton. iIntros (ret) "Htl HL HT".
unfold cctx_elt_interp.
iApply ("HC" $! (_ cont(_, _)) with "[%] Htl HL [> -]").
{ by apply elem_of_list_singleton. }
rewrite /tctx_interp !big_sepL_singleton /=.
iDestruct "HT" as (v) "[HP Hown]". iExists v. iFrame "HP".
iDestruct (box_type_incl with "[$Hty]") as "(_ & #Hincl & _)".
by iApply "Hincl".
- iClear "Hf". rewrite /tctx_interp
-{2}(fst_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?length_vec_to_list //
-{2}(snd_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?length_vec_to_list //
!zip_with_fmap_r !(zip_with_zip (λ _ _, (_ _) _ _)) !big_sepL_fmap.
iApply (big_sepL_impl with "HT"). iIntros "!>".
iIntros (i [p [ty1' ty2']]) "#Hzip H /=".
iDestruct "H" as (v) "[? Hown]". iExists v. iFrame.
rewrite !lookup_zip_with.
iDestruct "Hzip" as %(? & ? & ([? ?] & (? & Hty'1 &
(? & Hty'2 & [=->->])%bind_Some)%bind_Some & [=->->->])%bind_Some)%bind_Some.
iDestruct (big_sepL_lookup with "Htys") as "#Hty'".
{ rewrite lookup_zip_with /=. erewrite Hty'2. simpl. by erewrite Hty'1. }
iDestruct (box_type_incl with "[$Hty']") as "(_ & #Hincl & _)".
by iApply "Hincl".
Qed.
(* This proper and the next can probably not be inferred, but oh well. *)
Global Instance fn_subtype' {A n} E0 L0 :
Proper (pointwise_relation A (fn_params_rel (n:=n) (subtype E0 L0)) ==>
subtype E0 L0) fn.
Proof.
intros fp1 fp2 Hfp. apply fn_subtype=>x ϝ. destruct (Hfp x) as (Htys & Hty & HE).
split; last split.
- rewrite (HE ϝ). solve_typing.
- eapply Forall2_impl; first eapply Htys. intros ??.
eapply subtype_weaken; last done. by apply submseteq_inserts_r.
- eapply subtype_weaken, Hty; last done. by apply submseteq_inserts_r.
Qed.
Global Instance fn_eqtype' {A n} E0 L0 :
Proper (pointwise_relation A (fn_params_rel (n:=n) (eqtype E0 L0)) ==>
eqtype E0 L0) fn.
Proof.
intros fp1 fp2 Hfp. split; eapply fn_subtype=>x ϝ; destruct (Hfp x) as (Htys & Hty & HE); (split; last split).
- rewrite (HE ϝ). solve_typing.
- eapply Forall2_impl; first eapply Htys. intros t1 t2 Ht.
eapply subtype_weaken; last apply Ht; last done. by apply submseteq_inserts_r.
- eapply subtype_weaken; last apply Hty; last done. by apply submseteq_inserts_r.
- rewrite (HE ϝ). solve_typing.
- symmetry in Htys. eapply Forall2_impl; first eapply Htys. intros t1 t2 Ht.
eapply subtype_weaken; last apply Ht; last done. by apply submseteq_inserts_r.
- eapply subtype_weaken; last apply Hty; last done. by apply submseteq_inserts_r.
Qed.
Lemma fn_subtype_specialize {A B n} (σ : A B) E0 L0 fp :
subtype E0 L0 (fn (n:=n) fp) (fn (fp σ)).
Proof.
apply subtype_simple_type=>//= qmax qL.
iIntros "_ !> _ * Hf". iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst.
iExists fb, kb, xb, e, _. iSplit; first done. iSplit; first done.
rewrite /typed_body. iNext. iIntros "*". iApply "Hf".
Qed.
(* In principle, proving this hard-coded to an empty L would be sufficient --
but then we would have to require elctx_sat as an Iris assumption. *)
Lemma type_call_iris' E L (κs : list lft) {A} x (ps : list path) qκs qmax qL tid
p (k : expr) (fp : A fn_params (length ps)) :
( ϝ, elctx_sat (((λ κ, ϝ κ) <$> κs) ++ E) L ((fp x).(fp_E) ϝ))
AsVal k
lft_ctx -∗ elctx_interp E -∗ na_own tid -∗ llctx_interp_noend qmax L qL -∗
qκs.[lft_intersect_list κs] -∗
tctx_elt_interp tid (p fn fp) -∗
([ list] y zip_with TCtx_hasty ps (box <$> vec_to_list (fp x).(fp_tys)),
tctx_elt_interp tid y) -∗
( ret, na_own tid top -∗ llctx_interp_noend qmax L qL -∗ qκs.[lft_intersect_list κs] -∗
(box (fp x).(fp_ty)).(ty_own) tid [ret] -∗
WP k [of_val ret] {{ _, cont_postcondition }}) -∗
WP (call: p ps k) {{ _, cont_postcondition }}.
Proof.
iIntros (HE [k' <-]) "#LFT #HE Htl HL Hκs Hf Hargs Hk".
wp_apply (wp_hasty with "Hf"). iIntros (v) "% Hf".
iApply (wp_app_vec _ _ (_::_) ((λ v, v = (λ: ["_r"], k' ["_r"])%V):::
vmap (λ ty (v : val), tctx_elt_interp tid (v box ty)) (fp x).(fp_tys))%I
with "[Hargs]").
- rewrite /=. iSplitR "Hargs".
{ simpl. iApply wp_value. by unlock. }
remember (fp_tys (fp x)) as tys. clear dependent k' p HE fp x.
iInduction ps as [|p ps] "IH" forall (tys); first by simpl.
simpl in tys. inv_vec tys=>ty tys. simpl.
iDestruct "Hargs" as "[HT Hargs]". iSplitL "HT".
+ iApply (wp_hasty with "HT"). iIntros (?). rewrite tctx_hasty_val. iIntros "? $".
+ iApply "IH". done.
- simpl. change (@length expr ps) with (length ps).
iIntros (vl'). inv_vec vl'=>kv vl; csimpl.
iIntros "[-> Hvl]". iDestruct "Hf" as (fb kb xb e ?) "[EQ [EQl #Hf]]".
iDestruct "EQ" as %[=->]. iDestruct "EQl" as %EQl.
revert vl fp HE. rewrite /= -EQl=>vl fp HE. wp_rec.
iMod (lft_create with "LFT") as (ϝ_inner) "[Htk #Hend]"; first done.
set (ϝ := ϝ_inner lft_intersect_list κs).
iSpecialize ("Hf" $! x ϝ _ vl). iDestruct (HE ϝ with "HL") as "#HE'".
destruct (Qp.lower_bound qκs 1) as (q0 & q'1 & q'2 & -> & Hsum1).
rewrite Hsum1. assert (q0 < 1)%Qp as Hq0.
{ apply Qp.lt_sum. eauto. }
clear Hsum1.
iDestruct "Htk" as "[Htk1 Htk2]".
iDestruct "Hκs" as "[Hκs1 Hκs2]".
iApply ("Hf" $! _ q0 with "LFT [] Htl [Hκs1 Htk1 Htk2] [Hk HL Hκs2]").
+ iApply "HE'". iFrame "HE".
iIntros "{$# Hf Hend HE' LFT HE %}". subst ϝ.
iApply big_sepL_forall.
iIntros (i [κ1 κ2] [κ [Hpair Helem]]%elem_of_list_lookup_2%elem_of_list_fmap).
injection Hpair as -> ->. iPureIntro. simpl.
eapply lft_incl_syn_trans; first by apply lft_intersect_incl_syn_r.
apply lft_intersect_list_elem_of_incl_syn.
done.
+ iSplitL; last done. iExists ϝ. rewrite left_id. iSplit; first done.
rewrite decide_False; last first.
{ apply Qp.lt_nge. done. }
subst ϝ. rewrite -!lft_tok_sep. iFrame. iIntros "[Htk1 _]".
rewrite -lft_dead_or. rewrite -bi.or_intro_l. iApply "Hend". iFrame.
+ iIntros (y) "IN {Hend}". iDestruct "IN" as %->%elem_of_list_singleton.
iIntros (args) "Htl [Hϝ _] [Hret _]". inv_vec args=>r.
iDestruct "Hϝ" as (κ') "(EQ & Htk & _)". iDestruct "EQ" as %EQ.
rewrite /= left_id in EQ. subst κ' ϝ.
rewrite decide_False; last first.
{ apply Qp.lt_nge. done. }
rewrite -lft_tok_sep. iDestruct "Htk" as "[_ Hκs1]". wp_rec.
iApply ("Hk" with "Htl HL [$Hκs1 $Hκs2]"). rewrite tctx_hasty_val. done.
+ rewrite /tctx_interp vec_to_list_map !zip_with_fmap_r
(zip_with_zip (λ v ty, (v, _))) zip_with_zip !big_sepL_fmap.
iApply (big_sepL_mono' with "Hvl"); last done. by iIntros (i [v ty']).
Qed.
Lemma type_call_iris E (κs : list lft) {A} x (ps : list path) qκs tid
f (k : expr) (fp : A fn_params (length ps)) :
( ϝ, elctx_sat (((λ κ, ϝ κ) <$> κs) ++ E) [] ((fp x).(fp_E) ϝ))
AsVal k
lft_ctx -∗ elctx_interp E -∗ na_own tid -∗
qκs.[lft_intersect_list κs] -∗
(fn fp).(ty_own) tid [f] -∗
([ list] y zip_with TCtx_hasty ps (box <$> vec_to_list (fp x).(fp_tys)),
tctx_elt_interp tid y) -∗
( ret, na_own tid top -∗ qκs.[lft_intersect_list κs] -∗
(box (fp x).(fp_ty)).(ty_own) tid [ret] -∗
WP k [of_val ret] {{ _, cont_postcondition }}) -∗
WP (call: f ps k) {{ _, cont_postcondition }}.
Proof.
iIntros (HE Hk') "#LFT #HE Htl Hκs Hf Hargs Hk". rewrite -tctx_hasty_val.
iApply (type_call_iris' with "LFT HE Htl [] Hκs Hf Hargs [Hk]"); [done..| |].
- instantiate (1 := 1%Qp). instantiate (1 := 1%Qp). by rewrite /llctx_interp_noend.
- iIntros "* Htl _". iApply "Hk". done.
Qed.
Lemma type_call' E L (κs : list lft) T p (ps : list path)
{A} (fp : A fn_params (length ps)) (k : val) x :
Forall (lctx_lft_alive E L) κs
( ϝ, elctx_sat (((λ κ, ϝ κ) <$> κs) ++ E) L ((fp x).(fp_E) ϝ))
typed_body E L [k cont(L, λ v : vec _ 1, ((v!!!0%fin:val) box (fp x).(fp_ty)) :: T)]
((p fn fp) ::
zip_with TCtx_hasty ps (box <$> vec_to_list (fp x).(fp_tys)) ++
T)
(call: p ps k).
Proof.
iIntros (Hκs HE tid qmax) "#LFT #HE Htl HL HC (Hf & Hargs & HT)".
iMod (lctx_lft_alive_tok_list _ _ κs with "HE HL") as (q) "(Hκs & HL & Hclose)"; [done..|].
iApply (type_call_iris' with "LFT HE Htl HL Hκs Hf Hargs"); [done|].
iIntros (r) "Htl HL Hκs Hret". iMod ("Hclose" with "Hκs HL") as "HL".
iSpecialize ("HC" with "[]"); first by (iPureIntro; apply elem_of_list_singleton).
iApply ("HC" $! [#r] with "Htl HL").
rewrite tctx_interp_cons tctx_hasty_val. iFrame.
Qed.
(* Specialized type_call': Adapted for use by solve_typing.
κs is still expected to be given manually. *)
Lemma type_call {A} κs x E L C T T' T'' p (ps : list path)
(fp : A fn_params (length ps)) k :
p fn fp T
Forall (lctx_lft_alive E L) κs
( ϝ, elctx_sat (((λ κ, ϝ κ) <$> κs) ++ E) L ((fp x).(fp_E) ϝ))
tctx_extract_ctx E L (zip_with TCtx_hasty ps
(box <$> vec_to_list (fp x).(fp_tys))) T T'
k cont(L, T'') C
( ret : val, tctx_incl E L ((ret box (fp x).(fp_ty))::T') (T'' [# ret]))
typed_body E L C T (call: p ps k).
Proof.
intros Hfn HL HE HTT' HC HT'T''.
iApply typed_body_mono; [| |done|by iApply type_call']; simpl.
- etrans.
+ eapply (incl_cctx_incl _ [_]); by intros ? ->%elem_of_list_singleton.
+ apply cctx_incl_cons; first done. intros args. by inv_vec args.
- etrans; last by apply (tctx_incl_frame_l [_]).
apply copy_elem_of_tctx_incl; last done. apply _.
Qed.
Lemma type_letcall {A} x E L C T T' p (ps : list path)
(fp : A fn_params (length ps)) b e :
Closed (b :b: []) e Closed [] p Forall (Closed []) ps
p fn fp T
Forall (lctx_lft_alive E L) (L.*1)
( ϝ, elctx_sat (((λ κ, ϝ κ) <$> (L.*1)) ++ E) L ((fp x).(fp_E) ϝ))
tctx_extract_ctx E L (zip_with TCtx_hasty ps
(box <$> vec_to_list (fp x).(fp_tys))) T T'
( ret : val, typed_body E L C ((ret box (fp x).(fp_ty))::T') (subst' b ret e)) -∗
typed_body E L C T (letcall: b := p ps in e).
Proof.
iIntros (?? Hpsc ????) "He".
iApply (type_cont_norec [_] _ (λ r, ((r!!!0%fin:val) box (fp x).(fp_ty)) :: T')).
- (* TODO : make [solve_closed] work here. *)
eapply is_closed_weaken; first done. set_solver+.
- (* TODO : make [solve_closed] work here. *)
rewrite /Closed /= !andb_True. split.
+ by eapply is_closed_weaken, list_subseteq_nil.
+ eapply Is_true_eq_left, forallb_forall, List.Forall_forall, Forall_impl=>//.
intros. eapply Is_true_eq_true, is_closed_weaken=>//. set_solver+.
- iIntros (k).
(* TODO : make [simpl_subst] work here. *)
change (subst' "_k" k (p ((λ: ["_r"], "_k" ["_r"])%E :: ps))) with
((subst "_k" k p) ((λ: ["_r"], k ["_r"])%E :: map (subst "_k" k) ps)).
rewrite is_closed_nil_subst //.
assert (map (subst "_k" k) ps = ps) as ->.
{ clear -Hpsc. induction Hpsc=>//=. rewrite is_closed_nil_subst //. congruence. }
iApply type_call; try done.
+ constructor.
+ done.
- simpl. iIntros (k ret). inv_vec ret=>ret. rewrite /subst_v /=.
rewrite ->(is_closed_subst []); last set_solver+; last first.
{ apply subst'_is_closed; last done. apply is_closed_of_val. }
(iApply typed_body_mono; last by iApply "He"); [|done..].
apply incl_cctx_incl. set_solver+.
Qed.
Lemma type_rec {A} E L fb (argsb : list binder) ef e n
(fp : A fn_params n) T `{!CopyC T, !SendC T, !Closed _ e} :
IntoVal ef (fnrec: fb argsb := e)
n = length argsb
( x ϝ (f : val) k (args : vec val (length argsb)),
typed_body ((fp x).(fp_E) ϝ) [ϝ []]
[k cont([ϝ []], λ v : vec _ 1, [(v!!!0%fin:val) box (fp x).(fp_ty)])]
((f fn fp) ::
zip_with (TCtx_hasty of_val) args
(box <$> vec_to_list (fp x).(fp_tys)) ++ T)
(subst_v (fb :: BNamed "return" :: argsb) (f ::: k ::: args) e)) -∗
typed_instruction_ty E L T ef (fn fp).
Proof.
iIntros (<- ->) "#Hbody /=". iIntros (tid qmax) "#LFT _ $ $ #HT". iApply wp_value.
rewrite tctx_interp_singleton. iLöb as "IH". iExists _. iSplit.
{ simpl. rewrite decide_True_pi. done. }
iExists fb, _, argsb, e, _. iSplit; first done. iSplit; first done. iNext.
iIntros (x ϝ k args) "!>". iIntros (tid' qmax') "_ HE Htl HL HC HT'".
iApply ("Hbody" with "LFT HE Htl HL HC").
rewrite tctx_interp_cons tctx_interp_app. iFrame "HT' IH".
by iApply sendc_change_tid.
Qed.
Lemma type_fn {A} E L (argsb : list binder) ef e n
(fp : A fn_params n) T `{!CopyC T, !SendC T, !Closed _ e} :
IntoVal ef (fn: argsb := e)
n = length argsb
( x ϝ k (args : vec val (length argsb)),
typed_body ((fp x).(fp_E) ϝ)
[ϝ []]
[k cont([ϝ []], λ v : vec _ 1, [(v!!!0%fin:val) box (fp x).(fp_ty)])]
(zip_with (TCtx_hasty of_val) args
(box <$> vec_to_list (fp x).(fp_tys)) ++ T)
(subst_v (BNamed "return" :: argsb) (k ::: args) e)) -∗
typed_instruction_ty E L T ef (fn fp).
Proof.
iIntros (??) "#He". iApply type_rec; try done. iIntros "!> *".
iApply typed_body_mono; last iApply "He"; try done.
eapply contains_tctx_incl. by constructor.
Qed.
End typing.
Global Hint Resolve fn_subtype : lrust_typing.
From iris.proofmode Require Import tactics.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Export type.
From lrust.typing Require Import bool programs.
Set Default Proof Using "Type".
From iris.prelude Require Import options.
Section int.
Context `{typeG Σ}.
Context `{!typeGS Σ}.
Program Definition int : type :=
{| st_own tid vl :=
......@@ -12,91 +12,75 @@ Section int.
| [ #(LitInt z)] => True
| _ => False
end%I |}.
Next Obligation. intros ? [|[[]|] []]; try iIntros "[]". auto. Qed.
Next Obligation. intros ? [|[[]|] []]; auto. Qed.
Next Obligation. intros ? [|[[]|] []]; apply _. Qed.
Global Instance int_wf : TyWf int := { ty_lfts := []; ty_wf_E := [] }.
Global Instance int_send : Send int.
Proof. done. Qed.
End int.
Section typing.
Context `{typeG Σ}.
Context `{!typeGS Σ}.
Lemma type_int_instr (z : Z) E L :
typed_instruction_ty E L [] #z int.
Lemma type_int_instr (z : Z) : typed_val #z int.
Proof.
iAlways. iIntros (tid qE) "_ _ $ $ $ _". wp_value.
by rewrite tctx_interp_singleton tctx_hasty_val.
iIntros (E L tid ?) "_ _ $ $ _". iApply wp_value.
by rewrite tctx_interp_singleton tctx_hasty_val' //.
Qed.
Lemma type_int (z : Z) E L C T x e :
Closed (x :b: []) e
( (v : val), typed_body E L C ((v int) :: T) (subst' x v e))
( (v : val), typed_body E L C ((v int) :: T) (subst' x v e)) -∗
typed_body E L C T (let: x := #z in e).
Proof.
intros. eapply type_let; [done|apply type_int_instr|solve_typing|done].
Qed.
Proof. iIntros. iApply type_let; [apply type_int_instr|solve_typing|done]. Qed.
Lemma type_plus_instr E L p1 p2 :
typed_instruction_ty E L [p1 int; p2 int] (p1 + p2) int.
typed_instruction_ty E L [p1 int; p2 int] (p1 + p2) int.
Proof.
iAlways. iIntros (tid qE) "_ _ $ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
iIntros "[Hp1 Hp2]".
wp_bind p1. iApply (wp_hasty with "Hp1").
iIntros ([[]|]) "_ H1"; try iDestruct "H1" as "[]".
wp_bind p2. iApply (wp_hasty with "Hp2").
iIntros ([[]|]) "_ H2"; try iDestruct "H2" as "[]".
iIntros (tid ?) "_ _ $ $ [Hp1 [Hp2 _]]".
wp_apply (wp_hasty with "Hp1"). iIntros ([[]|]) "_ H1"; try done.
wp_apply (wp_hasty with "Hp2"). iIntros ([[]|]) "_ H2"; try done.
wp_op. by rewrite tctx_interp_singleton tctx_hasty_val' //.
Qed.
Lemma type_plus E L C T T' p1 p2 x e :
Closed (x :b: []) e
tctx_extract_ctx E L [p1 int; p2 int] T T'
( (v : val), typed_body E L C ((v int) :: T') (subst' x v e))
( (v : val), typed_body E L C ((v int) :: T') (subst' x v e)) -∗
typed_body E L C T (let: x := p1 + p2 in e).
Proof.
intros. eapply type_let; [done|apply type_plus_instr|solve_typing|done].
Qed.
Proof. iIntros. iApply type_let; [iApply type_plus_instr|solve_typing|done]. Qed.
Lemma type_minus_instr E L p1 p2 :
typed_instruction_ty E L [p1 int; p2 int] (p1 - p2) int.
typed_instruction_ty E L [p1 int; p2 int] (p1 - p2) int.
Proof.
iAlways. iIntros (tid qE) "_ _ $ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
iIntros "[Hp1 Hp2]".
wp_bind p1. iApply (wp_hasty with "Hp1").
iIntros ([[]|]) "_ H1"; try iDestruct "H1" as "[]".
wp_bind p2. iApply (wp_hasty with "Hp2").
iIntros ([[]|]) "_ H2"; try iDestruct "H2" as "[]".
iIntros (tid ?) "_ _ $ $ [Hp1 [Hp2 _]]".
wp_apply (wp_hasty with "Hp1"). iIntros ([[]|]) "_ H1"; try done.
wp_apply (wp_hasty with "Hp2"). iIntros ([[]|]) "_ H2"; try done.
wp_op. by rewrite tctx_interp_singleton tctx_hasty_val' //.
Qed.
Lemma type_minus E L C T T' p1 p2 x e :
Closed (x :b: []) e
tctx_extract_ctx E L [p1 int; p2 int] T T'
( (v : val), typed_body E L C ((v int) :: T') (subst' x v e))
( (v : val), typed_body E L C ((v int) :: T') (subst' x v e)) -∗
typed_body E L C T (let: x := p1 - p2 in e).
Proof.
intros. eapply type_let; [done|apply type_minus_instr|solve_typing|done].
Qed.
Proof. iIntros. iApply type_let; [apply type_minus_instr|solve_typing|done]. Qed.
Lemma type_le_instr E L p1 p2 :
typed_instruction_ty E L [p1 int; p2 int] (p1 p2) bool.
typed_instruction_ty E L [p1 int; p2 int] (p1 p2) bool.
Proof.
iAlways. iIntros (tid qE) "_ _ $ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
iIntros "[Hp1 Hp2]".
wp_bind p1. iApply (wp_hasty with "Hp1").
iIntros ([[]|]) "_ H1"; try iDestruct "H1" as "[]".
wp_bind p2. iApply (wp_hasty with "Hp2").
iIntros ([[]|]) "_ H2"; try iDestruct "H2" as "[]".
wp_op; intros _; by rewrite tctx_interp_singleton tctx_hasty_val' //.
iIntros (tid ?) "_ _ $ $ [Hp1 [Hp2 _]]".
wp_apply (wp_hasty with "Hp1"). iIntros ([[]|]) "_ H1"; try done.
wp_apply (wp_hasty with "Hp2"). iIntros ([[]|]) "_ H2"; try done.
wp_op; case_bool_decide; by rewrite tctx_interp_singleton tctx_hasty_val' //.
Qed.
Lemma type_le E L C T T' p1 p2 x e :
Closed (x :b: []) e
tctx_extract_ctx E L [p1 int; p2 int] T T'
( (v : val), typed_body E L C ((v bool) :: T') (subst' x v e))
( (v : val), typed_body E L C ((v bool) :: T') (subst' x v e)) -∗
typed_body E L C T (let: x := p1 p2 in e).
Proof.
intros. eapply type_let; [done|apply type_le_instr|solve_typing|done].
Qed.
Proof. iIntros. iApply type_let; [apply type_le_instr|solve_typing|done]. Qed.
End typing.
From iris.proofmode Require Import proofmode.
From iris.bi Require Import fractional.
From lrust.lang Require Import proofmode.
From lrust.typing Require Export base.
From lrust.lifetime Require Import frac_borrow.
From iris.prelude Require Import options.
Definition elctx_elt : Type := lft * lft.
Notation elctx := (list elctx_elt).
Declare Scope lrust_elctx_scope.
Delimit Scope lrust_elctx_scope with EL.
(* We need to define [elctx] and [llctx] as notations to make eauto
work. But then, Coq is not able to bind them to their
notations, so we have to use [Arguments] everywhere. *)
Bind Scope lrust_elctx_scope with elctx_elt.
Notation "κ1 ⊑ₑ κ2" := (@pair lft lft κ1 κ2) (at level 70).
Definition llctx_elt : Type := lft * list lft.
Notation llctx := (list llctx_elt).
Notation "κ ⊑ₗ κl" := (@pair lft (list lft) κ κl) (at level 70).
Section lft_contexts.
Context `{!invGS Σ, !lftGS Σ lft_userE}.
Implicit Type (κ : lft).
(* External lifetime contexts. *)
Definition elctx_elt_interp (x : elctx_elt) : iProp Σ :=
(x.1 ˢʸⁿ x.2)⌝%I.
Definition elctx_interp (E : elctx) : iProp Σ :=
([ list] x E, elctx_elt_interp x)%I.
Global Instance elctx_interp_permut :
Proper (() ==> (⊣⊢)) elctx_interp.
Proof. intros ???. by apply big_opL_permutation. Qed.
Global Instance elctx_interp_persistent E :
Persistent (elctx_interp E).
Proof. apply _. Qed.
(* Local lifetime contexts. *)
(* To support [type_call_iris'], the local lifetime [x.1] may be an
intersection of not just the atomic lifetime [κ0] but also of some extra
lifetimes [κextra], of which a smaller fraction [qextra] is owned (multiplied
by [q] to remain fractional). Since [x.2] is empty, [type_call_iris'] can show
that [κ0 ⊓ κextra] is the same after execution the function as it was before,
which is sufficient to extract the lifetime tokens for [κextra] again. *)
Definition llctx_elt_interp (qmax : Qp) (x : llctx_elt) : iProp Σ :=
let κ' := lft_intersect_list (x.2) in
(* This is [qeff := min 1 qmax], i.e., we cap qmax at [1] -- but the old
std++ we use here does not yet have [min] on [Qp].
We do the cap so that we never need to have more than [1] of this token. *)
let qeff := (if decide (1 qmax) then 1 else qmax)%Qp in
( κ0, x.1 = κ' κ0 qeff.[κ0] (qeff.[κ0] ={lftN lft_userE}[lft_userE]▷=∗ [κ0]))%I.
(* Local lifetime contexts without the "ending a lifetime" viewshifts -- these
are fractional. *)
Definition llctx_elt_interp_noend (qmax : Qp) (x : llctx_elt) (q : Qp) : iProp Σ :=
let κ' := lft_intersect_list (x.2) in
let qeff := (if decide (1 qmax) then 1 else qmax)%Qp in
( κ0, x.1 = κ' κ0 (qeff * q).[κ0])%I.
Global Instance llctx_elt_interp_noend_fractional qmax x :
Fractional (llctx_elt_interp_noend qmax x).
Proof.
destruct x as [κ κs]. iIntros (q q'). iSplit; iIntros "H".
- iDestruct "H" as (κ0) "(% & Hq)".
rewrite Qp.mul_add_distr_l.
iDestruct "Hq" as "[Hq Hq']".
iSplitL "Hq"; iExists _; by iFrame "∗%".
- iDestruct "H" as "[Hq Hq']".
iDestruct "Hq" as (κ0) "(% & Hq)".
iDestruct "Hq'" as (κ0') "(% & Hq')". simpl in *.
rewrite (inj ((lft_intersect_list κs) ⊓.) κ0' κ0); last congruence.
iExists κ0. rewrite Qp.mul_add_distr_l. by iFrame "∗%".
Qed.
Lemma llctx_elt_interp_acc_noend qmax x :
llctx_elt_interp qmax x
llctx_elt_interp_noend qmax x 1 (llctx_elt_interp_noend qmax x 1 -∗ llctx_elt_interp qmax x).
Proof.
destruct x as [κ κs].
iIntros "H". iDestruct "H" as (κ0) "(% & Hq & Hend)". iSplitL "Hq".
{ iExists κ0. rewrite Qp.mul_1_r. by iFrame "∗%". }
iIntros "H". iDestruct "H" as (κ0') "(% & Hq')". simpl in *.
rewrite (inj ((lft_intersect_list κs) ⊓.) κ0' κ0); last congruence.
iExists κ0. rewrite Qp.mul_1_r. by iFrame "∗%".
Qed.
Definition llctx_interp (qmax : Qp) (L : llctx) : iProp Σ :=
([ list] x L, llctx_elt_interp qmax x)%I.
Definition llctx_interp_noend (qmax : Qp) (L : llctx) (q : Qp) : iProp Σ :=
([ list] x L, llctx_elt_interp_noend qmax x q)%I.
Global Instance llctx_interp_fractional qmax L :
Fractional (llctx_interp_noend qmax L).
Proof. intros ??. rewrite -big_sepL_sep. by setoid_rewrite <-fractional. Qed.
Global Instance llctx_interp_as_fractional qmax L q :
AsFractional (llctx_interp_noend qmax L q) (llctx_interp_noend qmax L) q.
Proof. split; first done. apply _. Qed.
Global Instance frame_llctx_interp p qmax L q1 q2 q :
FrameFractionalQp q1 q2 q
Frame p (llctx_interp_noend qmax L q1) (llctx_interp_noend qmax L q2)
(llctx_interp_noend qmax L q) | 5.
Proof. apply: frame_fractional. Qed.
Global Instance llctx_interp_permut qmax :
Proper (() ==> (⊣⊢)) (llctx_interp qmax).
Proof. intros ???. by apply big_opL_permutation. Qed.
Lemma llctx_interp_acc_noend qmax L :
llctx_interp qmax L
llctx_interp_noend qmax L 1 (llctx_interp_noend qmax L 1 -∗ llctx_interp qmax L).
Proof.
rewrite /llctx_interp. setoid_rewrite llctx_elt_interp_acc_noend at 1.
rewrite big_sepL_sep. iIntros "[$ Hclose]". iIntros "Hnoend".
iCombine "Hclose Hnoend" as "H".
rewrite /llctx_interp_noend -big_sepL_sep.
setoid_rewrite bi.wand_elim_l. done.
Qed.
Lemma lctx_equalize_lft_sem qmax κ1 κ2 `{!frac_borG Σ} :
lft_ctx -∗ llctx_elt_interp qmax (κ1 [κ2]%list) ={}=∗
κ1 κ2 κ2 κ1.
Proof.
iIntros "#LFT". iDestruct 1 as (κ) "(% & Hκ & _)"; simplify_eq/=.
iMod (lft_eternalize with "Hκ") as "#Hincl".
iModIntro. iSplit.
- iApply lft_incl_trans; iApply lft_intersect_incl_l.
- iApply (lft_incl_glb with "[]"); first iApply (lft_incl_glb with "[]").
+ iApply lft_incl_refl.
+ iApply lft_incl_static.
+ iApply lft_incl_trans; last done. iApply lft_incl_static.
Qed.
Lemma lctx_equalize_lft_sem_static qmax κ `{!frac_borG Σ} :
lft_ctx -∗ llctx_elt_interp qmax (κ []%list) ={}=∗
static κ.
Proof.
iIntros "#LFT". iDestruct 1 as (κ') "(% & Hκ & _)"; simplify_eq/=.
iMod (lft_eternalize with "Hκ") as "#Hincl".
iModIntro.
- iApply (lft_incl_glb with "[]"); simpl.
+ iApply lft_incl_refl.
+ done.
Qed.
(* Lifetime inclusion *)
Context (E : elctx) (L : llctx).
Definition lctx_lft_incl κ κ' : Prop :=
qmax qL, llctx_interp_noend qmax L qL (elctx_interp E -∗ κ ˢʸⁿ κ')%I.
Definition lctx_lft_eq κ κ' : Prop :=
lctx_lft_incl κ κ' lctx_lft_incl κ' κ.
Lemma lctx_lft_incl_incl_noend κ κ' : lctx_lft_incl κ κ' lctx_lft_incl κ κ'.
Proof. done. Qed.
Lemma lctx_lft_incl_incl κ κ' qmax :
lctx_lft_incl κ κ' llctx_interp qmax L -∗ (elctx_interp E -∗ κ ˢʸⁿ κ')%I.
Proof.
iIntros (Hincl) "HL".
iDestruct (llctx_interp_acc_noend with "HL") as "[HL _]".
iApply Hincl. done.
Qed.
Lemma lctx_lft_incl_refl κ : lctx_lft_incl κ κ.
Proof.
iIntros (qmax qL) "_ !> _".
iPureIntro. apply lft_incl_syn_refl.
Qed.
Global Instance lctx_lft_incl_preorder : PreOrder lctx_lft_incl.
Proof.
split; first by intros ?; apply lctx_lft_incl_refl.
iIntros (??? H1 H2 ??) "HL".
iDestruct (H1 with "HL") as "#H1".
iDestruct (H2 with "HL") as "#H2".
iClear "∗". iIntros "!> #HE".
iDestruct ("H1" with "HE") as "%". iDestruct ("H2" with "HE") as "%".
iPureIntro.
by eapply lft_incl_syn_trans.
Qed.
Global Instance lctx_lft_incl_proper :
Proper (lctx_lft_eq ==> lctx_lft_eq ==> iff) lctx_lft_incl.
Proof. intros ??[] ??[]. split; intros; by etrans; [|etrans]. Qed.
Global Instance lctx_lft_eq_equivalence : Equivalence lctx_lft_eq.
Proof.
split.
- by split.
- intros ?? Heq; split; apply Heq.
- intros ??? H1 H2. split; etrans; (apply H1 || apply H2).
Qed.
Lemma lctx_lft_incl_static κ : lctx_lft_incl κ static.
Proof. iIntros (qmax qL) "_ !> _". iPureIntro. apply lft_incl_syn_static. Qed.
Lemma lctx_lft_incl_local κ κ' κs :
κ κs L κ' κs lctx_lft_incl κ κ'.
Proof.
iIntros (? Hκ'κs qmax qL) "H".
iDestruct (big_sepL_elem_of with "H") as "H"; first done.
iDestruct "H" as (κ'') "[EQ _]". iDestruct "EQ" as %EQ.
simpl in EQ; subst. iIntros "!> #HE".
iPureIntro.
eapply lft_incl_syn_trans; first apply lft_intersect_incl_syn_l.
by apply lft_intersect_list_elem_of_incl_syn.
Qed.
Lemma lctx_lft_incl_local' κ κ' κ'' κs :
κ κs L κ' κs lctx_lft_incl κ' κ'' lctx_lft_incl κ κ''.
Proof. intros. etrans; last done. by eapply lctx_lft_incl_local. Qed.
Lemma lctx_lft_incl_external κ κ' : κ κ' E lctx_lft_incl κ κ'.
Proof.
iIntros (???) "_ !> #HE".
rewrite /elctx_interp /elctx_elt_interp big_sepL_elem_of //. done.
Qed.
Lemma lctx_lft_incl_external' κ κ' κ'' :
κ κ' E lctx_lft_incl κ' κ'' lctx_lft_incl κ κ''.
Proof. intros. etrans; last done. by eapply lctx_lft_incl_external. Qed.
Lemma lctx_lft_incl_intersect κ κ' κ'' :
lctx_lft_incl κ κ' lctx_lft_incl κ κ''
lctx_lft_incl (κ κ) (κ' κ'').
Proof.
iIntros (Hκ' Hκ'' ??) "HL".
iDestruct (Hκ' with "HL") as "#Hκ'".
iDestruct (Hκ'' with "HL") as "#Hκ''".
iIntros "!> #HE".
iDestruct ("Hκ'" with "HE") as "%".
iDestruct ("Hκ''" with "HE") as "%".
iPureIntro.
by apply lft_intersect_mono_syn.
Qed.
Lemma lctx_lft_incl_intersect_l κ κ' κ'' :
lctx_lft_incl κ κ'
lctx_lft_incl (κ κ'') κ'.
Proof.
iIntros (Hκ' ??) "HL".
iDestruct (Hκ' with "HL") as "#Hκ'".
iIntros "!> #HE". iDestruct ("Hκ'" with "HE") as "%". iPureIntro.
eapply lft_incl_syn_trans; last eassumption.
by apply lft_intersect_incl_syn_l.
Qed.
Lemma lctx_lft_incl_intersect_r κ κ' κ'' :
lctx_lft_incl κ κ'
lctx_lft_incl (κ'' κ) κ'.
Proof.
iIntros (Hκ' ??) "HL".
iDestruct (Hκ' with "HL") as "#Hκ'".
iIntros "!> #HE". iDestruct ("Hκ'" with "HE") as "%". iPureIntro.
eapply lft_incl_syn_trans; last eassumption.
by apply lft_intersect_incl_syn_r.
Qed.
(* Lifetime aliveness *)
Definition lctx_lft_alive (κ : lft) : Prop :=
F qmax qL, lftN F elctx_interp E -∗ llctx_interp_noend qmax L qL ={F}=∗
q', q'.[κ] (q'.[κ] ={F}=∗ llctx_interp_noend qmax L qL).
Lemma lctx_lft_alive_tok_noend κ F qmax q :
lftN F
lctx_lft_alive κ elctx_interp E -∗ llctx_interp_noend qmax L q ={F}=∗
q', q'.[κ] llctx_interp_noend qmax L q'
(q'.[κ] -∗ llctx_interp_noend qmax L q' ={F}=∗ llctx_interp_noend qmax L q).
Proof.
iIntros (? Hal) "#HE [HL1 HL2]".
iMod (Hal with "HE HL1") as (q') "[Htok Hclose]"; first done.
destruct (Qp.lower_bound (q/2) q') as (qq & q0 & q'0 & Hq & ->). rewrite Hq.
iExists qq. iDestruct "HL2" as "[$ HL]". iDestruct "Htok" as "[$ Htok]".
iIntros "!> Htok' HL'". iCombine "HL'" "HL" as "HL". rewrite -Hq. iFrame.
iApply "Hclose". iFrame.
Qed.
Lemma lctx_lft_alive_tok κ F qmax :
lftN F
lctx_lft_alive κ elctx_interp E -∗ llctx_interp qmax L ={F}=∗
q', q'.[κ] llctx_interp_noend qmax L q'
(q'.[κ] -∗ llctx_interp_noend qmax L q' ={F}=∗ llctx_interp qmax L).
Proof.
iIntros (? Hal) "#HE HL".
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (lctx_lft_alive_tok_noend with "HE HL") as (q') "(Hκ & HL & Hclose)"; [done..|].
iModIntro. iExists q'. iFrame. iIntros "Hκ HL".
iApply "HLclose". iApply ("Hclose" with "Hκ"). done.
Qed.
Lemma lctx_lft_alive_tok_noend_list κs F qmax q :
lftN F Forall lctx_lft_alive κs
elctx_interp E -∗ llctx_interp_noend qmax L q ={F}=∗
q', q'.[lft_intersect_list κs] llctx_interp_noend qmax L q'
(q'.[lft_intersect_list κs] -∗ llctx_interp_noend qmax L q' ={F}=∗ llctx_interp_noend qmax L q).
Proof.
iIntros (? Hκs) "#HE". iInduction κs as [|κ κs] "IH" forall (q Hκs).
{ iIntros "HL !>". iExists _. iFrame "HL". iSplitL; first iApply lft_tok_static.
iIntros "_ $". done. }
inversion_clear Hκs.
iIntros "HL". iMod (lctx_lft_alive_tok_noend κ with "HE HL")as (q') "(Hκ & HL & Hclose1)"; [solve_typing..|].
iMod ("IH" with "[//] HL") as (q'') "(Hκs & HL & Hclose2)".
destruct (Qp.lower_bound q' q'') as (qq & q0 & q'0 & -> & ->).
iExists qq. iDestruct "HL" as "[$ HL2]". iDestruct "Hκ" as "[Hκ1 Hκ2]".
iDestruct "Hκs" as "[Hκs1 Hκs2]". iModIntro. simpl. rewrite -lft_tok_sep. iSplitL "Hκ1 Hκs1".
{ by iFrame. }
iIntros "[Hκ1 Hκs1] HL1". iMod ("Hclose2" with "[$Hκs1 $Hκs2] [$HL1 $HL2]") as "HL".
iMod ("Hclose1" with "[$Hκ1 $Hκ2] HL") as "HL". done.
Qed.
Lemma lctx_lft_alive_tok_list κs F qmax :
lftN F Forall lctx_lft_alive κs
elctx_interp E -∗ llctx_interp qmax L ={F}=∗
q', q'.[lft_intersect_list κs] llctx_interp_noend qmax L q'
(q'.[lft_intersect_list κs] -∗ llctx_interp_noend qmax L q' ={F}=∗ llctx_interp qmax L).
Proof.
iIntros (? Hal) "#HE HL".
iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]".
iMod (lctx_lft_alive_tok_noend_list with "HE HL") as (q') "(Hκ & HL & Hclose)"; [done..|].
iModIntro. iExists q'. iFrame. iIntros "Hκ HL".
iApply "HLclose". iApply ("Hclose" with "Hκ"). done.
Qed.
Lemma lctx_lft_alive_static : lctx_lft_alive static.
Proof.
iIntros (F qmax qL ?) "_ $". iExists 1%Qp. iSplitL; last by auto.
by iApply lft_tok_static.
Qed.
Lemma lctx_lft_alive_local κ κs:
κ κs L Forall lctx_lft_alive κs lctx_lft_alive κ.
Proof.
iIntros ([i HL]%elem_of_list_lookup_1 Hκs F qmax qL ?) "#HE HL".
iDestruct "HL" as "[HL1 HL2]". rewrite {2}/llctx_interp_noend /llctx_elt_interp.
iDestruct (big_sepL_lookup_acc with "HL2") as "[Hκ Hclose]"; first done.
iDestruct "Hκ" as (κ0) "(EQ & Htok)". simpl. iDestruct "EQ" as %->.
iAssert ( q', q'.[lft_intersect_list κs]
(q'.[lft_intersect_list κs] ={F}=∗ llctx_interp_noend qmax L (qL / 2)))%I
with "[> HE HL1]" as "H".
{ move:(qL/2)%Qp=>qL'. clear HL.
iInduction Hκs as [|κ κs ?] "IH" forall (qL').
- iExists 1%Qp. iFrame. iSplitR; last by auto. iApply lft_tok_static.
- iDestruct "HL1" as "[HL1 HL2]".
iMod ( with "HE HL1") as (q') "[Htok' Hclose]"; first done.
iMod ("IH" with "HL2") as (q'') "[Htok'' Hclose']".
destruct (Qp.lower_bound q' q'') as (q0 & q'2 & q''2 & -> & ->).
iExists q0. rewrite -lft_tok_sep. iDestruct "Htok'" as "[$ Hr']".
iDestruct "Htok''" as "[$ Hr'']". iIntros "!>[Hκ Hfold]".
iMod ("Hclose" with "[$Hκ $Hr']") as "$". iApply "Hclose'". iFrame. }
iDestruct "H" as (q') "[Htok' Hclose']". rewrite -{5}(Qp.div_2 qL).
set (qeff := (if decide (1 qmax) then 1 else qmax)%Qp).
destruct (Qp.lower_bound q' (qeff * (qL/2))) as (q0 & q'2 & q''2 & -> & Hmax).
iExists q0. rewrite -(lft_tok_sep q0). rewrite Hmax.
iDestruct "Htok" as "[$ Htok]".
iDestruct "Htok'" as "[$ Htok']". iIntros "!>[Hfold Hκ0]".
iMod ("Hclose'" with "[$Hfold $Htok']") as "$".
rewrite /llctx_interp /llctx_elt_interp. iApply "Hclose".
iExists κ0. rewrite Hmax. iFrame. auto.
Qed.
Lemma lctx_lft_alive_incl κ κ':
lctx_lft_alive κ lctx_lft_incl κ κ' lctx_lft_alive κ'.
Proof.
iIntros (Hal Hinc F qmax qL ?) "#HE HL".
iAssert (κ κ')%I with "[#]" as "#Hincl".
{ rewrite Hinc. iDestruct ("HL" with "HE") as "%".
by iApply lft_incl_syn_sem. }
iMod (Hal with "HE HL") as (q') "[Htok Hclose]"; first done.
iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose']"; first done.
iExists q''. iIntros "{$Htok}!>Htok". iApply ("Hclose" with "[> -]").
by iApply "Hclose'".
Qed.
Lemma lctx_lft_alive_external κ κ':
κ κ' E lctx_lft_alive κ lctx_lft_alive κ'.
Proof.
intros. by eapply lctx_lft_alive_incl, lctx_lft_incl_external.
Qed.
(* External lifetime context satisfiability *)
Definition elctx_sat E' : Prop :=
qmax qL, llctx_interp_noend qmax L qL -∗ (elctx_interp E -∗ elctx_interp E').
Lemma elctx_sat_nil : elctx_sat [].
Proof. iIntros (??) "_ !> _". by rewrite /elctx_interp /=. Qed.
Lemma elctx_sat_lft_incl E' κ κ' :
lctx_lft_incl κ κ' elctx_sat E' elctx_sat ((κ κ') :: E').
Proof.
iIntros (Hκκ' HE' qmax qL) "HL".
iDestruct (Hκκ' with "HL") as "#Hincl".
iDestruct (HE' with "HL") as "#HE'".
iClear "∗". iIntros "!> #HE". iSplit.
- by iApply "Hincl".
- by iApply "HE'".
Qed.
Lemma elctx_sat_app E1 E2 :
elctx_sat E1 elctx_sat E2 elctx_sat (E1 ++ E2).
Proof.
iIntros (HE1 HE2 ??) "HL".
iDestruct (HE1 with "HL") as "#HE1".
iDestruct (HE2 with "HL") as "#HE2".
iClear "∗". iIntros "!> #HE".
iDestruct ("HE1" with "HE") as "#$".
iApply ("HE2" with "HE").
Qed.
Lemma elctx_sat_refl : elctx_sat E.
Proof. iIntros (??) "_ !> ?". done. Qed.
End lft_contexts.
Global Arguments lctx_lft_incl {_ _} _ _ _ _.
Global Arguments lctx_lft_eq {_ _} _ _ _ _.
Global Arguments lctx_lft_alive {_ _ _} _ _ _.
Global Arguments elctx_sat {_ _} _ _ _.
Global Arguments lctx_lft_incl_incl {_ _ _ _ _} _ _.
Global Arguments lctx_lft_incl_incl_noend {_ _ _ _} _ _.
Global Arguments lctx_lft_alive_tok {_ _ _ _ _} _ _ _.
Global Arguments lctx_lft_alive_tok_noend {_ _ _ _ _} _ _ _.
Lemma elctx_sat_submseteq `{!invGS Σ, !lftGS Σ lft_userE} E E' L :
E' ⊆+ E elctx_sat E L E'.
Proof. iIntros (HE' ??) "_ !> H". by iApply big_sepL_submseteq. Qed.
Global Hint Resolve
lctx_lft_incl_refl lctx_lft_incl_static lctx_lft_incl_local'
lctx_lft_incl_external' lctx_lft_incl_intersect
lctx_lft_incl_intersect_l lctx_lft_incl_intersect_r
lctx_lft_alive_static lctx_lft_alive_local lctx_lft_alive_external
elctx_sat_nil elctx_sat_lft_incl elctx_sat_app elctx_sat_refl
: lrust_typing.
Global Hint Extern 10 (lctx_lft_eq _ _ _ _) => split : lrust_typing.
Global Hint Resolve elctx_sat_submseteq | 100 : lrust_typing.
Global Hint Opaque elctx_sat lctx_lft_alive lctx_lft_incl : 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 arc.
From lrust.lifetime Require Import at_borrow.
From lrust.typing Require Export type.
From lrust.typing Require Import typing option.
From iris.prelude Require Import options.
Definition arcN := lrustN .@ "arc".
Definition arc_invN := arcN .@ "inv".
Definition arc_shrN := arcN .@ "shr".
Definition arc_endN := arcN .@ "end".
Section arc.
Context `{!typeGS Σ, !arcG Σ}.
(* Preliminary definitions. *)
Let P1 ν q := (q.[ν])%I.
Local Instance P1_fractional ν : Fractional (P1 ν).
Proof. unfold P1. apply _. Qed.
Let P2 l n := ( l(2 + n) (l + 2) ↦∗: λ vl, length vl = n)%I.
Definition arc_persist tid ν (γ : gname) (l : loc) (ty : type) : iProp Σ :=
(is_arc (P1 ν) (P2 l ty.(ty_size)) arc_invN γ l
(* We use this disjunction, and not simply [ty_shr] here, *)
(* because [weak_new] cannot prove ty_shr, even for a dead *)
(* lifetime. *)
(ty.(ty_shr) ν tid (l + 2) [ν])
(1.[ν] ={lftN lft_userE arc_endN}[lft_userE]▷=∗
[ν] (l + 2) ↦∗: ty.(ty_own) tid l(2 + ty.(ty_size))))%I.
Global Instance arc_persist_ne tid ν γ l n :
Proper (type_dist2 n ==> dist n) (arc_persist tid ν γ l).
Proof.
unfold arc_persist, P1, P2.
solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Lemma arc_persist_type_incl tid ν γ l ty1 ty2:
type_incl ty1 ty2 -∗ arc_persist tid ν γ l ty1 -∗ arc_persist tid ν γ l ty2.
Proof.
iIntros "#(Heqsz & Hinclo & Hincls) #(?& Hs & Hvs)".
iDestruct "Heqsz" as %->. iFrame "#". iSplit.
- iDestruct "Hs" as "[?|?]"; last auto. iLeft. by iApply "Hincls".
- iIntros "!> Hν". iMod ("Hvs" with "Hν") as "H". iModIntro. iNext.
iMod "H" as "($ & H & $)". iDestruct "H" as (vl) "[??]". iExists _.
iFrame. by iApply "Hinclo".
Qed.
Lemma arc_persist_send tid tid' ν γ l ty `{!Send ty,!Sync ty} :
arc_persist tid ν γ l ty arc_persist tid' ν γ l ty.
Proof.
iIntros "#($ & ? & Hvs)". rewrite sync_change_tid. iFrame "#".
iIntros "!> Hν". iMod ("Hvs" with "Hν") as "H". iModIntro. iNext.
iMod "H" as "($ & H & $)". iDestruct "H" as (vl) "?". iExists _.
by rewrite send_change_tid.
Qed.
(* Model of arc *)
(* The ty_own part of the arc type cointains either the
shared protocol or the full ownership of the
content. The reason is that get_mut does not have the
masks to rebuild the invariants. *)
Definition full_arc_own l ty tid : iProp Σ:=
(l #1 (l + 1) #1 l(2 + ty.(ty_size))
(l + 2) ↦∗: ty.(ty_own) tid)%I.
Definition shared_arc_own l ty tid : iProp Σ:=
( γ ν q, arc_persist tid ν γ l ty arc_tok γ q q.[ν])%I.
Lemma arc_own_share E l ty tid :
lftN E
lft_ctx -∗ full_arc_own l ty tid ={E}=∗ shared_arc_own l ty tid.
Proof.
iIntros (?) "#LFT (Hl1 & Hl2 & H† & Hown)".
iMod (lft_create with "LFT") as (ν) "[Hν #Hν†]"=>//.
(* TODO: We should consider changing the statement of
bor_create to dis-entangle the two masks such that this is no
longer necessary. *)
iApply fupd_trans. iApply (fupd_mask_mono (lftN))=>//.
iMod (bor_create _ ν with "LFT Hown") as "[HP HPend]"=>//. iModIntro.
iMod (ty_share with "LFT HP Hν") as "[#? Hν]"; first solve_ndisj.
iMod (create_arc (P1 ν) (P2 l ty.(ty_size)) arc_invN with "Hl1 Hl2 Hν")
as (γ q') "(#? & ? & ?)".
iExists _, _, _. iFrame "∗#". iCombine "H†" "HPend" as "H".
iMod (inv_alloc arc_endN _ ([ν] _)%I with "[H]") as "#INV";
first by iRight; iApply "H". iIntros "!> !> Hν".
iMod (inv_acc with "INV") as "[[H†|[$ Hvs]] Hclose]";
[set_solver|iDestruct (lft_tok_dead with "Hν H†") as ">[]"|].
rewrite difference_union_distr_l_L difference_diag_L (right_id_L )
difference_disjoint_L; last first.
{ apply disjoint_union_l; solve_ndisj. }
iMod ("Hν†" with "Hν") as "H". iModIntro. iNext. iApply fupd_trans.
iMod "H" as "#Hν".
iMod fupd_mask_subseteq as "Hclose2"; last iMod ("Hvs" with "Hν") as "$".
{ set_solver+. }
iIntros "{$Hν} !>".
iMod "Hclose2" as "_". iApply "Hclose". auto.
Qed.
Program Definition arc (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] => full_arc_own l ty tid shared_arc_own l ty tid
| _ => False end;
ty_shr κ tid l :=
(l' : loc), &frac{κ} (λ q, l {q} #l')
F q, ⌜↑shrN lftN F -∗ q.[κ]
={F}[F∖↑shrN]▷=∗ q.[κ] γ ν q', κ ν
arc_persist tid ν γ l' ty
&at{κ, arc_shrN}(arc_tok γ q')
|}%I.
Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". 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.
(* Ideally, we'd change ty_shr to say "l ↦{q} #l" in the fractured borrow,
but that would be additional work here... *)
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 "[>[] _]".
setoid_rewrite heap_pointsto_vec_singleton.
iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx.
iDestruct "Hb" as (i) "(#Hpb&Hpbown)".
set (C := ( _ _ _, _ _ &at{_,_} _)%I).
iMod (inv_alloc shrN _ (idx_bor_own 1 i C)%I
with "[Hpbown]") as "#Hinv"; first by iLeft.
iIntros "!> !> * % Htok".
iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iAssert (&{κ} _)%I with "[Hbtok]" as "Hb".
{ rewrite bor_unfold_idx. iExists _. by iFrame. }
iClear "H↦ Hinv Hpb".
iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose2]"; first solve_ndisj.
iModIntro. iNext. iAssert (shared_arc_own l' ty tid)%I with "[>HP]" as "HX".
{ iDestruct "HP" as "[?|$]"; last done. iApply arc_own_share; solve_ndisj. }
iDestruct "HX" as (γ ν q') "[#Hpersist H]".
iMod ("Hclose2" with "[] H") as "[HX $]"; first by unfold shared_arc_own; auto 10.
iAssert C with "[>HX]" as "#$".
{ iExists _, _, _. iFrame "Hpersist".
iMod (bor_sep with "LFT HX") as "[Hrc Hlft]"; first solve_ndisj.
iDestruct (frac_bor_lft_incl with "LFT [> Hlft]") as "$".
{ iApply (bor_fracture with "LFT"); first solve_ndisj. by rewrite Qp.mul_1_r. }
iApply (bor_share with "Hrc"); solve_ndisj. }
iApply ("Hclose1" with "[]"). by auto.
- iMod ("Hclose1" with "[]") as "_"; first by auto.
iApply step_fupd_intro; first solve_ndisj. by iFrame.
Qed.
Next Obligation.
iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]".
iExists _. iSplit; first by iApply frac_bor_shorten.
iModIntro. iIntros (F q) "% Htok".
iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hshr" with "[] Htok") as "Hshr2"; first done.
iModIntro. iNext. iMod "Hshr2" as "[Htok HX]".
iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν ?) "(? & ? & ?)".
iExists _, _, _. iModIntro. iFrame. iSplit.
- by iApply lft_incl_trans.
- by iApply at_bor_shorten.
Qed.
Global Instance arc_wf ty `{!TyWf ty} : TyWf (arc ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance arc_type_contractive : TypeContractive arc.
Proof.
constructor=>/=; unfold arc, full_arc_own, shared_arc_own;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance arc_ne : NonExpansive arc.
Proof. apply type_contractive_ne, _. Qed.
Lemma arc_subtype ty1 ty2 :
type_incl ty1 ty2 type_incl (arc ty1) (arc ty2).
Proof.
iIntros "#Hincl". iPoseProof "Hincl" as "(#Hsz & #Hoincl & #Hsincl)".
iSplit; first done. iSplit; iModIntro.
- iIntros "%tid %vl Hvl". destruct vl as [|[[|vl|]|] [|]]; try done.
iDestruct "Hvl" as "[(Hl1 & Hl2 & H† & Hc) | Hvl]".
{ iLeft. iFrame. iDestruct "Hsz" as %->.
iFrame. iApply (heap_pointsto_pred_wand with "Hc"). iApply "Hoincl". }
iDestruct "Hvl" as (γ ν q) "(#Hpersist & Htk & Hν)".
iRight. iExists _, _, _. iFrame "#∗". by iApply arc_persist_type_incl.
- iIntros "%κ %tid %l #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'.
iIntros "{$Hfrac} !> * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H".
iModIntro. iNext. iMod "H" as "[$ H]".
iDestruct "H" as (γ ν q') "(Hlft & Hpersist & Hna)".
iExists _, _, _. iFrame. by iApply arc_persist_type_incl.
Qed.
Global Instance arc_mono E L :
Proper (subtype E L ==> subtype E L) arc.
Proof.
iIntros (ty1 ty2 Hsub ??) "HL". iDestruct (Hsub with "HL") as "#Hsub".
iIntros "!> #HE". iApply arc_subtype. by iApply "Hsub".
Qed.
Global Instance arc_proper E L :
Proper (eqtype E L ==> eqtype E L) arc.
Proof. intros ??[]. by split; apply arc_mono. Qed.
Global Instance arc_send ty :
Send ty Sync ty Send (arc ty).
Proof.
intros Hse Hsy tid tid' vl. destruct vl as [|[[|l|]|] []]=>//=.
unfold full_arc_own, shared_arc_own.
repeat (apply send_change_tid || apply bi.exist_mono ||
(apply arc_persist_send; apply _) || f_equiv || intros ?).
Qed.
Global Instance arc_sync ty :
Send ty Sync ty Sync (arc ty).
Proof.
intros Hse Hsy ν tid tid' l.
repeat (apply send_change_tid || apply bi.exist_mono ||
(apply arc_persist_send; apply _) || f_equiv || intros ?).
Qed.
(* Model of weak *)
Program Definition weak (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] => γ ν, arc_persist tid ν γ l ty weak_tok γ
| _ => False end;
ty_shr κ tid l :=
(l' : loc), &frac{κ} (λ q, l {q} #l')
F q, ⌜↑shrN lftN F -∗ q.[κ]
={F}[F∖↑shrN]▷=∗ q.[κ] γ ν,
arc_persist tid ν γ l' ty &at{κ, arc_shrN}(weak_tok γ)
|}%I.
Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". 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.
(* Ideally, we'd change ty_shr to say "l ↦{q} #l" in the fractured borrow,
but that would be additional work here... *)
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 "[>[] _]".
setoid_rewrite heap_pointsto_vec_singleton.
iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx.
iDestruct "Hb" as (i) "(#Hpb&Hpbown)".
set (C := ( _ _, _ &at{_,_} _)%I).
iMod (inv_alloc shrN _ (idx_bor_own 1 i C)%I
with "[Hpbown]") as "#Hinv"; first by iLeft.
iIntros "!> !> * % Htok".
iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iAssert (&{κ} _)%I with "[Hbtok]" as "Hb".
{ rewrite bor_unfold_idx. iExists _. by iFrame. }
iClear "H↦ Hinv Hpb".
iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose2]"; first solve_ndisj.
iDestruct "HP" as (γ ν) "[#Hpersist H]".
iMod ("Hclose2" with "[] H") as "[HX $]"; first by auto 10.
iModIntro. iNext. iAssert C with "[>HX]" as "#$".
{ iExists _, _. iFrame "Hpersist". iApply bor_share; solve_ndisj. }
iApply ("Hclose1" with "[]"). by auto.
- iMod ("Hclose1" with "[]") as "_"; first by auto.
iApply step_fupd_intro; first solve_ndisj. by iFrame.
Qed.
Next Obligation.
iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]".
iExists _. iSplit; first by iApply frac_bor_shorten.
iModIntro. iIntros (F q) "% Htok".
iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hshr" with "[] Htok") as "Hshr2"; first done.
iModIntro. iNext. iMod "Hshr2" as "[Htok HX]".
iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν) "[??]".
iExists _, _. iModIntro. iFrame. by iApply at_bor_shorten.
Qed.
Global Instance weak_wf ty `{!TyWf ty} : TyWf (weak ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance weak_type_contractive : TypeContractive weak.
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance weak_ne : NonExpansive weak.
Proof. apply type_contractive_ne, _. Qed.
Lemma weak_subtype ty1 ty2 :
type_incl ty1 ty2 -∗ type_incl (weak ty1) (weak ty2).
Proof.
iIntros "#Hincl". iPoseProof "Hincl" as "(#Hsz & #Hoincl & #Hsincl)".
iSplit; first done. iSplit; iModIntro.
- iIntros "%tid %vl Hvl". destruct vl as [|[[|vl|]|] [|]]; try done.
iDestruct "Hvl" as (γ ν) "(#Hpersist & Htk)".
iExists _, _. iFrame "#∗". by iApply arc_persist_type_incl.
- iIntros "%κ %tid %l #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'.
iIntros "{$Hfrac} !> * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H".
iModIntro. iNext. iMod "H" as "[$ H]". iDestruct "H" as (γ ν) "[Hpersist Hna]".
iExists _, _. iFrame. by iApply arc_persist_type_incl.
Qed.
Global Instance weak_mono E L :
Proper (subtype E L ==> subtype E L) weak.
Proof.
iIntros (ty1 ty2 Hsub ??) "HL". iDestruct (Hsub with "HL") as "#Hsub".
iIntros "!> #HE". iApply weak_subtype. by iApply "Hsub".
Qed.
Global Instance weak_proper E L :
Proper (eqtype E L ==> eqtype E L) weak.
Proof. intros ??[]. by split; apply weak_mono. Qed.
Global Instance weak_send ty :
Send ty Sync ty Send (weak ty).
Proof.
intros Hse Hsy tid tid' vl. destruct vl as [|[[|l|]|] []]=>//=.
repeat (apply send_change_tid || apply bi.exist_mono ||
(apply arc_persist_send; apply _) || f_equiv || intros ?).
Qed.
Global Instance weak_sync ty :
Send ty Sync ty Sync (weak ty).
Proof.
intros Hse Hsy ν tid tid' l.
repeat (apply send_change_tid || apply bi.exist_mono ||
(apply arc_persist_send; apply _) || f_equiv || intros ?).
Qed.
(* Code : constructors *)
Definition arc_new ty : val :=
fn: ["x"] :=
let: "arcbox" := new [ #(2 + ty.(ty_size))%nat ] in
let: "arc" := new [ #1 ] in
"arcbox" + #0 <- #1;;
"arcbox" + #1 <- #1;;
"arcbox" + #2 <-{ty.(ty_size)} !"x";;
"arc" <- "arcbox";;
delete [ #ty.(ty_size); "x"];; return: ["arc"].
Lemma arc_new_type ty `{!TyWf ty} :
typed_val (arc_new ty) (fn(; ty) arc ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (rc); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]".
rewrite !tctx_hasty_val.
iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x.
iDestruct (ownptr_uninit_own with "Hrcbox")
as (lrcbox vlrcbox) "(% & Hrcbox↦ & Hrcbox†)". subst rcbox. inv_vec vlrcbox=>???.
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]".
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]".
rewrite shift_loc_assoc.
iDestruct (ownptr_uninit_own with "Hrc") as (lrc vlrc) "(% & Hrc↦ & Hrc†)". subst rc.
inv_vec vlrc=>?. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write.
wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz.
wp_apply (wp_memcpy with "[$Hrcbox↦2 $Hx↦]"); [by auto using length_vec_to_list..|].
iIntros "[Hrcbox↦2 Hx↦]". wp_seq. wp_write.
(* Finish up the proof. *)
iApply (type_type _ _ _ [ #lx box (uninit (ty_size ty)); #lrc box (arc 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_singleton. iFrame. iLeft.
rewrite freeable_sz_full_S. by iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition weak_new ty : val :=
fn: [] :=
let: "arcbox" := new [ #(2 + ty.(ty_size))%nat ] in
let: "w" := new [ #1 ] in
"arcbox" + #0 <- #0;;
"arcbox" + #1 <- #1;;
"w" <- "arcbox";;
return: ["w"].
Lemma weak_new_type ty `{!TyWf ty} :
typed_val (weak_new ty) (fn() weak ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg. simpl_subst.
iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (rc); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrc [Hrcbox _]]". rewrite !tctx_hasty_val.
iDestruct (ownptr_uninit_own with "Hrcbox")
as (lrcbox vlrcbox) "(% & Hrcbox↦ & Hrcbox†)". subst rcbox. inv_vec vlrcbox=>??? /=.
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]".
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]".
rewrite shift_loc_assoc.
iDestruct (ownptr_uninit_own with "Hrc") as (lrc vlrc) "(% & Hrc↦ & Hrc†)". subst rc.
inv_vec vlrc=>?. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
iMod (lft_create with "LFT") as (ν) "[Hν Hν†]"; first done.
wp_bind (_ + _)%E. iSpecialize ("Hν†" with "Hν").
iApply wp_mask_mono; last iApply (wp_step_fupd with "Hν†"); [set_solver+..|].
wp_op. iIntros "#H† !>". rewrite shift_loc_0. wp_write. wp_op. wp_write. wp_write.
iApply (type_type _ _ _ [ #lrc box (weak ty)]
with "[] LFT HE Hna HL Hk [>-]"); last first.
{ rewrite tctx_interp_singleton tctx_hasty_val' //=. iFrame.
iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iMod (create_weak (P1 ν) (P2 lrcbox ty.(ty_size))
with "Hrcbox↦0 Hrcbox↦1 [-]") as (γ) "[Ha Htok]".
{ rewrite freeable_sz_full_S. iFrame.
by rewrite length_vec_to_list. }
iExists γ, ν. iFrame. iSplitL; first by auto. iIntros "!>!>!> Hν".
iDestruct (lft_tok_dead with "Hν H†") as "[]". }
iApply type_jump; solve_typing.
Qed.
(* Code : deref *)
Definition arc_deref : val :=
fn: ["arc"] :=
let: "x" := new [ #1 ] in
let: "arc'" := !"arc" in
"x" <- (!"arc'" + #2);;
delete [ #1; "arc" ];; return: ["x"].
Lemma arc_deref_type ty `{!TyWf ty} :
typed_val arc_deref (fn( α, ; &shr{α}(arc ty)) &shr{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (x); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [#Hrc' [Hx _]]]".
rewrite !tctx_hasty_val [[rcx]]lock.
iDestruct (ownptr_uninit_own with "Hx") as (lrc2 vlrc2) "(% & Hx & Hx†)".
subst x. inv_vec vlrc2=>?. rewrite heap_pointsto_vec_singleton.
destruct rc' as [[|rc'|]|]; try done. simpl of_val.
iDestruct "Hrc'" as (l') "[Hrc' #Hdelay]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hdelay" with "[] Hα1"); last iApply (wp_step_fupd with "Hdelay"); [done..|].
iMod (frac_bor_acc with "LFT Hrc' Hα2") as (q') "[Hrc'↦ Hclose2]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose2" with "[$Hrc'↦]") as "Hα2". iIntros "!> [Hα1 #Hproto] !>".
iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & _)".
iDestruct "Hpersist" as "(_ & [Hshr|Hν†] & _)"; last first.
{ iMod (lft_incl_dead with "Hαν Hν†") as "Hα†"; first done.
iDestruct (lft_tok_dead with "Hα1 Hα†") as "[]". }
wp_op. wp_write. iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ rcx box (&shr{α}(arc ty)); #lrc2 box (&shr{α}ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame "Hrcx". iFrame "Hx†". iExists [_]. rewrite heap_pointsto_vec_singleton.
iFrame "Hx". by iApply ty_shr_mono. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Code : getters *)
Definition arc_strong_count : val :=
fn: ["arc"] :=
let: "r" := new [ #1 ] in
let: "arc'" := !"arc" in
let: "arc''" := !"arc'" in
"r" <- strong_count ["arc''"];;
delete [ #1; "arc" ];; return: ["r"].
Lemma arc_strong_count_type ty `{!TyWf ty} :
typed_val arc_strong_count (fn( α, ; &shr{α}(arc ty)) int).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|]. wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let.
wp_apply (strong_count_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I
with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]".
{ iIntros "!> Hα".
iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|].
iExists _. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro; first set_solver.
iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". }
iIntros (c) "[Hα _]". iMod ("Hclose1" with "Hα HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(arc ty)); #c int; r box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition arc_weak_count : val :=
fn: ["arc"] :=
let: "r" := new [ #1 ] in
let: "arc'" := !"arc" in
let: "arc''" := !"arc'" in
"r" <- weak_count ["arc''"];;
delete [ #1; "arc" ];; return: ["r"].
Lemma arc_weak_count_type ty `{!TyWf ty} :
typed_val arc_weak_count (fn( α, ; &shr{α}(arc ty)) int).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|]. wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let.
wp_apply (weak_count_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I
with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]".
{ iIntros "!> Hα".
iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|].
iExists _. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro; first set_solver.
iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". }
iIntros (c) "[Hα _]". iMod ("Hclose1" with "Hα HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(arc ty)); #c int; r box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Code : clone, [up/down]grade. *)
Definition arc_clone : val :=
fn: ["arc"] :=
let: "r" := new [ #1 ] in
let: "arc'" := !"arc" in
let: "arc''" := !"arc'" in
clone_arc ["arc''"];;
"r" <- "arc''";;
delete [ #1; "arc" ];; return: ["r"].
Lemma arc_clone_type ty `{!TyWf ty} :
typed_val arc_clone (fn( α, ; &shr{α}(arc ty)) arc ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|]. wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let.
wp_apply (clone_arc_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I
with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]".
{ iIntros "!> Hα".
iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|].
iExists _. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro; first set_solver.
iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". }
iIntros (q'') "[Hα Hown]". wp_seq. iMod ("Hclose1" with "Hα HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(arc ty)); #l' arc ty; r box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. simpl. unfold shared_arc_own. auto 10 with iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition weak_clone : val :=
fn: ["w"] :=
let: "r" := new [ #1 ] in
let: "w'" := !"w" in
let: "w''" := !"w'" in
clone_weak ["w''"];;
"r" <- "w''";;
delete [ #1; "w" ];; return: ["r"].
Lemma weak_clone_type ty `{!TyWf ty} :
typed_val weak_clone (fn( α, ; &shr{α}(weak ty)) weak ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|]. wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν) "#[Hpersist Hrctokb]". iModIntro. wp_let.
wp_apply (clone_weak_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I
with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]".
{ iIntros "!> Hα".
iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>$ Hclose1]"; [solve_ndisj..|].
iMod fupd_mask_subseteq as "Hclose2"; last iModIntro; first set_solver.
iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". }
iIntros "[Hα Hown]". wp_seq. iMod ("Hclose1" with "Hα HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(weak ty)); #l' weak ty; r box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. simpl. eauto 10 with iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition downgrade : val :=
fn: ["arc"] :=
let: "r" := new [ #1 ] in
let: "arc'" := !"arc" in
let: "arc''" := !"arc'" in
downgrade ["arc''"];;
"r" <- "arc''";;
delete [ #1; "arc" ];; return: ["r"].
Lemma downgrade_type ty `{!TyWf ty} :
typed_val downgrade (fn( α, ; &shr{α}(arc ty)) weak ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|]. wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let.
wp_apply (downgrade_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I
with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]".
{ iIntros "!> Hα".
iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|].
iExists _. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro; first set_solver.
iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". }
iIntros "[Hα Hown]". wp_seq. iMod ("Hclose1" with "Hα HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(arc ty)); #l' weak ty; r box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. simpl. eauto 10 with iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition upgrade : val :=
fn: ["w"] :=
let: "r" := new [ #2 ] in
let: "w'" := !"w" in
let: "w''" := !"w'" in
if: upgrade ["w''"] then
"r" <-{Σ some} "w''";;
delete [ #1; "w" ];; return: ["r"]
else
"r" <-{Σ none} ();;
delete [ #1; "w" ];; return: ["r"].
Lemma upgrade_type ty `{!TyWf ty} :
typed_val upgrade (fn( α, ; &shr{α}(weak ty)) option (arc ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|]. wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν) "#[Hpersist Hrctokb]". iModIntro. wp_let.
wp_apply (upgrade_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I
with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]".
{ iIntros "!> Hα".
iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>$ Hclose1]"; [solve_ndisj..|].
iMod fupd_mask_subseteq as "Hclose2"; last iModIntro; first set_solver.
iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". }
iIntros ([] q') "[Hα Hown]"; wp_if; iMod ("Hclose1" with "Hα HL") as "HL".
- (* Finish up the proof (sucess). *)
iApply (type_type _ _ _ [ x box (&shr{α}(weak ty)); #l' arc ty;
r box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. simpl. unfold shared_arc_own. eauto 10 with iFrame. }
iApply (type_sum_assign (option (arc ty))); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
- (* Finish up the proof (fail). *)
iApply (type_type _ _ _ [ x box (&shr{α}(weak ty)); r box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
unlock. iFrame. }
iApply (type_sum_unit (option (arc ty))); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Code : drop *)
Definition arc_drop ty : val :=
fn: ["arc"] :=
let: "r" := new [ #(option ty).(ty_size) ] in
let: "arc'" := !"arc" in
"r" <-{Σ none} ();;
(if: drop_arc ["arc'"] then
"r" <-{ty.(ty_size),Σ some} !("arc'" + #2);;
if: drop_weak ["arc'"] then
delete [ #(2 + ty.(ty_size)); "arc'" ]
else #☠
else #☠);;
delete [ #1; "arc"];; return: ["r"].
Lemma arc_drop_type ty `{!TyWf ty} :
typed_val (arc_drop ty) (fn(; arc ty) option ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new (option ty).(ty_size)); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iApply (type_sum_unit (option ty)); [solve_typing..|].
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hr [Hrcx [Hrc' _]]]".
rewrite !tctx_hasty_val. destruct rc' as [[|rc'|]|]=>//=.
iAssert (shared_arc_own rc' ty tid)%I with "[>Hrc']" as "Hrc'".
{ iDestruct "Hrc'" as "[?|$]"; last done. iApply arc_own_share; solve_ndisj. }
iDestruct "Hrc'" as (γ ν q) "(#Hpersist & Htok & Hν)".
wp_bind (drop_arc _). iApply (drop_arc_spec with "[] [$Htok Hν]");
[by iDestruct "Hpersist" as "[$?]"|done|].
iNext. iIntros (b) "Hdrop". wp_bind (if: _ then _ else _)%E.
iApply (wp_wand _ _ _ (λ _, ty_own (box (option ty)) tid [r])%I with "[Hdrop Hr]").
{ destruct b; wp_if=>//. destruct r as [[]|]; try done.
(* FIXME: 'simpl' reveals a match here. Didn't we forbid that for ty_own? *)
rewrite {3}[option]lock. simpl. rewrite -{2 3}lock. (* FIXME: Tried using contextual pattern, did not work. *)
iDestruct "Hr" as "[Hr ?]". iDestruct "Hr" as ([|d vl]) "[H↦ Hown]";
first by iDestruct "Hown" as (???) "[>% ?]".
rewrite heap_pointsto_vec_cons. iDestruct "H↦" as "[H↦0 H↦1]".
iDestruct "Hpersist" as "(? & ? & H†)". wp_bind (_ <- _)%E.
iDestruct "Hdrop" as "[Hν Hdrop]". iSpecialize ("H†" with "Hν").
iApply wp_mask_mono; last iApply (wp_step_fupd with "H†"); [solve_ndisj..|].
iDestruct "Hown" as (???) "(_ & Hlen & _)". wp_write. iIntros "(#Hν & Hown & H†)!>".
wp_seq. wp_op. wp_op. iDestruct "Hown" as (?) "[H↦ Hown]".
iDestruct (ty_size_eq with "Hown") as %?. rewrite Nat.max_0_r.
iDestruct "Hlen" as %[= ?]. wp_apply (wp_memcpy with "[$H↦1 $H↦]"); [congruence..|].
iIntros "[? Hrc']". wp_seq. iMod ("Hdrop" with "[Hrc' H†]") as "Htok".
{ unfold P2. auto with iFrame. }
wp_apply (drop_weak_spec with "[//] Htok"). unlock. iIntros ([]); last first.
{ iIntros "_". wp_if. unlock. iFrame. iExists (_::_). rewrite heap_pointsto_vec_cons.
iFrame. iExists 1%nat, _, []. rewrite /= right_id_L Nat.max_0_r.
auto 10 with iFrame. }
iIntros "([H† H1] & H2 & H3)". iDestruct "H1" as (vl1) "[H1 Heq]".
iDestruct "Heq" as %<-. wp_if.
wp_apply (wp_delete _ _ _ (_::_::vl1) with "[H1 H2 H3 H†]").
{ simpl. lia. }
{ rewrite 2!heap_pointsto_vec_cons shift_loc_assoc. auto with iFrame. }
iFrame. iIntros "_". iExists (_::_). rewrite heap_pointsto_vec_cons. iFrame.
iExists 1%nat, _, []. rewrite right_id_L. iFrame. iSplit; [by auto|simpl].
auto with lia. }
iIntros (?) "Hr". wp_seq.
(* Finish up the proof. *)
iApply (type_type _ _ _ [ rcx box (uninit 1); r box (option ty) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. unlock.
rewrite tctx_hasty_val. iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition weak_drop ty : val :=
fn: ["arc"] :=
let: "r" := new [ #0] in
let: "arc'" := !"arc" in
(if: drop_weak ["arc'"] then delete [ #(2 + ty.(ty_size)); "arc'" ]
else #☠);;
delete [ #1; "arc"];; return: ["r"].
Lemma weak_drop_type ty `{!TyWf ty} :
typed_val (weak_drop ty) (fn(; weak ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 0); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. destruct rc' as [[|rc'|]|]=>//=.
iDestruct "Hrc'" as (γ ν) "[#Hpersist Htok]". wp_bind (drop_weak _).
iApply (drop_weak_spec with "[] [Htok]"); [by iDestruct "Hpersist" as "[$?]"|by auto|].
iIntros "!> %b Hdrop". wp_bind (if: _ then _ else _)%E.
iApply (wp_wand _ _ _ (λ _, True)%I with "[Hdrop]").
{ destruct b; wp_if=>//. iDestruct "Hdrop" as "((? & H↦) & ? & ?)".
iDestruct "H↦" as (vl) "[? Heq]". iDestruct "Heq" as %<-.
wp_apply (wp_delete _ _ _ (_::_::vl) with "[-]"); [simpl;lia| |done].
rewrite !heap_pointsto_vec_cons shift_loc_assoc. iFrame. }
iIntros (?) "_". wp_seq.
(* Finish up the proof. *)
iApply (type_type _ _ _ [ rcx box (uninit 1); r box (uninit 0) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. unlock.
rewrite tctx_hasty_val. iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Code : other primitives *)
Definition arc_try_unwrap ty : val :=
fn: ["arc"] :=
let: "r" := new [ #(ty_size Σ[ ty; arc ty ]) ] in
let: "arc'" := !"arc" in
if: try_unwrap ["arc'"] then
(* Return content *)
"r" <-{ty.(ty_size),Σ 0} !("arc'" + #2);;
(* Decrement the "strong weak reference" *)
(if: drop_weak ["arc'"] then delete [ #(2 + ty.(ty_size)); "arc'" ]
else #☠);;
delete [ #1; "arc"];; return: ["r"]
else
"r" <-{Σ 1} "arc'";;
delete [ #1; "arc"];; return: ["r"].
Lemma arc_try_unwrap_type ty `{!TyWf ty} :
typed_val (arc_try_unwrap ty) (fn(; arc ty) Σ[ ty; arc ty ]).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new (Σ[ ty; arc ty ]).(ty_size));
[solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. destruct rc' as [[|rc'|]|]=>//=.
iAssert (shared_arc_own rc' ty tid)%I with "[>Hrc']" as "Hrc'".
{ iDestruct "Hrc'" as "[?|$]"; last done. iApply arc_own_share; solve_ndisj. }
iDestruct "Hrc'" as (γ ν q) "(#Hpersist & Htok & Hν)".
wp_apply (try_unwrap_spec with "[] [Hν Htok]");
[by iDestruct "Hpersist" as "[$?]"|iFrame|].
iIntros ([]) "H"; wp_if.
- iDestruct "H" as "[Hν Hend]". rewrite -(lock [r]). destruct r as [[|r|]|]=>//=.
iDestruct "Hr" as "[Hr >Hfree]". iDestruct "Hr" as (vl0) "[>Hr Hown]".
iDestruct "Hown" as ">Hlen".
destruct vl0 as [|? vl0]=>//; iDestruct "Hlen" as %[=Hlen0].
rewrite heap_pointsto_vec_cons. iDestruct "Hr" as "[Hr0 Hr1]".
iDestruct "Hpersist" as "(Ha & _ & H†)". wp_bind (_ <- _)%E.
iSpecialize ("H†" with "Hν").
iApply wp_mask_mono; last iApply (wp_step_fupd with "H†"); [solve_ndisj..|].
wp_write. iIntros "(#Hν & Hown & H†) !>". wp_seq. wp_op. wp_op.
rewrite -(firstn_skipn ty.(ty_size) vl0) heap_pointsto_vec_app.
iDestruct "Hr1" as "[Hr1 Hr2]". iDestruct "Hown" as (vl) "[Hrc' Hown]".
iDestruct (ty_size_eq with "Hown") as %Hlen.
wp_apply (wp_memcpy with "[$Hr1 $Hrc']"); rewrite /= ?firstn_length_le; try lia.
iIntros "[Hr1 Hrc']". wp_seq. wp_bind (drop_weak _).
iMod ("Hend" with "[$H† Hrc']") as "Htok"; first by eauto.
iApply (drop_weak_spec with "Ha Htok").
iIntros "!> %b Hdrop". wp_bind (if: _ then _ else _)%E.
iApply (wp_wand _ _ _ (λ _, True)%I with "[Hdrop]").
{ destruct b; wp_if=>//. iDestruct "Hdrop" as "((? & H↦) & ? & ?)".
iDestruct "H↦" as (vl') "[? Heq]". iDestruct "Heq" as %<-.
wp_apply (wp_delete _ _ _ (_::_::vl') with "[-]"); [simpl; lia| |done].
rewrite !heap_pointsto_vec_cons shift_loc_assoc. iFrame. }
iIntros (?) "_". wp_seq.
iApply (type_type _ _ _ [ rcx box (uninit 1); #r box (Σ[ ty; arc ty ]) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. unlock.
iFrame. iCombine "Hr1" "Hr2" as "Hr1". iCombine "Hr0" "Hr1" as "Hr".
rewrite -[in _ + ty.(ty_size)]Hlen -heap_pointsto_vec_app
-heap_pointsto_vec_cons tctx_hasty_val' //. iFrame.
iExists O, _, _. iSplit; first by auto. iFrame. iIntros "!> !% /=".
rewrite length_app length_drop. lia. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
- iApply (type_type _ _ _ [ rcx box (uninit 1); #rc' arc ty;
r box (uninit (Σ[ ty; arc ty ]).(ty_size)) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iRight. iExists _, _, _. auto with iFrame. }
iApply (type_sum_assign Σ[ ty; arc ty ]); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition arc_get_mut : val :=
fn: ["arc"] :=
let: "r" := new [ #2 ] in
let: "arc'" := !"arc" in
let: "arc''" := !"arc'" in
if: is_unique ["arc''"] then
Skip;;
(* Return content *)
"r" <-{Σ some} "arc''" + #2;;
delete [ #1; "arc"];; return: ["r"]
else
"r" <-{Σ none} ();;
delete [ #1; "arc"];; return: ["r"].
Lemma arc_get_mut_type ty `{!TyWf ty} :
typed_val arc_get_mut (fn( α, ; &uniq{α}(arc ty)) option (&uniq{α}ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. destruct rc' as [[|rc'|]|]=>//=.
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & HL & Hclose1)";
[solve_typing..|].
iMod (bor_exists with "LFT Hrc'") as (rcvl) "Hrc'"=>//.
iMod (bor_sep with "LFT Hrc'") as "[Hrc'↦ Hrc]"=>//.
destruct rcvl as [|[[|rc|]|][|]]; try by
iMod (bor_persistent with "LFT Hrc Hα") as "[>[] ?]".
rewrite heap_pointsto_vec_singleton.
iMod (bor_acc with "LFT Hrc'↦ Hα") as "[Hrc'↦ Hclose2]"=>//. wp_read.
iMod ("Hclose2" with "Hrc'↦") as "[_ Hα]".
iMod (bor_acc_cons with "LFT Hrc Hα") as "[Hrc Hclose2]"=>//. wp_let.
iAssert (shared_arc_own rc ty tid)%I with "[>Hrc]" as "Hrc".
{ iDestruct "Hrc" as "[Hrc|$]"=>//. by iApply arc_own_share. }
iDestruct "Hrc" as (γ ν q') "[#(Hi & Hs & #Hc) Htoks]".
wp_apply (is_unique_spec with "Hi Htoks"). iIntros ([]) "H"; wp_if.
- iDestruct "H" as "(Hrc & Hrc1 & Hν)". iSpecialize ("Hc" with "Hν"). wp_bind Skip.
iApply wp_mask_mono; last iApply (wp_step_fupd with "Hc"); [solve_ndisj..|].
wp_seq. iIntros "(#Hν & Hown & H†) !>". wp_seq.
iMod ("Hclose2" with "[Hrc Hrc1 H†] Hown") as "[Hb Hα]".
{ iIntros "!> Hown !>". iLeft. iFrame. }
iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ rcx box (uninit 1); #rc + #2 &uniq{α}ty;
r box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. unlock. iFrame. }
iApply (type_sum_assign (option (&uniq{α}ty))); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
- iMod ("Hclose2" with "[] [H]") as "[_ Hα]";
[by iIntros "!> H"; iRight; iApply "H"|iExists _, _, _; iFrame "∗#"|].
iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ rcx box (uninit 1); r box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. unlock. iFrame. }
iApply (type_sum_unit (option (&uniq{α}ty))); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition arc_make_mut (ty : type) (clone : val) : val :=
fn: ["arc"] :=
let: "r" := new [ #1 ] in
let: "arc'" := !"arc" in
let: "arc''" := !"arc'" in
(case: try_unwrap_full ["arc''"] of
[ "r" <- "arc''" + #2;;
delete [ #1; "arc"];; return: ["r"];
let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in
"rcbox" + #0 <- #1;;
"rcbox" + #1 <- #1;;
"rcbox" + #2 <-{ty.(ty_size)} !"arc''" + #2;;
"arc'" <- "rcbox";;
"r" <- "rcbox" + #2;;
delete [ #1; "arc"];; return: ["r"];
letalloc: "x" <- "arc''" + #2 in
letcall: "c" := clone ["x"]%E in (* FIXME : why do I need %E here ? *)
let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in
"rcbox" + #0 <- #1;;
"rcbox" + #1 <- #1;;
"rcbox" + #2 <-{ty.(ty_size)} !"c";;
delete [ #ty.(ty_size); "c"];;
"arc'" <- "rcbox";;
letalloc: "rcold" <- "arc''" in
(* FIXME : here, we are dropping the old arc pointer. In the
case another strong reference has been dropped while
cloning, it is possible that we are actually dropping the
last reference here. This means that we may have to drop an
instance of [ty] here. Instead, we simply forget it. *)
let: "drop" := arc_drop ty in
letcall: "content" := "drop" ["rcold"]%E in
delete [ #(option ty).(ty_size); "content"];;
"r" <- "rcbox" + #2;;
delete [ #1; "arc"];; return: ["r"]
]).
Lemma arc_make_mut_type ty `{!TyWf ty} clone :
typed_val clone (fn( α, ; &shr{α}ty) ty)
typed_val (arc_make_mut ty clone) (fn( α, ; &uniq{α}(arc ty)) &uniq{α} ty).
Proof.
intros Hclone E L. iApply type_fn; [solve_typing..|]. rewrite [(2 + ty_size ty)%nat]lock.
iIntros "/= !>". iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. destruct rc' as [[|rc'|]|]=>//=.
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|].
iMod (bor_acc_cons with "LFT Hrc' Hα1") as "[Hrc' Hclose2]"=>//.
iDestruct "Hrc'" as (rcvl) "[Hrc'↦ Hrc]".
destruct rcvl as [|[[|rc|]|][|]]; try by iDestruct "Hrc" as ">[]".
rewrite heap_pointsto_vec_singleton. wp_read.
iAssert (shared_arc_own rc ty tid)%I with "[>Hrc]" as "Hrc".
{ iDestruct "Hrc" as "[Hrc|$]"=>//. by iApply arc_own_share. }
iDestruct "Hrc" as (γ ν q') "[#(Hi & Hs & #Hc) Htoks]". wp_let.
wp_apply (try_unwrap_full_spec with "Hi Htoks"). iIntros (x).
pose proof (fin_to_nat_lt x). destruct (fin_to_nat x) as [|[|[]]]; last lia.
- iIntros "(Hrc0 & Hrc1 & HP1)". wp_case. wp_bind (_ + _)%E.
iSpecialize ("Hc" with "HP1").
iApply wp_mask_mono; last iApply (wp_step_fupd with "Hc"); [solve_ndisj..|].
wp_op. iIntros "(#Hν & Hrc2 & H†)". iModIntro.
iMod ("Hclose2" with "[Hrc'↦ Hrc0 Hrc1 H†] Hrc2") as "[Hb Hα1]".
{ iIntros "!> Hrc2". iExists [_]. rewrite heap_pointsto_vec_singleton.
iFrame. iLeft. by iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
iApply (type_type _ _ _ [ rcx box (uninit 1); #(rc + 2) &uniq{α}ty;
r box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
- iIntros "[Hν Hweak]". wp_case. iSpecialize ("Hc" with "Hν"). wp_bind (new _).
iApply wp_mask_mono; last iApply (wp_step_fupd with "Hc"); [solve_ndisj..|].
wp_apply wp_new=>//; first lia. iIntros (l) "(H† & Hlvl) (#Hν & Hown & H†') !>".
rewrite -!lock Nat2Z.id.
wp_let. wp_op. rewrite !heap_pointsto_vec_cons shift_loc_assoc shift_loc_0.
iDestruct "Hlvl" as "(Hrc0 & Hrc1 & Hrc2)". wp_write. wp_op. wp_write.
wp_op. wp_op. iDestruct "Hown" as (vl) "[H↦ Hown]".
iDestruct (ty_size_eq with "Hown") as %Hsz.
wp_apply (wp_memcpy with "[$Hrc2 $H↦]").
{ by rewrite repeat_length. } { by rewrite Hsz. }
iIntros "[H↦ H↦']". wp_seq. wp_write.
iMod ("Hclose2" $! ((l + 2) ↦∗: ty_own ty tid)%I
with "[Hrc'↦ Hrc0 Hrc1 H†] [H↦ Hown]") as "[Hb Hα1]"; [|by auto with iFrame|].
{ iIntros "!> H↦". iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iLeft. iFrame. iDestruct "H†" as "[?|%]"=>//. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
iApply (type_type _ _ _ [ rcx box (uninit 1); (#l + #2) &uniq{α}ty;
r box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
- iIntros "[Htok Hν]". wp_case. wp_apply wp_new=>//.
iIntros (l) "/= (H† & Hl)". wp_let. wp_op.
rewrite heap_pointsto_vec_singleton. wp_write. wp_let.
wp_bind (of_val clone).
iApply (wp_wand with "[Hna]").
{ iApply (Hclone _ [] $! _ 1%Qp with "LFT HE Hna"); rewrite /llctx_interp /tctx_interp //. }
clear Hclone clone. iIntros (clone) "(Hna & _ & [Hclone _])". rewrite tctx_hasty_val.
iDestruct "Hs" as "[Hs|Hν']"; last by iDestruct (lft_tok_dead with "Hν Hν'") as "[]".
iDestruct (lft_intersect_acc with "Hα2 Hν") as (q'') "[Hαν Hclose3]".
rewrite -[α ν](right_id_L).
iApply (type_call_iris _ [α ν] (α ν) [_] with
"LFT HE Hna Hαν Hclone [Hl H†]"); [solve_typing| |].
{ rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full.
iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iApply ty_shr_mono; last done. iApply lft_intersect_incl_r. }
iIntros ([[|cl|]|]) "Hna Hαν Hcl //". wp_rec.
iDestruct "Hcl" as "[Hcl Hcl†]". iDestruct "Hcl" as (vl) "[Hcl↦ Hown]".
iDestruct (ty_size_eq with "Hown") as %Hsz.
iDestruct ("Hclose3" with "Hαν") as "[Hα2 Hν]".
wp_apply wp_new=>//; first lia. iIntros (l') "(Hl'† & Hl')". wp_let. wp_op.
rewrite shift_loc_0. rewrite -!lock Nat2Z.id.
rewrite !heap_pointsto_vec_cons shift_loc_assoc.
iDestruct "Hl'" as "(Hl' & Hl'1 & Hl'2)". wp_write. wp_op. wp_write. wp_op.
wp_apply (wp_memcpy with "[$Hl'2 $Hcl↦]").
{ by rewrite repeat_length. } { by rewrite Hsz. }
iIntros "[Hl'2 Hcl↦]". wp_seq. rewrite freeable_sz_full.
wp_apply (wp_delete with "[$Hcl↦ Hcl†]");
[lia|by replace (length vl) with (ty.(ty_size))|].
iIntros "_". wp_seq. wp_write.
iMod ("Hclose2" $! ((l' + 2) ↦∗: ty_own ty tid)%I with
"[Hrc'↦ Hl' Hl'1 Hl'†] [Hl'2 Hown]") as "[Hl' Hα1]".
{ iIntros "!> H". iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iLeft. iFrame. iDestruct "Hl'†" as "[?|%]"=>//. }
{ iExists _. iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
iApply (type_type _ _ _ [ #rc arc ty; #l' + #2 &uniq{α}ty;
r box (uninit 1); rcx box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 3!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
!tctx_hasty_val' //. unlock. iFrame. iRight.
iExists _, _, _. iFrame "∗#". }
iApply type_letalloc_1; [solve_typing..|]. iIntros (rcold). simpl_subst.
iApply type_let.
{ apply arc_drop_type. }
{ solve_typing. }
iIntros (drop). simpl_subst.
iApply (type_letcall ()); [solve_typing..|]. iIntros (content). simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End arc.
From iris.algebra Require Import auth numbers.
From iris.proofmode Require Import proofmode.
From lrust.lang Require Import proofmode notation lib.new_delete.
From lrust.lifetime Require Import meta.
From lrust.typing Require Import typing.
From lrust.typing.lib Require Import option.
From iris.prelude Require Import options.
Definition brandidx_stR := max_natUR.
Class brandidxG Σ := BrandedIdxG {
brandidx_inG :: inG Σ (authR brandidx_stR);
brandidx_gsingletonG :: lft_metaG Σ;
}.
Definition brandidxΣ : gFunctors
:= #[GFunctor (authR brandidx_stR); lft_metaΣ].
Global Instance subG_brandidxΣ {Σ} : subG brandidxΣ Σ brandidxG Σ.
Proof. solve_inG. Qed.
Definition brandidxN := lrustN .@ "brandix".
Definition brandvecN := lrustN .@ "brandvec".
Section brandedvec.
Context `{!typeGS Σ, !brandidxG Σ}.
Implicit Types (q : Qp) (α : lft) (γ : gname) (n m : nat).
Local Notation iProp := (iProp Σ).
Definition brandvec_inv α n : iProp :=
( γ, lft_meta α γ own γ ( MaxNat n))%I.
Program Definition brandvec (α : lft) : type :=
{| ty_size := int.(ty_size);
ty_own tid vl :=
( n, brandvec_inv α n vl = [ #n ])%I;
ty_shr κ tid l :=
( n, &at{κ,brandvecN}(brandvec_inv α n) &frac{κ}(λ q, l ↦∗{q} [ #n ]))%I;
|}.
Next Obligation. iIntros "* H". iDestruct "H" as (?) "[_ %]". by subst. Qed.
Next Obligation.
iIntros (ty E κ l tid q ?) "#LFT Hown Hκ".
iMod (bor_exists with "LFT Hown") as (vl) "Hown"; first solve_ndisj.
iMod (bor_sep with "LFT Hown") as "[Hshr Hown]"; first solve_ndisj.
iMod (bor_exists with "LFT Hown") as (n) "Hown"; first solve_ndisj.
iMod (bor_sep with "LFT Hown") as "[Hghost Hown]"; first solve_ndisj.
iMod (bor_share _ brandvecN with "Hghost") as "Hghost"; first solve_ndisj.
iMod (bor_persistent with "LFT Hown Hκ") as "[> % $]"; first solve_ndisj.
subst. iExists n. iFrame.
by iApply (bor_fracture with "LFT"); first solve_ndisj.
Qed.
Next Obligation.
iIntros (ty ?? tid l) "#H⊑ H".
iDestruct "H" as (n) "[Hghost Hown]".
iExists n. iSplitR "Hown".
- by iApply (at_bor_shorten with "H⊑").
- by iApply (frac_bor_shorten with "H⊑").
Qed.
Global Instance brandvec_wf α : TyWf (brandvec α) :=
{ ty_lfts := [α]; ty_wf_E := [] }.
Global Instance brandvec_ne : NonExpansive brandvec.
Proof. solve_proper. Qed.
Global Instance brandvec_mono E L :
Proper (lctx_lft_eq E L ==> subtype E L) brandvec.
Proof. apply lft_invariant_subtype. Qed.
Global Instance brandvec_equiv E L :
Proper (lctx_lft_eq E L ==> eqtype E L) brandvec.
Proof. apply lft_invariant_eqtype. Qed.
Global Instance brandvec_send α :
Send (brandvec α).
Proof. by unfold brandvec, Send. Qed.
Global Instance brandvec_sync α :
Sync (brandvec α).
Proof. by unfold brandvec, Sync. Qed.
(** [γ] is (a lower bound of) the length of the vector; as an in-bounds index,
that must be at least one more than the index value. *)
Definition brandidx_inv α m : iProp :=
( γ, lft_meta α γ own γ ( MaxNat (m+1)%nat))%I.
Program Definition brandidx α :=
{| ty_size := int.(ty_size);
ty_own tid vl := ( m, brandidx_inv α m vl = [ #m])%I;
ty_shr κ tid l := ( m, &frac{κ}(λ q, l ↦∗{q} [ #m]) brandidx_inv α m)%I;
|}.
Next Obligation. iIntros "* H". iDestruct "H" as (?) "[_ %]". by subst. Qed.
Next Obligation.
iIntros (ty E κ l tid q ?) "#LFT Hown Hκ".
iMod (bor_exists with "LFT Hown") as (vl) "Hown"; first solve_ndisj.
iMod (bor_sep with "LFT Hown") as "[Hown Hghost]"; first solve_ndisj.
iMod (bor_persistent with "LFT Hghost Hκ") as "[>Hghost $]"; first solve_ndisj.
iDestruct "Hghost" as (m) "[Hghost %]". subst vl.
iExists m. iFrame.
by iApply (bor_fracture with "LFT"); first solve_ndisj.
Qed.
Next Obligation.
iIntros (ty ?? tid l) "#H⊑ H".
iDestruct "H" as (m) "[H ?]".
iExists m. iFrame.
by iApply (frac_bor_shorten with "H⊑").
Qed.
Global Instance brandidx_wf α : TyWf (brandidx α) :=
{ ty_lfts := [α]; ty_wf_E := [] }.
Global Instance brandidx_ne : NonExpansive brandidx.
Proof. solve_proper. Qed.
Global Instance brandidx_mono E L :
Proper (lctx_lft_eq E L ==> subtype E L) brandidx.
Proof. apply lft_invariant_subtype. Qed.
Global Instance brandidx_equiv E L :
Proper (lctx_lft_eq E L ==> eqtype E L) brandidx.
Proof. apply lft_invariant_eqtype. Qed.
Global Instance brandidx_send α :
Send (brandidx α).
Proof. by unfold brandidx, Send. Qed.
Global Instance brandidx_sync α :
Sync (brandidx α).
Proof. by unfold brandidx, Sync. Qed.
Global Instance brandidx_copy α :
Copy (brandidx α).
Proof.
split; first by apply _.
iIntros (κ tid E ? l q ? HF) "#LFT #Hshr Htok Hlft".
iDestruct (na_own_acc with "Htok") as "[$ Htok]"; first solve_ndisj.
iDestruct "Hshr" as (γ) "[Hmem Hinv]".
iMod (frac_bor_acc with "LFT Hmem Hlft") as (q') "[>Hmt Hclose]"; first solve_ndisj.
iDestruct "Hmt" as "[Hmt1 Hmt2]".
iModIntro. iExists _.
iSplitL "Hmt1"; first by simpl; iFrame "∗#".
iIntros "Htok2 Hmt1". iDestruct "Hmt1" as (vl') "[>Hmt1 #Hown']".
iDestruct ("Htok" with "Htok2") as "$".
iAssert ( 1 = length vl')%I as ">%".
{ iNext. by iDestruct ("Hown'") as "[%x [Hown' ->]]". }
destruct vl' as [|v' vl']; first done.
destruct vl' as [|]; last first. { simpl in *. lia. }
rewrite !heap_pointsto_vec_singleton.
iDestruct (heap_pointsto_agree with "[$Hmt1 $Hmt2]") as %->.
iCombine "Hmt1" "Hmt2" as "Hmt".
iApply "Hclose". done.
Qed.
Lemma brandinv_brandidx_lb α m n :
brandvec_inv α n -∗ brandidx_inv α m -∗ m < n⌝%nat.
Proof.
iIntros "Hn Hm".
iDestruct "Hn" as (γn) "(Hidx1 & Hn)".
iDestruct "Hm" as (γm) "(Hidx2 & Hm)".
iDestruct (lft_meta_agree with "Hidx1 Hidx2") as %<-.
iCombine "Hn Hm" gives %[?%max_nat_included ?]%auth_both_valid_discrete.
iPureIntro. simpl in *. lia.
Qed.
End brandedvec.
Section typing.
Context `{!typeGS Σ, !brandidxG Σ}.
Implicit Types (q : Qp) (α : lft) (γ : gname) (n m : nat).
Local Notation iProp := (iProp Σ).
Definition brandvec_new (call_once : val) (R_F : type) : val :=
fn: ["f"] :=
let: "call_once" := call_once in
letalloc: "n" <- #0 in
letcall: "r" := "call_once" ["f";"n"]%E in
letalloc: "r'" <-{ R_F.(ty_size)} !"r" in
delete [ #R_F.(ty_size); "r"];;
return: ["r'"].
Lemma brandvec_new_type F R_F call_once `{!TyWf F, !TyWf R_F} :
typed_val call_once (fn( α, ; F, brandvec α) R_F)
typed_val (brandvec_new call_once R_F) (fn(; F) R_F).
Proof.
iIntros (Hf E L). iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret args). inv_vec args=> env. simpl_subst.
iApply type_let; [apply Hf|solve_typing|]. iIntros (f); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hc (Hf & Henv & _)".
wp_let.
wp_op.
wp_case.
wp_alloc n as "Hn" "Hdead".
wp_let.
rewrite !heap_pointsto_vec_singleton.
wp_write.
iAssert ( E : coPset, ⌜↑lftN E
|={E}=> α, tctx_elt_interp tid ((LitV n : expr) box (brandvec α)) 1.[α])%I
with "[Hn Hdead]" as "Hn'".
{ iIntros (E' Hlft).
iMod (own_alloc ( (MaxNat 0))) as (γ) "Hown"; first by apply auth_auth_valid.
iMod (lft_create_meta γ with "LFT") as (α) "[#Hsing [Htok #Hα]]"; first done.
iExists α.
rewrite !tctx_hasty_val.
rewrite ownptr_own.
rewrite -heap_pointsto_vec_singleton.
iFrame "Htok".
iExists n, (Vector.cons (LitV 0) Vector.nil).
iSplitR; first done.
simpl.
rewrite freeable_sz_full.
by iFrame "#∗". }
iMod ("Hn'" with "[%]") as (α) "[Hn Htok]"; [solve_typing..|].
wp_let.
iMod (lctx_lft_alive_tok ϝ with "HE HL")
as (qϝf) "(Hϝf & HL & Hclosef)"; [solve_typing..|].
iDestruct (lft_intersect_acc with "Htok Hϝf") as (?) "[Hαϝ Hcloseα]".
rewrite !tctx_hasty_val.
iApply (type_call_iris _ [α; ϝ] α [_;_] _ _ _ _
with "LFT HE Hna [Hαϝ] Hf [Hn Henv]"); try solve_typing.
+ by rewrite /= right_id.
+ rewrite /= right_id.
rewrite !tctx_hasty_val.
by iFrame.
+ iIntros (r) "Hna Hf Hown".
simpl_subst.
iDestruct ("Hcloseα" with "[Hf]") as "[Htok Hf]"; first by rewrite right_id.
iMod ("Hclosef" with "Hf HL") as "HL".
wp_let.
iApply (type_type _ _ _ [ r box R_F ] with "[] LFT HE Hna HL Hc");
try solve_typing; last by rewrite !tctx_interp_singleton tctx_hasty_val.
iApply type_letalloc_n; [solve_typing..|].
iIntros (r').
simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition brandvec_get_index : val :=
fn: ["v"; "i"] :=
let: "r" := new [ #2 ] in
let: "v'" := !"v" in
let: "i'" := !"i" in
delete [ #1; "v" ];; delete [ #1; "i" ];;
let: "inbounds" := (if: #0 "i'" then "i'" < !"v'" else #false) in
if: "inbounds" then
"r" <-{Σ some} "i'";;
return: ["r"]
else
"r" <-{Σ none} ();;
return: ["r"].
Lemma brandvec_get_index_type :
typed_val brandvec_get_index (fn( '(α, β), ; &shr{β} (brandvec α), int) option (brandidx α))%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret args). inv_vec args=>v i. simpl_subst.
iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (v'); simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (i'); simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hi' & #Hv' & Hr & _)".
rewrite !tctx_hasty_val. clear.
destruct i' as [[| |i']|]; try done. iClear "Hi'".
wp_op.
rewrite bool_decide_decide.
destruct (decide (0 i')) as [Hpos|Hneg]; last first.
{ wp_case. wp_let. wp_case.
iApply (type_type _ _ _ [ r box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_singleton tctx_hasty_val. done. }
iApply (type_sum_unit (option (brandidx _))); [solve_typing..|].
iApply type_jump; solve_typing. }
destruct (Z_of_nat_complete _ Hpos) as [i ->]. clear Hpos.
wp_case. wp_op.
iDestruct (shr_is_ptr with "Hv'") as % [l ?]. simplify_eq.
iDestruct "Hv'" as (m) "#[Hghost Hmem]".
iMod (lctx_lft_alive_tok β with "HE HL") as () "(Hβ & HL & Hclose)"; [solve_typing..|].
iMod (frac_bor_acc with "LFT Hmem Hβ") as (qβ') "[>Hn'↦ Hcloseβ]"; first done.
rewrite !heap_pointsto_vec_singleton.
wp_read.
iMod ("Hcloseβ" with "Hn'↦") as "Hβ".
wp_op.
rewrite bool_decide_decide.
destruct (decide (i + 1 m)) as [Hle|Hoob]; last first.
{ wp_let. wp_case.
iMod ("Hclose" with "Hβ HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_singleton tctx_hasty_val. done. }
iApply (type_sum_unit (option (brandidx _))); [solve_typing..|].
iApply type_jump; solve_typing. }
wp_let. wp_case.
iApply fupd_wp.
iMod (at_bor_acc_tok with "LFT Hghost Hβ") as "[>Hidx Hcloseg]"; [solve_ndisj..|].
iDestruct "Hidx" as (γ) "(#Hidx & Hown)".
iAssert (|==> own γ ( (MaxNat m)) own γ ( (MaxNat m)))%I with "[Hown]" as "> [Hown Hlb]".
{ rewrite -own_op. iApply (own_update with "Hown"). apply auth_update_alloc.
by apply max_nat_local_update. }
iMod ("Hcloseg" with "[Hown]") as "Hβ".
{ iExists _. eauto with iFrame. }
iMod ("Hclose" with "Hβ HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 2); #i brandidx _ ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. iFrame.
rewrite tctx_hasty_val'; last done.
iExists _. iSplit; last done.
iExists _. iFrame "Hidx".
iApply base_logic.lib.own.own_mono; last done.
apply: auth_frag_mono. apply max_nat_included. simpl. lia. }
iApply (type_sum_assign (option (brandidx _))); [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition brandidx_get : val :=
fn: ["s";"c"] :=
let: "len" := !"s" in
let: "idx" := !"c" in
delete [ #1; "s" ];; delete [ #1; "c" ];;
if: !"idx" < !"len" then
let: "r" := new [ #0] in return: ["r"]
else
!#☠ (* stuck *).
Lemma brandidx_get_type :
typed_val brandidx_get (fn( '(α, β), ; &shr{β} (brandvec α), &shr{β} (brandidx α)) unit)%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret args). inv_vec args=> s c. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (n); simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (m); simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iIntros (tid qmax) "#LFT #HE Hna HL HC (Hm & Hn & _)".
rewrite !tctx_hasty_val.
iDestruct (shr_is_ptr with "Hm") as %[lm ?]. simplify_eq.
iDestruct (shr_is_ptr with "Hn") as %[ln ?]. simplify_eq.
simpl in *.
iDestruct "Hm" as (m) "[Hm Hmidx]".
iDestruct "Hn" as (n) "[Hnidx Hn]".
iMod (lctx_lft_alive_tok β with "HE HL") as () "((Hβ1 & Hβ2 & Hβ3) & HL & Hclose)"; [solve_typing..|].
iMod (frac_bor_acc with "LFT Hn Hβ1") as (qβn) "[>Hn↦ Hcloseβ1]"; first solve_ndisj.
iMod (frac_bor_acc with "LFT Hm Hβ2") as (qβm) "[>Hm↦ Hcloseβ2]"; first solve_ndisj.
rewrite !heap_pointsto_vec_singleton.
wp_read.
wp_op.
wp_read.
wp_op.
wp_case.
iApply fupd_wp.
iMod (at_bor_acc_tok with "LFT Hnidx Hβ3") as "[>Hnidx Hcloseβ3]"; [solve_ndisj..|].
iDestruct (brandinv_brandidx_lb with "Hnidx Hmidx") as %Hle.
iMod ("Hcloseβ3" with "Hnidx") as "Hβ3".
iMod ("Hcloseβ2" with "Hm↦") as "Hβ2".
iMod ("Hcloseβ1" with "Hn↦") as "Hβ1".
iCombine "Hβ2 Hβ3" as "Hβ2".
iMod ("Hclose" with "[$Hβ1 $Hβ2] HL") as "HL".
rewrite bool_decide_true; last by lia.
iApply (type_type _ _ _ []
with "[] LFT HE Hna HL HC []"); last by rewrite tctx_interp_nil.
iApply (type_new _); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_jump; solve_typing.
Qed.
Definition brandvec_push : val :=
fn: ["s"] :=
let: "n" := !"s" in
delete [ #1; "s" ];;
let: "r" := new [ #1] in
let: "oldlen" := !"n" in
"n" <- "oldlen"+#1;;
"r" <- "oldlen";;
return: ["r"].
Lemma brandvec_push_type :
typed_val brandvec_push (fn( '(α, β), ; &uniq{β} (brandvec α)) brandidx α).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret args). inv_vec args=>(*n *)s. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (n); simpl_subst.
iApply type_delete; [solve_typing..|].
iApply (type_new _); [solve_typing..|]; iIntros (r); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL HC (Hr & Hn & _)".
rewrite !tctx_hasty_val.
iDestruct (uniq_is_ptr with "Hn") as %[ln H]. simplify_eq.
iMod (lctx_lft_alive_tok β with "HE HL") as () "(Hβ & HL & Hclose)"; [solve_typing..|].
iMod (bor_acc with "LFT Hn Hβ") as "[H↦ Hclose']"; first solve_ndisj.
iDestruct "H↦" as (vl) "[> H↦ Hn]".
iDestruct "Hn" as (n) "[Hn > %]". simplify_eq.
rewrite !heap_pointsto_vec_singleton.
wp_read.
wp_let.
wp_op.
wp_write.
iDestruct "Hn" as (γ) "[#Hidx Hown]".
iMod (own_update _ _ ( MaxNat (n+1) _) with "Hown") as "[Hown Hlb]".
{ apply auth_update_alloc.
apply max_nat_local_update.
simpl. lia.
}
iDestruct "Hlb" as "#Hlb".
iMod ("Hclose'" with "[H↦ Hidx Hown]") as "[Hn Hβ]".
{ iExists (#(n+1)::nil).
rewrite heap_pointsto_vec_singleton.
iFrame "#∗". iPureIntro; do 3 f_equal; lia.
}
iMod ("Hclose" with "Hβ HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 1); #n brandidx _]
with "[] LFT HE Hna HL HC [Hr]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. iFrame.
rewrite tctx_hasty_val'; last done.
iExists _. iSplit; last done.
iExists _. eauto with iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End typing.
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 cell.
Context `{!typeGS Σ}.
Program Definition cell (ty : type) :=
{| ty_size := ty.(ty_size);
ty_own := ty.(ty_own);
ty_shr κ tid l := (&na{κ, tid, shrN.@l}(l ↦∗: ty.(ty_own) tid))%I |}.
Next Obligation. apply ty_size_eq. Qed.
Next Obligation.
iIntros (ty E κ l tid q ?) "#LFT Hown $". by iApply (bor_na with "Hown").
Qed.
Next Obligation.
iIntros (ty ?? tid l) "#H⊑ H". by iApply (na_bor_shorten with "H⊑").
Qed.
Global Instance cell_wf ty `{!TyWf ty} : TyWf (cell ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance cell_type_ne : TypeNonExpansive cell.
Proof. solve_type_proper. Qed.
Global Instance cell_ne : NonExpansive cell.
Proof.
constructor;
solve_proper_core ltac:(fun _ => (eapply ty_size_ne; try reflexivity) || f_equiv).
Qed.
Global Instance cell_mono E L : Proper (eqtype E L ==> subtype E L) cell.
Proof.
move=>?? /eqtype_unfold EQ. iIntros (??) "HL".
iDestruct (EQ with "HL") as "#EQ". iIntros "!> #HE".
iDestruct ("EQ" with "HE") as "(% & #Hown & #Hshr)".
iSplit; [done|iSplit; iIntros "!> * H"].
- iApply ("Hown" with "H").
- iApply na_bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H";
iDestruct "H" as (vl) "[??]"; iExists vl; iFrame; by iApply "Hown".
Qed.
Lemma cell_mono' E L ty1 ty2 : eqtype E L ty1 ty2 subtype E L (cell ty1) (cell ty2).
Proof. eapply cell_mono. Qed.
Global Instance cell_proper E L : Proper (eqtype E L ==> eqtype E L) cell.
Proof. by split; apply cell_mono. Qed.
Lemma cell_proper' E L ty1 ty2 : eqtype E L ty1 ty2 eqtype E L (cell ty1) (cell ty2).
Proof. eapply cell_proper. Qed.
(* The real [Cell] in Rust is never [Copy], but that is mostly for reasons of
ergonomics -- it is widely agreed that it would be sound for [Cell] to be
[Copy]. [Cell] is also rather unique as an interior mutable [Copy] type, so
proving this is a good benchmark. *)
Global Instance cell_copy ty :
Copy ty Copy (cell ty).
Proof.
intros Hcopy. split; first by intros; simpl; unfold ty_own; apply _.
iIntros (κ tid E F l q ??) "#LFT #Hshr Htl Htok". iExists 1%Qp. simpl in *.
(* Size 0 needs a special case as we can't keep the thread-local invariant open. *)
destruct (ty_size ty) as [|sz] eqn:Hsz; simpl in *.
{ iMod (na_bor_acc with "LFT Hshr Htok Htl") as "(Hown & Htl & Hclose)"; [solve_ndisj..|].
iDestruct "Hown" as (vl) "[H↦ #Hown]".
simpl. assert (F = F) as -> by set_solver+.
iDestruct (ty_size_eq with "Hown") as "#>%". rewrite ->Hsz in *.
iMod ("Hclose" with "[H↦] Htl") as "[$ $]".
{ iExists vl. by iFrame. }
iModIntro. iSplitL "".
{ iNext. iExists vl. destruct vl; last done. iFrame "Hown".
by iApply heap_pointsto_vec_nil. }
by iIntros "$ _". }
(* Now we are in the non-0 case. *)
iMod (na_bor_acc with "LFT Hshr Htok Htl") as "($ & Htl & Hclose)"; [solve_ndisj..|].
iDestruct (na_own_acc with "Htl") as "($ & Hclose')"; first by set_solver.
iIntros "!> Htl Hown". iPoseProof ("Hclose'" with "Htl") as "Htl".
by iMod ("Hclose" with "Hown Htl") as "[$ $]".
Qed.
Global Instance cell_send ty :
Send ty Send (cell ty).
Proof. by unfold cell, Send. Qed.
End cell.
Section typing.
Context `{!typeGS Σ}.
(** The next couple functions essentially show owned-type equalities, as they
are all different types for the identity function. *)
(* Constructing a cell. *)
Definition cell_new : val := fn: ["x"] := return: ["x"].
Lemma cell_new_type ty `{!TyWf ty} : typed_val cell_new (fn(; ty) cell ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_jump; [solve_typing..|].
iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
Qed.
(* The other direction: getting ownership out of a cell. *)
Definition cell_into_inner : val := fn: ["x"] := return: ["x"].
Lemma cell_into_inner_type ty `{!TyWf ty} :
typed_val cell_into_inner (fn(; cell ty) ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_jump; [solve_typing..|].
iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
Qed.
Definition cell_get_mut : val :=
fn: ["x"] := return: ["x"].
Lemma cell_get_mut_type ty `{!TyWf ty} :
typed_val cell_get_mut (fn( α, ; &uniq{α} (cell ty)) &uniq{α} ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_jump; [solve_typing..|].
iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
Qed.
Definition cell_from_mut : val :=
fn: ["x"] := return: ["x"].
Lemma cell_from_mut_type ty `{!TyWf ty} :
typed_val cell_from_mut (fn( α, ; &uniq{α} ty) &uniq{α} (cell ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_jump; [solve_typing..|].
iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
Qed.
Definition cell_into_box : val :=
fn: ["x"] := return: ["x"].
Lemma cell_into_box_type ty `{!TyWf ty} :
typed_val cell_into_box (fn(;box (cell ty)) box ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_jump; [solve_typing..|].
iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
Qed.
Definition cell_from_box : val :=
fn: ["x"] := return: ["x"].
Lemma cell_from_box_type ty `{!TyWf ty} :
typed_val cell_from_box (fn(; box ty) box (cell ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply type_jump; [solve_typing..|].
iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
Qed.
(** Reading from a cell *)
Definition cell_get ty : val :=
fn: ["x"] :=
let: "x'" := !"x" in
letalloc: "r" <-{ty.(ty_size)} !"x'" in
delete [ #1; "x"];; return: ["r"].
Lemma cell_get_type ty `{!TyWf ty} `(!Copy ty) :
typed_val (cell_get ty) (fn( α, ; &shr{α} (cell ty)) 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.
iApply type_letalloc_n; [solve_typing| |iIntros (r); simpl_subst].
{ rewrite typed_read_eq.
have Hrd := (read_shr _ _ _ (cell ty)). rewrite typed_read_eq in Hrd.
apply Hrd; solve_typing. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(** Writing to a cell *)
Definition cell_replace ty : val :=
fn: ["c"; "x"] :=
let: "c'" := !"c" in
letalloc: "r" <-{ty.(ty_size)} !"c'" in
"c'" <-{ty.(ty_size)} !"x";;
delete [ #1; "c"] ;; delete [ #ty.(ty_size); "x"] ;; return: ["r"].
Lemma cell_replace_type ty `{!TyWf ty} :
typed_val (cell_replace ty) (fn( α, ; &shr{α}(cell ty), ty) ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>c x. simpl_subst.
iApply type_deref; [solve_typing..|].
iIntros (c'); simpl_subst.
iApply (type_new (ty_size ty)); [solve_typing..|]; iIntros (r); simpl_subst.
(* Drop to Iris level. *)
iIntros (tid qmax) "#LFT #HE Htl HL HC".
rewrite 3!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iIntros "(Hr & Hc & #Hc' & Hx)".
destruct c' as [[|c'|]|]; try done. destruct x as [[|x|]|]; try done.
destruct r as [[|r|]|]; try done.
iMod (lctx_lft_alive_tok α with "HE HL") as (q') "(Htok & HL & Hclose1)"; [solve_typing..|].
iMod (na_bor_acc with "LFT Hc' Htok Htl") as "(Hc'↦ & Htl & Hclose2)"; [solve_ndisj..|].
iDestruct "Hc'↦" as (vc') "[>Hc'↦ Hc'own]".
iDestruct (ty_size_eq with "Hc'own") as "#>%".
iDestruct "Hr" as "[Hr↦ Hr†]". iDestruct "Hr↦" as (vr) "[>Hr↦ Hrown]".
iDestruct (ty_size_eq with "Hrown") as ">Heq". iDestruct "Heq" as %Heq.
(* FIXME: Changing the order of $Hr↦ $Hc'↦ breaks applying...?? *)
wp_apply (wp_memcpy with "[$Hr↦ $Hc'↦]").
{ rewrite Heq //. }
{ f_equal. done. }
iIntros "[Hr↦ Hc'↦]". wp_seq.
iDestruct "Hx" as "[Hx↦ Hx†]". iDestruct "Hx↦" as (vx) "[Hx↦ Hxown]".
iDestruct (ty_size_eq with "Hxown") as "#%".
wp_apply (wp_memcpy with "[$Hc'↦ $Hx↦]"); try by f_equal.
iIntros "[Hc'↦ Hx↦]". wp_seq.
iMod ("Hclose2" with "[Hc'↦ Hxown] Htl") as "[Htok Htl]"; first by auto with iFrame.
iMod ("Hclose1" with "Htok HL") as "HL".
(* Now go back to typing level. *)
iApply (type_type _ _ _
[c box (&shr{α}(cell ty)); #x box (uninit ty.(ty_size)); #r box ty]
with "[] LFT HE Htl HL HC"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
iFrame "Hc". rewrite !tctx_hasty_val' //. iSplitL "Hx↦ Hx†"; by iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(** Create a shared cell from a mutable borrow.
Called alias::one in Rust.
This is really just [cell_from_mut] followed by sharing. *)
Definition fake_shared_cell : val :=
fn: ["x"] :=
let: "x'" := !"x" in
letalloc: "r" <- "x'" in
delete [ #1; "x"];; return: ["r"].
Lemma fake_shared_cell_type ty `{!TyWf ty} :
typed_val fake_shared_cell (fn( α, ; &uniq{α} ty) &shr{α}(cell ty)).
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 HT".
rewrite tctx_interp_singleton tctx_hasty_val.
iApply (type_type _ _ _ [ x box (&uniq{α}(cell ty)) ]
with "[] LFT HE Hna HL Hk [HT]"); last first.
{ by rewrite tctx_interp_singleton tctx_hasty_val. }
iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
iApply (type_letalloc_1 (&shr{α}(cell ty))); [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End typing.
Global Hint Resolve cell_mono' cell_proper' : lrust_typing.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Export type.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Section diverging_static.
Context `{!typeGS Σ}.
(* Show that we can convert any live borrow to 'static with an infinite
loop. *)
Definition diverging_static_loop (call_once : val) : val :=
fn: ["x"; "f"] :=
let: "call_once" := call_once in
letcall: "ret" := "call_once" ["f"; "x"]%E in
withcont: "loop":
"loop" []
cont: "loop" [] :=
"loop" [].
Lemma diverging_static_loop_type T F call_once
`{!TyWf T, !TyWf F} :
(* F : FnOnce(&'static T), as witnessed by the impl call_once *)
typed_val call_once (fn(; F, &shr{static} T) unit)
typed_val (diverging_static_loop call_once)
(fn( α, λ ϝ, ty_outlives_E T static; &shr{α} T, F) ).
Proof.
intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x f. simpl_subst.
iApply type_let; [apply Hf|solve_typing|]. iIntros (call); simpl_subst.
(* Drop to Iris *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk (Hcall & Hx & Hf & _)".
(* We need a ϝ token to show that we can call F despite it being non-'static. *)
iMod (lctx_lft_alive_tok ϝ with "HE HL") as () "(Hϝ & HL & _)";
[solve_typing..|].
iMod (lctx_lft_alive_tok_noend α with "HE HL") as (q) "(Hα & _ & _)";
[solve_typing..|].
iMod (lft_eternalize with "Hα") as "#Hα".
iAssert (type_incl (box (&shr{α} T)) (box (&shr{static} T))) as "#[_ [Hincl _]]".
{ iApply box_type_incl. iNext. iApply shr_type_incl; first done.
iApply type_incl_refl. }
wp_let. rewrite tctx_hasty_val.
iApply (type_call_iris _ [ϝ] () [_; _] with "LFT HE Hna [Hϝ] Hcall [Hx Hf]").
- solve_typing.
- by rewrite /= (right_id static).
- simpl. iFrame. iSplit; last done. rewrite !tctx_hasty_val.
iApply "Hincl". done.
- iClear "Hα Hincl". iIntros (r) "Htl Hϝ1 Hret". wp_rec.
(* Go back to the type system. *)
iApply (type_type _ [] _ [] with "[] LFT HE Htl [] Hk [-]"); last first.
{ rewrite /tctx_interp /=. done. }
{ rewrite /llctx_interp /=. done. }
iApply (type_cont [] [] (λ _, [])).
+ iIntros (?). simpl_subst. iApply type_jump; solve_typing.
+ iIntros "!>" (? args). inv_vec args. simpl_subst. iApply type_jump; solve_typing.
Qed.
(** With the right typing rule, we can prove a variant of the above where the
lifetime is in an arbitrary invariant position, without even leaving the
type system. This is incompatible with branding!
Our [TyWf] is not working well for type constructors (we have no way of
representing the fact that well-formedness is somewhat "uniform"), so we
instead work with a constant [Euser] of lifetime inclusion assumptions (in
general it could change with the type parameter but only in monotone ways)
and a single [κuser] of lifetimes that have to be alive (making κuser a
list would require some induction-based reasoning principles that we do
not have, but showing that it works for one lifetime is enough to
demonstrate the principle). *)
Lemma diverging_static_loop_type_bad (T : lft type) F κuser Euser call_once
`{!TyWf F} :
(* The "bad" type_equivalize_lft_static rule *)
( E L C T κ e,
( typed_body ((static κ) :: E) L C T e)
typed_body E ((κ []) :: L) C T e)
(* Typing of this function *)
let _ : ( κ, TyWf (T κ)) := λ κ, {| ty_lfts := [κ; κuser]; ty_wf_E := Euser |} in
( κ1 κ2, (T κ1).(ty_size) = (T κ2).(ty_size))
( E L κ1 κ2, lctx_lft_eq E L κ1 κ2 subtype E L (T κ1) (T κ2))
typed_val call_once (fn(; F, T static) unit)
typed_val (diverging_static_loop call_once)
(fn( α, ; T α, F) ).
Proof.
intros type_equivalize_lft_static_bad HWf HTsz HTsub Hf E L.
iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x f. simpl_subst.
iApply type_let; [apply Hf|solve_typing|]. iIntros (call); simpl_subst.
iApply (type_cont [_] [] (λ r, [(r!!!0%fin:val) box (unit)])).
{ iIntros (k). simpl_subst.
iApply type_equivalize_lft_static_bad.
iApply (type_call [ϝ] ()); solve_typing. }
iIntros "!> %k %args". inv_vec args=>r. simpl_subst.
iApply (type_cont [] [] (λ r, [])).
{ iIntros (kloop). simpl_subst. iApply type_jump; solve_typing. }
iIntros "!> %k' %args". inv_vec args. simpl_subst.
iApply type_jump; solve_typing.
Qed.
End diverging_static.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Section fake_shared.
Context `{!typeGS Σ}.
Definition fake_shared_box : val :=
fn: ["x"] := Skip ;; return: ["x"].
Lemma fake_shared_box_type ty `{!TyWf ty} :
typed_val fake_shared_box
(fn( '(α, β), ; &shr{α}(&shr{β} ty)) &shr{α}(box ty)).
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 HT".
rewrite tctx_interp_singleton tctx_hasty_val.
iDestruct (lctx_lft_incl_incl α β with "HL HE") as "%"; [solve_typing..|].
iAssert ( ty_own (own_ptr 1 (&shr{α}(box ty))) tid [x])%I with "[HT]" as "HT".
{ destruct x as [[|l|]|]=>//=. iDestruct "HT" as "[H $]".
iNext. iDestruct "H" as ([|[[]|][]]) "[H↦ H]"; try done.
iExists _. iFrame. iDestruct "H" as (vl) "[#Hf H]".
iNext. destruct vl as [|[[|l'|]|][]]; try done. iExists l'. iSplit.
{ iApply frac_bor_iff; last done. iIntros "!>!> *".
rewrite heap_pointsto_vec_singleton. iSplit; auto. }
iDestruct "H" as "#H". iIntros "!> * % $". iApply step_fupd_intro; first set_solver.
simpl. iApply ty_shr_mono; last done.
by iApply lft_incl_syn_sem. }
do 2 wp_seq.
iApply (type_type [] _ _ [ x box (&shr{α}(box ty)) ]
with "[] LFT [] Hna HL Hk [HT]"); last first.
{ by rewrite tctx_interp_singleton tctx_hasty_val. }
{ by rewrite /elctx_interp. }
iApply type_jump; simpl; solve_typing.
Qed.
Definition fake_shared_uniq : val :=
fn: ["x"] := Skip ;; return: ["x"].
Lemma fake_shared_uniq_type ty `{!TyWf ty} :
typed_val fake_shared_box
(fn( '(α, β), ; &shr{α}(&shr{β} ty)) &shr{α}(&uniq{β} ty)).
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 HT".
rewrite tctx_interp_singleton tctx_hasty_val.
iDestruct (lctx_lft_incl_incl α β with "HL HE") as "%"; [solve_typing..|].
(* FIXME: WTF, using &uniq{β} here does not work. *)
iAssert ( ty_own (own_ptr 1 (&shr{α} (uniq_bor β ty))) tid [x])%I with "[HT]" as "HT".
{ destruct x as [[|l|]|]=>//=. iDestruct "HT" as "[H $]".
iNext. iDestruct "H" as ([|[[]|][]]) "[H↦ H]"; try done.
iExists _. iFrame. iDestruct "H" as (vl) "[#Hf H]".
iNext. destruct vl as [|[[|l'|]|][]]; try done. iExists l'. iSplit.
{ iApply frac_bor_iff; last done. iIntros "!>!> *".
rewrite heap_pointsto_vec_singleton. iSplit; auto. }
iDestruct "H" as "#H". iIntros "!> * % $". iApply step_fupd_intro; first set_solver.
simpl. iApply ty_shr_mono; last done.
by iApply lft_intersect_incl_l.
}
do 2 wp_seq.
iApply (type_type [] _ _ [ x box (&shr{α}(&uniq{β} ty)) ]
with "[] LFT [] Hna HL Hk [HT]"); last first.
{ by rewrite tctx_interp_singleton tctx_hasty_val. }
{ by rewrite /elctx_interp. }
iApply type_jump; simpl; solve_typing.
Qed.
End fake_shared.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import csum frac excl agree coPset.
From iris.bi Require Import lib.fractional.
From lrust.lang Require Import proofmode notation.
From lrust.lifetime Require Import meta.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
Definition ghostcellN := lrustN .@ "ghostcell".
Definition ghosttokenN := lrustN .@ "ghosttoken".
(* This library should be valid under weak memory, but the current
model of Lambdarust-weak makes this challenging to prove. The
reasons are:
- The SC model in this file relies crucially on using the atomic accessors
of borrows, which are much weaker (when they exist) in weakmem. The
reason of this weakness is that, when we are not in SC, we cannot tel
that a lifetime is "either alive or dead", since it can be "dead", but
in a way which is not yet synchronized with the current thread.
So let's search for another model:
- We have to prove that successive mutable accesses to the ghost cell are
synchronized (i.e., they are in happens-before relationship).
- The only reason these accesses are indeed synchronized is that the *ghost
token* is transmitted via mechanisms guaranteeing happens-before
relationship (like every resource transfer in Rust).
- In particular, when the borrow of a ghost token ends, the ghost token
needs to change its internal state to remember the view of the death of
its lifetime. Indeed, the death of the corresponding lifetime is an
essential link in the synchronization chain between two successive
accesses to the ghost cell.
- In the lifetime logic, the only way to allow updates at the death of a
lifetime is through the "consequence view-shift" of [bor_acc_strong]:
▷ Q -∗ [†κ'] ={userE}=∗ ▷ P
- However, unfortunately, this "consequence view-shift" take a subjective
version of the death token in weak_mem:
▷ Q -∗ subj [†κ'] ={userE}=∗ ▷ P
As a result, the view at which this view shift is used is independent
from the actual "view of the death of the lifetime", so we have no
information about it.
- The reason why this is a subjective token in weak-mem is because of what
happens when a non-atomic lifetime dies (a non-atomic lifetime is the
intersection of other lifetimes). Indeed, a non-atomic lifetime can die
"several times", once for each component, and each of these death can
occur at non-comparable views. The only relevant view that we would pass
to the consequence view-shift is the view used for the inheritance of
the borrow, but that view is not necessarily well-defined since the
borrow can be the result of a merging of other borrows.
Hence, it sounds like that the main blocking point is the merging rule of
the lifetime logic, which we use very rarely in the model (only for
merging actual borrows in [tctx_merge_uniq_prod2]). So it might be mossible
to remove the merging rule from the lifetime logic to get rid of the
subjective modality in the consequence view-shift and finally get a model
for GhostCell under weak memory...
*)
(* States:
Not borrowed, borrowed.
Shared states for all borrowed entities:
Not shared,
((lifetime at which shared), borrow fraction).
"Is borrowed" should agree.
*)
Definition ghosttoken_stR :=
(csumR (exclR unitO) ((prodR (agreeR lftO) fracR))).
Class ghostcellG Σ := GhostcellG {
ghostcell_inG :: inG Σ (ghosttoken_stR);
ghostcell_gsingletonG :: lft_metaG Σ;
}.
Definition ghostcellΣ : gFunctors
:= #[GFunctor ghosttoken_stR ; lft_metaΣ].
Global Instance subG_ghostcellΣ {Σ} : subG ghostcellΣ Σ ghostcellG Σ.
Proof. solve_inG. Qed.
(* There are two possible states:
- None: The GhostToken is currently not used or used in some
GhostCell::borrow_mut, to get mutable access.
- Some: The GhostToken is currently used in many GhostCell::get, to get
shared access. *)
Definition ghosttoken_st := option (lft * frac).
Definition ghosttoken_st_to_R (st : ghosttoken_st) : ghosttoken_stR :=
(match st with
| None => Cinl (Excl ())
| Some p => Cinr ((to_agree p.1, p.2))
end).
Section ghostcell.
Context `{!typeGS Σ, !ghostcellG Σ}.
Implicit Types (α β: lft) (γ: gname) (q: Qp) (ty: type) (l: loc).
Local Instance ghosttoken_fractional γ κ' :
Fractional (λ q, own γ (ghosttoken_st_to_R (Some (κ',q))))%I.
Proof.
rewrite /Fractional=>q1 q2.
rewrite -own_op /ghosttoken_st_to_R /=. f_equiv.
rewrite -Cinr_op. eapply Cinr_proper.
rewrite -pair_op. f_equiv; [|done..].
rewrite agree_idemp. done.
Qed.
Program Definition ghosttoken α :=
tc_opaque
{| ty_size := 0;
ty_own tid vl :=
(vl = [] γ, lft_meta α γ
own γ (ghosttoken_st_to_R None))%I;
ty_shr κ tid l :=
( γ, lft_meta α γ
κ', κ κ'
&frac{κ'}(λ q, own γ (ghosttoken_st_to_R (Some (κ',q)))))%I;
|}.
Next Obligation. by iIntros (??[|]) "[% ?]". Qed.
Next Obligation.
iIntros (α E κ l tid q ?) "#LFT Hvl Htok".
iMod (bor_exists with "LFT Hvl") as (vl) "Hvl"; first done.
iMod (bor_sep with "LFT Hvl") as "[Hvl Hset]"; first done.
iMod (bor_sep with "LFT Hset") as "[Hp Hset]"; first done.
iMod (bor_exists with "LFT Hset") as (γ) "Hset"; first done.
iMod (bor_sep with "LFT Hset") as "[Hidx Hset]"; first done.
iMod (bor_persistent with "LFT Hidx Htok") as "[>#Hidx Htok]"; first solve_ndisj.
iMod (bor_acc_strong with "LFT Hset Htok") as (κ') "(#Hκ & >Hset & Hclose)"; first done.
rewrite bi.sep_exist_r. iExists γ. iFrame "Hidx".
iMod (own_update _ _ (ghosttoken_st_to_R (Some (κ', 1%Qp))) with "Hset") as "Hset".
{ rewrite /ghosttoken_st_to_R /=. apply cmra_update_exclusive. done. }
iMod ("Hclose" with "[] Hset") as "[Hset $]".
{ iIntros "!> >Hset _".
iMod (own_update _ _ (ghosttoken_st_to_R None) with "Hset") as "Hset".
{ rewrite /ghosttoken_st_to_R /=. apply cmra_update_exclusive. done. }
done. }
iExists κ'. iMod (bor_fracture with "LFT [Hset]") as "$"; first done.
{ done. }
eauto.
Qed.
Next Obligation.
iIntros (?????) "#Hκ H".
iDestruct "H" as (γ) "[#? H]".
iDestruct "H" as (κ'') "[? ?]".
iFrame "∗#".
by iApply lft_incl_trans.
Qed.
Global Instance ghosttoken_wf α : TyWf (ghosttoken α) :=
{ ty_lfts := [α]; ty_wf_E := [] }.
Global Instance ghosttoken_ne : NonExpansive ghosttoken := _.
Global Instance ghosttoken_send α :
Send (ghosttoken α).
Proof. done. Qed.
Global Instance ghosttoken_sync α :
Sync (ghosttoken α).
Proof. done. Qed.
Global Instance ghosttoken_mono E L :
Proper (lctx_lft_eq E L ==> subtype E L) ghosttoken.
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic.
It would easily gain from some automation. *)
iIntros (α1 α2 [EQ1 EQ2] qmax qL) "HL".
iDestruct (EQ1 with "HL") as "#EQ1'".
iDestruct (EQ2 with "HL") as "#EQ2'".
iIntros "!> #HE".
iDestruct ("EQ1'" with "HE") as %EQ1'.
iDestruct ("EQ2'" with "HE") as %EQ2'.
iSplit; [|iSplit; iIntros "!> * H"]; simpl.
- iPureIntro; congruence.
- solve_proper_prepare.
iDestruct "H" as "[$ H]".
iDestruct "H" as (γ) "H".
iExists γ; iDestruct "H" as "[H $]".
by rewrite (lft_incl_syn_anti_symm _ _ EQ1' EQ2').
- iDestruct "H" as (γ) "H".
iExists γ; iDestruct "H" as "[H $]".
by rewrite (lft_incl_syn_anti_symm _ _ EQ1' EQ2').
Qed.
Global Instance ghosttoken_mono'
E L : Proper (lctx_lft_eq E L ==> eqtype E L) ghosttoken.
Proof.
intros.
split; by apply ghosttoken_mono.
Qed.
(** β is the total sharing lifetime of the GhostCell.
(In the proofs below, β is often something else!) *)
Definition ghostcell_inv tid l β α ty : iProp Σ :=
( st',
match st' with
| None => &{β}(l ↦∗: ty.(ty_own) tid) (* ghostcell is not currently being accessed *)
| Some rw => γ, lft_meta α γ
match rw with
| true => (* ghostcell is currently being write-accessed *)
(* The κ' will be set to the lifetime β in borrow_mut, the lifetime
for which we exclusively have the borrow token. If we own the
ghost token fragment (in any state), then at any time either κ'
still runs (so we get a contradiction when opening the borrow
here) or we can run the inheritance to get back the full
borrow. *)
κ', &{κ'} (own γ (ghosttoken_st_to_R None))
([κ'] ={lftN}=∗ &{β} (l ↦∗: ty_own ty tid))
| false => (* ghostcell is currently being read-accessed *)
(* The κ' will be set to the total lifetime for which the token is
shared. If we own the ghost token fragment in [None] state, then
we can deduce κ' is dead and we get back the full borrow. If we
own it in [Some] state we know the κ' is equal to the one in our
token, which outlives the lifetime of the token borrow that we got
(that is ensures in the GhostToken sharing interpretation), which
means it lives long enough to return it to the caller at that
lifetime. *)
κ', &frac{κ'} (λ q', own γ (ghosttoken_st_to_R (Some (κ', q'))))
([κ'] ={lftN}=∗ &{β}(l ↦∗: ty.(ty_own) tid))
ty.(ty_shr) (β κ') tid l
end
end)%I.
Global Instance ghostcell_inv_type_ne n tid l β α :
Proper (type_dist2 n ==> dist n) (ghostcell_inv tid l β α).
Proof.
solve_proper_core
ltac:(fun _ => exact: type_dist2_S || f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance ghostcell_inv_ne tid l β α : NonExpansive (ghostcell_inv tid l β α).
Proof.
intros n ???. by eapply ghostcell_inv_type_ne, type_dist_dist2.
Qed.
Lemma ghostcell_inv_proper E L κ1 κ2 ty1 ty2 qmax qL :
lctx_lft_eq E L κ1 κ2
eqtype E L ty1 ty2
llctx_interp_noend qmax L qL -∗ tid l β, (elctx_interp E -∗
ghostcell_inv tid l β κ1 ty1 -∗ ghostcell_inv tid l β κ2 ty2).
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic. *)
(* It would easily gain from some automation. *)
rewrite eqtype_unfold. iIntros ([Hlft1 Hlft2] Hty) "HL".
iDestruct (Hty with "HL") as "#Hty".
iDestruct (Hlft1 with "HL") as "#Hlft1".
iDestruct (Hlft2 with "HL") as "#Hlft2".
iIntros (tid l β) "!> #HE H".
iDestruct ("Hty" with "HE") as "(% & #Hown & #Hshr)".
iDestruct ("Hlft1" with "HE") as %Hκ1.
iDestruct ("Hlft2" with "HE") as %Hκ2.
iAssert ( (&{β}(l ↦∗: ty_own ty1 tid) -∗
&{β}(l ↦∗: 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 (rw) "H"; iExists rw; destruct rw as [rw|]; iFrame "∗";
last by iApply "Hb"; last first.
iDestruct "H" as (γc) "(#Hsing&H)".
iExists γc.
iSplitR.
{ by rewrite (lft_incl_syn_anti_symm _ _ Hκ1 Hκ2). }
destruct rw.
{ iDestruct "H" as (κ') "H'".
iExists κ'; iDestruct "H'" as "($ & H')".
iIntros "Hκ'". iMod ("H'" with "Hκ'"). by iApply "Hb". }
iDestruct "H" as (κ') "(Hag&Hh&H)"; iExists κ'; iFrame "Hag".
iSplitL "Hh"; last by iApply "Hshr".
iIntros "Hν". iMod ("Hh" with "Hν") as "Hh".
by iApply "Hb".
Qed.
(* Idea:
Ghostcell is basically a refcell/rwlock. Main difference being that current r/w state is
not directly tied to any physical state (for ty_shr); in other words, you can't really do
anything with just a ghostcell. The ghosttoken is responsible for tracking the r/w state of any
ghostcells that are currently open. A share of a ghosttoken always tracks the exact fraction
that it is sharing. It may not actually be necessary for the ghosttoken to track these things
explicitly, since a read borrow is the same as a regular borrow.
Key point: you don't really need to prove you can go from a borrow back to the original
resource under normal circumstances. For atomic borrows you just need to show that the
thing inside the atomic borrow holds initially (which can just be the original borrow you
got plus some ownership of ghost state, if you want). To *update* an atomic borrow, you
just have to show you can go back to the original resource *or* the token is dead.
To update an original borrow (and not change the mask) you need to show that the token is
alive initially, and that the resource to which you want to update it can move back to P
with the empty mask and no access to the token you povided. bor_acc_atomic and friends do
other stuff.
(when the lifetime is dead it can
be removed from the track set).
*)
Program Definition ghostcell (α : lft) (ty : type) :=
tc_opaque
{| ty_size := ty.(ty_size);
ty_own tid vl := ty.(ty_own) tid vl;
ty_shr κ tid l :=
( β, κ β &at{β, ghostcellN}(ghostcell_inv tid l β α ty))%I |}.
Next Obligation.
iIntros (????) "H". by rewrite ty_size_eq.
Qed.
Next Obligation.
iIntros (α ty E κ l tid q ?) "#LFT Hvl [Htok Htok']".
iAssert ((q / 2).[κ] ghostcell_inv tid l κ α ty)%I with "[> -Htok]"
as "[$ HQ]"; last first.
{ iFrame "Htok".
iExists κ.
iSplitR; first by iApply lft_incl_refl.
iMod (bor_create _ κ with "LFT HQ") as "[HQ HQ']"; first done.
iApply bor_share; first solve_ndisj.
iFrame "∗". }
iFrame. iExists None. by iFrame.
Qed.
Next Obligation.
iIntros (??????) "#Hκ H". iDestruct "H" as (β) "H".
iExists β; iDestruct "H" as "[? $]".
by iApply lft_incl_trans.
Qed.
Global Instance ghostcell_wf α `{!TyWf ty} : TyWf (ghostcell α ty) :=
{ ty_lfts := α::(ty_lfts ty); ty_wf_E := ty_wf_E ty }.
Global Instance ghostcell_type_ne α : TypeNonExpansive (ghostcell α).
Proof.
constructor;
solve_proper_core ltac:(
fun _ => exact: type_dist2_S || (eapply ghostcell_inv_type_ne; try reflexivity) ||
(eapply ghostcell_inv_proj_type_ne; try reflexivity) ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance ghostcell_ne α : NonExpansive (ghostcell α).
Proof.
constructor; solve_proper_core ltac:(
fun _ => (eapply ty_size_ne; try reflexivity) || f_equiv).
Qed.
Global Instance ghostcell_mono E L : Proper (lctx_lft_eq E L ==> eqtype E L ==> subtype E L)
(ghostcell).
Proof.
(* TODO : this proof is essentially [solve_proper], but within the logic.
It would easily gain from some automation. *)
iIntros (κ κ' [EQ1 EQ1'] ty1 ty2 EQ2 qmax qL) "HL". generalize EQ2. rewrite eqtype_unfold=>EQ'2.
iDestruct (EQ1 with "HL") as "#EQ1".
iDestruct (EQ1' with "HL") as "#EQ1'".
iDestruct (EQ'2 with "HL") as "#EQ'".
iDestruct (ghostcell_inv_proper _ _ κ κ' with "HL") as "#Hty1ty2"; [done|done|].
iDestruct (ghostcell_inv_proper _ _ κ' κ with "HL") as "#Hty2ty1"; [done|by symmetry|].
iIntros "!> #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)".
iSplit; [|iSplit; iIntros "!> * H"]; simpl.
- iPureIntro; congruence.
- by iApply "Hown".
- iDestruct "H" as (a) "H".
iExists a; iDestruct "H" as "[$ H]".
iApply at_bor_iff; last done.
iNext; iModIntro; iSplit; iIntros "H"; by [iApply "Hty1ty2"|iApply "Hty2ty1"].
Qed.
Lemma ghostcell_mono' E L κ1 κ2 ty1 ty2 :
lctx_lft_eq E L κ1 κ2
eqtype E L ty1 ty2
subtype E L (ghostcell κ1 ty1) (ghostcell κ2 ty2).
Proof. intros. by eapply ghostcell_mono. Qed.
Global Instance ghostcell_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) ghostcell.
Proof. by split; apply ghostcell_mono. Qed.
Lemma ghostcell_proper' E L κ1 κ2 ty1 ty2 :
lctx_lft_eq E L κ1 κ2
eqtype E L ty1 ty2
eqtype E L (ghostcell κ1 ty1) (ghostcell κ2 ty2).
Proof. intros. by apply ghostcell_proper. Qed.
Global Instance ghostcell_send α ty :
Send ty Send (ghostcell α ty).
Proof. by unfold ghostcell, Send. Qed.
Global Instance ghostcell_sync α ty :
Send ty Sync ty Sync (ghostcell α ty).
Proof.
intros Hsend Hsync ????.
solve_proper_prepare.
do 3 f_equiv.
unfold ghostcell_inv.
rewrite at_bor_proper; first done.
do 7 f_equiv.
- do 9 f_equiv. iApply send_change_tid'.
- do 4 f_equiv; last by iApply sync_change_tid'.
do 6 f_equiv. iApply send_change_tid'.
- iApply send_change_tid'.
Qed.
Definition ghosttoken_new (call_once : val) (R_F : type) : val :=
fn: ["f"] :=
let: "call_once" := call_once in
let: "n" := new [ #0] in
letcall: "r" := "call_once" ["f";"n"]%E in
letalloc: "r'" <-{ R_F.(ty_size)} !"r" in
delete [ #R_F.(ty_size); "r"];;
return: ["r'"].
Lemma ghosttoken_new_type F R_F call_once `{!TyWf F, !TyWf R_F} :
typed_val call_once (fn( α, ; F, ghosttoken α) R_F)
typed_val (ghosttoken_new call_once R_F) (fn(; F) R_F).
Proof.
iIntros (Hf E L). iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret args). inv_vec args=>env. simpl_subst.
iApply type_let; [apply Hf|solve_typing|]. iIntros (f); simpl_subst.
iApply type_new_subtype; [solve_typing..|].
iIntros (n).
simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hc (Hn & Hf & Henv & _)".
iMod (lctx_lft_alive_tok (ϝ) with "HE HL")
as (qϝf) "(Hϝf & HL & Hclosef)"; [solve_typing..|].
iMod (own_alloc (ghosttoken_st_to_R None)) as (γ) "Hown"; first done.
iMod (lft_create_meta γ with "LFT") as (α) "[#Hidx [Htok #Hα]]"; first done.
wp_let.
rewrite !tctx_hasty_val.
iDestruct (lft_intersect_acc with "Htok Hϝf") as (?) "[Hαϝ Hcloseα]".
iApply (type_call_iris _ [α; ϝ] (α) [_;_] _ _ _ _
with "LFT HE Hna [Hαϝ] Hf [Hn Henv Hown]"); try solve_typing.
+ by rewrite /= (right_id static).
+ rewrite /= (right_id emp%I) !tctx_hasty_val ownptr_uninit_own.
iFrame "Henv".
rewrite ownptr_own.
iDestruct "Hn" as (l vl) "(% & Hl & Hfree)".
iExists l, vl.
iSplit; first done.
simpl_subst.
iFrame "∗#".
by refine (match vl with Vector.nil => _ end).
+ iIntros (r) "Hna Hf Hown".
simpl_subst.
iDestruct ("Hcloseα" with "[Hf]") as "[Htok Hf]"; [by rewrite (right_id static)|].
iMod ("Hclosef" with "Hf HL") as "HL".
wp_let.
iApply (type_type _ _ _ [ r box R_F ] with "[] LFT HE Hna HL Hc"); try solve_typing;
last by rewrite !tctx_interp_singleton tctx_hasty_val.
iApply type_letalloc_n; [solve_typing..|].
iIntros (r').
simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition ghostcell_new : val :=
fn: ["n"] :=
return: ["n"].
Lemma ghostcell_new_type `{!TyWf ty} :
typed_val ghostcell_new (fn( α, ; ty) (ghostcell α ty))%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret args). inv_vec args=>n. simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk Hn".
rewrite tctx_interp_singleton tctx_hasty_val.
iAssert (ty_own (box (ghostcell α ty)) tid [n])%I with "[Hn]" as "Hn".
{ rewrite !ownptr_own.
iDestruct "Hn" as (l vl) "(% & Hl & Hown & Hfree)".
by iExists l, vl; simpl; iFrame "Hl Hown Hfree". }
rewrite -tctx_hasty_val -tctx_interp_singleton.
iApply (type_type with "[] LFT HE Hna HL Hk Hn").
iApply type_jump; solve_typing.
Qed.
Definition ghostcell_borrow : val :=
fn: ["c";"s"] :=
(* Skips needed for laters *)
Skip ;; Skip ;;
return: ["c"].
Lemma ghostcell_share E β κ' tid lc γ κ α ty :
lftN E
(lft_ctx -∗
β κ -∗
β κ' -∗
&frac{κ'} (λ q : Qp, own γ (ghosttoken_st_to_R (Some (κ', q)))) -∗
().[β] -∗
lft_meta α γ -∗
&{κ} (lc ↦∗: ty_own ty tid) ={E}=∗
ghostcell_inv tid lc κ α ty
ty_shr ty β tid lc ().[β]).
Proof.
iIntros (HE) "#LFT #Hκ #Hβκ' #Hfractok [Hβ1 Hβ2] Hsing Hst'".
iMod (fupd_mask_subseteq (lftN)) as "Hclose"; first solve_ndisj.
iMod (lft_incl_acc with "Hκ Hβ1") as (q''2) "[Hq'' Hcloseq'']"; first solve_ndisj.
iMod (lft_incl_acc with "Hβκ' Hβ2") as (q''2_2) "[Hq''_2 Hcloseq''_2]";
first solve_ndisj.
iMod (rebor _ _ (κ κ') (lc ↦∗: ty_own ty tid)%I with "LFT [] [Hst']")
as "[Hvl Hh]"; [done|iApply lft_intersect_incl_l|done|].
iDestruct (lft_intersect_acc with "Hq'' Hq''_2") as (q''3) "[Hq'' Hcloseq''2]".
iMod (ty_share with "LFT Hvl Hq''") as "[#Hshr Hq'']"; first solve_ndisj.
iDestruct ("Hcloseq''2" with "Hq''") as "[Hq'' Hq''_2]".
iMod ("Hcloseq''" with "Hq''") as "$".
iMod ("Hcloseq''_2" with "Hq''_2") as "$".
iDestruct (ty_shr_mono with "[] Hshr") as "$"; first by iApply lft_incl_glb.
iMod "Hclose" as "_".
iExists (Some false), γ; iFrame "Hsing".
iExists κ'; iFrame "#∗". iIntros "!> !> Htok".
iApply "Hh". iApply lft_dead_or. by iRight.
Qed.
Lemma ghostcell_borrow_type `{!TyWf ty} :
typed_val ghostcell_borrow
(fn( '(α, β), ; &shr{β} (ghostcell α ty), &shr{β} (ghosttoken α))
&shr{β} ty)%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret args). inv_vec args=>c s. simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL HC (Hc & Hs & _)".
rewrite !tctx_hasty_val.
iMod (lctx_lft_alive_tok β with "HE HL")
as () "([Hβ1 [Hβ2 Hβ3]] & HL & Hclose)"; [solve_typing..|].
iAssert ( |={}[⊤∖↑ghostcellN]▷=> (ty_own (box (&shr{β} ty)) tid [c] ().[β]))%I
with "[Hβ1 Hβ2 Hβ3 Hs Hc]"as "Hc".
{ iClear "HE".
rewrite (ownptr_own (_ (_ (ghostcell _ _))))
(ownptr_own (_ (&shr{β} ty)))%T.
rewrite ownptr_own.
iDestruct "Hs" as (l' vl') "(_ & _ & Hats & _)".
iDestruct "Hc" as (l vl) "(% & Hl & Hatc & Hfree)".
subst.
inv_vec vl'=>ls'.
iAssert _ with "Hats" as "#Hats'".
iDestruct (shr_is_ptr with "Hats'") as (ls) "> H". iDestruct "H" as %H.
inversion H; subst; iClear (H) "Hats'".
inv_vec vl=>lc'.
iAssert _ with "Hatc" as "#Hatc'".
iDestruct (shr_is_ptr with "Hatc'") as (lc) "> H". iDestruct "H" as %H.
inversion H; simpl_subst; iClear (H) "Hatc'".
iDestruct "Hats" as (γ) "[Hsing Hats]".
iDestruct "Hats" as (κ') "[#Hβκ' #Hats]".
iDestruct "Hatc" as (κ) "[#Hκ Hatc]".
iDestruct (at_bor_shorten with "Hκ Hatc") as "Hatc'".
iIntros "!>".
iMod (at_bor_acc_tok with "LFT Hatc' Hβ1") as "[Hcell Hclosec]"; [solve_ndisj..|].
iDestruct "Hcell" as (st') "Hst'".
destruct st' as [rw|].
{ iDestruct "Hst'" as (γ') "(>Hsing' & Hst')".
iDestruct (lft_meta_agree with "Hsing Hsing'") as %<-.
iClear "Hsing'".
iIntros "!> !>".
iMod (fupd_mask_subseteq (lftN)) as "Hclose"; first solve_ndisj.
iDestruct (frac_bor_shorten with "Hβκ' Hats") as "Hats'".
iMod (frac_bor_acc with "LFT Hats' Hβ2") as (q') "[>Hset Hcloses] {Hats'}"; [solve_ndisj..|].
destruct rw as[|].
{ iDestruct "Hst'" as (κ'0) "(Hbor & Hdead)".
iMod (bor_acc_atomic with "LFT Hbor") as "[[> Hst' Hcloseb]|[H† Hcloseb]]"; first done.
- iCombine "Hset Hst'" gives %[].
- iMod "Hcloseb" as "_".
iMod ("Hdead" with "H†") as "Hst'".
iMod ("Hcloses" with "Hset") as "Hβ2".
iMod "Hclose" as "_".
iMod (ghostcell_share with "LFT Hκ Hβκ' Hats Hβ3 Hsing Hst'")
as "(Hinv' & Hshr & Hβ3)"; first solve_ndisj.
iMod ("Hclosec" with "Hinv'") as "$".
iFrame "Hβ3 Hβ2".
iModIntro.
iExists l, (Vector.cons #lc Vector.nil).
by iFrame "Hshr Hl Hfree". }
iDestruct "Hst'" as (κ'0) "(Hκ'0bor & Hνκ & #Hshrκ)".
iMod (frac_bor_atomic_acc with "LFT Hκ'0bor")
as "[Hsucc|[Hκ'0† >_]]"; first done; last first.
{ iMod ("Hνκ" with "Hκ'0†") as "Hst'".
iClear "Hshrκ".
iMod ("Hcloses" with "Hset") as "Hβ2".
iMod "Hclose" as "_".
iMod (ghostcell_share with "LFT Hκ Hβκ' Hats Hβ3 Hsing Hst'")
as "(Hinv' & Hshr & Hβ3)"; first solve_ndisj.
iMod ("Hclosec" with "Hinv'") as "$".
iFrame "Hβ3 Hβ2".
iModIntro.
iExists l, (Vector.cons #lc Vector.nil).
by iFrame "Hshr Hl Hfree". }
iDestruct "Hsucc" as (q'0) "[>Hownκ'0 Hcloseκ'0]".
iCombine "Hset Hownκ'0" gives %Hvalidκ'0.
(* Argue that we have the same κ' here as the already existing sharing protocol. *)
assert (Hκ'κ'0 : κ' = κ'0).
{ move: Hvalidκ'0. rewrite /ghosttoken_st_to_R /=.
rewrite -Cinr_op /valid /cmra_valid /=.
rewrite pair_valid /=.
intros [?%to_agree_op_inv_L _]. done. }
subst κ'0.
iMod ("Hcloseκ'0" with "Hownκ'0") as "_".
iDestruct (ty_shr_mono with "[] Hshrκ") as "Hshrβ"; first by iApply lft_incl_glb.
iMod ("Hcloses" with "Hset") as "Hβ2".
iMod "Hclose" as "_".
iMod ("Hclosec" with "[Hνκ Hsing]") as "$".
{ iNext.
iExists (Some false).
iExists γ.
iFrame "Hsing".
iExists κ'.
by iFrame "Hνκ Hshrκ". }
iFrame "Hβ3 Hβ2".
iExists l, (Vector.cons #lc Vector.nil).
by iFrame "Hshrβ Hl Hfree". }
iIntros "!> !>".
iMod (fupd_mask_subseteq (lftN)) as "Hclose"; first solve_ndisj.
iMod "Hclose" as "_".
iMod (ghostcell_share with "LFT Hκ Hβκ' Hats Hβ3 Hsing Hst'")
as "(Hinv' & Hshr & Hβ3)"; first solve_ndisj.
iMod ("Hclosec" with "Hinv'") as "$".
iFrame "Hβ3 Hβ2".
iModIntro.
iExists l, (Vector.cons #lc Vector.nil).
by iFrame "Hshr Hl Hfree". }
wp_let.
iApply lifting.wp_pure_step_fupd; first done.
iMod "Hc".
iIntros "!> !>".
iMod "Hc".
iDestruct "Hc" as "[Hshr Hβ]".
iMod ("Hclose" with "Hβ HL") as "HL".
iIntros "!> _".
do 2 wp_let.
iApply (type_type _ _ _ [c box (&shr{β} ty)]
with "[] LFT HE Hna HL HC [Hshr]"); last first.
{ by rewrite tctx_interp_singleton tctx_hasty_val. }
iApply type_jump; solve_typing.
Qed.
Definition ghostcell_borrow_mut : val :=
fn: ["c";"s"] :=
(* Skips needed for laters *)
Skip ;; Skip ;;
return: ["c"].
Lemma ghostcell_borrow_mut_type `{!TyWf ty} :
typed_val ghostcell_borrow_mut
(fn( '(α, β), ; &shr{β} (ghostcell α ty), &uniq{β} (ghosttoken α))
&uniq{β} ty)%T.
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret args). inv_vec args=>c s. simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL HC (Hc & Hs & _)".
rewrite !tctx_hasty_val.
iMod (lctx_lft_alive_tok β with "HE HL") as () "([Hβ1 [Hβ2 Hβ3]] & HL & Hclose)";
[solve_typing..|].
iAssert ( |={}[⊤∖↑ghostcellN]▷=> (ty_own (box (&uniq{β} ty)) tid [c] ().[β]))%I
with "[Hβ1 Hβ2 Hβ3 Hs Hc]"as "Hc".
{ iClear "HE".
rewrite (ownptr_own (_ (_ (ghostcell _ _))))
(ownptr_own (_ (&uniq{β} ty)))%T.
rewrite ownptr_own.
iDestruct "Hs" as (l' vl') "(_ & _ & Hats & _)".
iDestruct "Hc" as (l vl) "(% & Hl & Hatc & Hfree)".
subst.
inv_vec vl'=>ls'.
destruct ls' as [[|ls|]|]; try by iDestruct "Hats" as "> []".
inv_vec vl=>lc'.
iAssert _ with "Hatc" as "#Hatc'".
iDestruct (shr_is_ptr with "Hatc'") as (lc) "> H". iDestruct "H" as %H.
inversion H; simpl_subst; iClear (H) "Hatc'".
iDestruct "Hatc" as (κ) "[#Hκ Hatc]".
iDestruct (at_bor_shorten with "Hκ Hatc") as "Hatc'".
iIntros "!>".
iMod (at_bor_acc_tok with "LFT Hatc' Hβ1") as "[Hcell Hclosec]"; [solve_ndisj..|].
iDestruct "Hcell" as (st') "Hst'".
iMod (bor_acc_strong with "LFT Hats Hβ2") as (κ'1) "[#Hκ'κ'1 [Hats Hcloses]]"; first solve_ndisj.
iDestruct "Hats" as (vl) "(> Hls & > % & >Hats)".
subst vl.
iDestruct "Hats" as (γ) "[#Hsing Hset]".
iAssert ((|={ ghostcellN}▷=>
own γ (Cinl (Excl ()))
(&{κ} (lc ↦∗: ty_own ty tid))) (* ∨
(_ ∗ ▷ ghostcell_inv tid lc κ α ty ∗ |={⊤}▷=> ▷ ▷ False) *) )%I
with "[Hst' Hset]" as "Hown".
{ destruct st' as [rw|]; last first.
{ eauto with iFrame. }
iDestruct "Hst'" as (γ') "(>Hsing' & Hst')".
iDestruct (lft_meta_agree with "Hsing Hsing'") as %<-.
iClear "Hsing'".
iIntros "!> !>".
iMod (fupd_mask_subseteq (lftN)) as "Hclose"; first solve_ndisj.
destruct rw as [|].
{ iDestruct "Hst'" as (κ'0) "(Hbor & Hdead)".
iMod (bor_acc_atomic with "LFT Hbor") as "[[> Hst' Hcloseb]|[H† Hcloseb]]";
first done.
{ iCombine "Hset Hst'" gives %[]. }
iMod "Hcloseb" as "_".
iMod ("Hdead" with "H†") as "Hst'".
iFrame. done. }
iDestruct "Hst'" as (κ') "(Hst' & Hνκ & #Hshrκ)".
iDestruct "Hst'" as "Hκ'0bor".
iClear "Hshrκ".
iMod (frac_bor_atomic_acc with "LFT Hκ'0bor") as "[Hsucc|[H Hcloseb]]"; first done.
{ iDestruct "Hsucc" as (q) "[>Htok _]".
iCombine "Hset Htok" gives %[]. }
iMod "Hcloseb" as "_".
iMod ("Hνκ" with "H") as "$".
iMod "Hclose".
by iFrame. }
iMod "Hown".
iIntros "!> !>".
iMod "Hown" as "[Hset Hown]".
iMod ("Hcloses" $! (own γ (ghosttoken_st_to_R None)) with "[Hls] Hset") as "[Hset Hβ]".
{ iIntros "!> >Hset _". iExists []. eauto 10 with iFrame. }
iMod (fupd_mask_subseteq (lftN)) as "Hclose"; first solve_ndisj.
iMod (rebor with "LFT Hκ Hown") as "[Hbor Hcloseβ]"; first solve_ndisj.
iMod "Hclose" as "_".
iMod ("Hclosec" with "[Hcloseβ Hset]") as "$".
{ iExists (Some true), γ.
iFrame "Hsing". iExists β; iFrame. iNext.
iApply bor_shorten; last done. done.
}
iFrame "Hβ3 Hβ".
iExists l, (Vector.cons #lc Vector.nil).
iAssert (⌜#l = #l)%I as "$"; first done.
by iFrame "Hl Hfree Hbor". }
wp_let.
iApply lifting.wp_pure_step_fupd; first done.
iMod "Hc".
iIntros "!> !>".
iMod "Hc".
iDestruct "Hc" as "[Hshr Hβ]".
iMod ("Hclose" with "Hβ HL") as "HL".
iIntros "!> _".
do 2 wp_let.
iApply (type_type _ _ _ [c box (&uniq{β} ty)]
with "[] LFT HE Hna HL HC [Hshr]"); last first.
{ by rewrite tctx_interp_singleton tctx_hasty_val. }
iApply type_jump; solve_typing.
Qed.
Definition ghostcell_as_mut : val :=
fn: ["c"] :=
return: ["c"].
Lemma ghostcell_as_mut_type `{!TyWf ty} :
typed_val ghostcell_as_mut (fn( '(α, β), ; &uniq{β} (ghostcell α ty))
&uniq{β} ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret args). inv_vec args=>c. simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk Hc".
rewrite tctx_interp_singleton tctx_hasty_val.
iAssert (ty_own (box (&uniq{β} ty)) tid [c])%I with "[Hc]" as "Hc".
{ rewrite !ownptr_own.
iDestruct "Hc" as (l vl) "(% & Hl & Hown & Hfree)".
iExists l, vl; rewrite /ty_own /=. by iFrame "Hl Hown Hfree". }
rewrite -tctx_hasty_val -tctx_interp_singleton.
iApply (type_type with "[] LFT HE Hna HL Hk Hc").
iApply type_jump; solve_typing.
Qed.
End ghostcell.
Global Hint Resolve ghostcell_mono' ghostcell_proper' : lrust_typing.
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 joinN := lrustN .@ "join".
Section join.
Context `{!typeGS Σ, !spawnG Σ}.
(* This model is very far from rayon::join, which uses a persistent
work-stealing thread-pool. Still, the important bits are right:
One of the closures is executed in another thread, and the
closures can refer to on-stack data (no 'static' bound). *)
Definition join (call_once_A call_once_B : val) (R_A R_B : type) : val :=
fn: ["fA"; "fB"] :=
let: "call_once_A" := call_once_A in
let: "call_once_B" := call_once_B in
let: "join" := spawn [λ: ["c"],
letcall: "r" := "call_once_A" ["fA"]%E in
finish ["c"; "r"]] in
letcall: "retB" := "call_once_B" ["fB"]%E in
let: "retA" := spawn.join ["join"] in
(* Put the results in a pair. *)
let: "ret" := new [ #(R_A.(ty_size) + R_B.(ty_size)) ] in
"ret" + #0 <-{R_A.(ty_size)} !"retA";;
"ret" + #R_A.(ty_size) <-{R_B.(ty_size)} !"retB";;
delete [ #R_A.(ty_size); "retA"] ;; delete [ #R_B.(ty_size); "retB"] ;;
return: ["ret"].
Lemma join_type A B R_A R_B call_once_A call_once_B
`{!TyWf A, !TyWf B, !TyWf R_A, !TyWf R_B}
`(!Send A, !Send B, !Send R_A, !Send R_B) :
(* A : FnOnce() -> R_A, as witnessed by the impl call_once_A *)
typed_val call_once_A (fn(; A) R_A)
(* B : FnOnce() -> R_B, as witnessed by the impl call_once_B *)
typed_val call_once_B (fn(; B) R_B)
typed_val (join call_once_A call_once_B R_A R_B) (fn(; A, B) Π[R_A; R_B]).
Proof using Type*.
intros HfA HfB E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>envA envB. simpl_subst.
iApply type_let; [apply HfA|solve_typing|]. iIntros (fA); simpl_subst.
iApply type_let; [apply HfB|solve_typing|]. iIntros (fB); simpl_subst.
(* Drop to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk (HfB & HfA & HenvA & HenvB & _)".
iMod (lctx_lft_alive_tok ϝ with "HE HL") as (qϝ1) "(Hϝ1 & HL & Hclose1)";
[solve_typing..|].
(* FIXME: using wp_apply here breaks calling solve_to_val. *)
wp_bind (spawn _).
iApply ((spawn_spec joinN (λ v, (box R_A).(ty_own) tid [v] (qϝ1).[ϝ])%I) with "[HfA HenvA Hϝ1]").
{ (* The new thread. *)
simpl_subst. iIntros (c) "Hfin". iMod na_alloc as (tid') "Htl". wp_let.
wp_let. unlock. rewrite !tctx_hasty_val.
iApply (type_call_iris _ [ϝ] () [_] with "LFT HE Htl [Hϝ1] HfA [HenvA]").
- rewrite outlives_product. solve_typing.
- by rewrite /= (right_id static).
- by rewrite big_sepL_singleton tctx_hasty_val send_change_tid.
- iIntros (r) "Htl Hϝ1 Hret".
wp_rec. iApply (finish_spec with "[$Hfin Hret Hϝ1]"); last auto.
rewrite right_id. iFrame. by iApply @send_change_tid. }
iNext. iIntros (c) "Hjoin". wp_let. wp_let.
iMod (lctx_lft_alive_tok_noend ϝ with "HE HL") as (qϝ2) "(Hϝ2 & HL & Hclose2)";
[solve_typing..|].
rewrite !tctx_hasty_val.
iApply (type_call_iris _ [ϝ] () [_] with "LFT HE Hna [Hϝ2] HfB [HenvB]").
{ rewrite outlives_product. solve_typing. }
{ by rewrite /= (right_id static). }
{ by rewrite big_sepL_singleton tctx_hasty_val. }
(* The return continuation after calling fB in the main thread. *)
iIntros (retB) "Hna Hϝ2 HretB". rewrite /= (right_id static).
iMod ("Hclose2" with "Hϝ2 HL") as "HL". wp_rec.
wp_apply (join_spec with "Hjoin"). iIntros (retA) "[HretA Hϝ1]".
iMod ("Hclose1" with "Hϝ1 HL") as "HL". wp_let.
(* Switch back to type system mode. *)
iApply (type_type _ _ _ [ retA box R_A; retB box R_B ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply (type_new_subtype (Π[uninit R_A.(ty_size); uninit R_B.(ty_size)]));
(* FIXME: solve_typing should handle this without any aid. *)
rewrite ?Nat2Z.inj_add; [solve_typing..|].
iIntros (r); simpl_subst.
iApply (type_memcpy R_A); [solve_typing..|].
iApply (type_memcpy R_B); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End join.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From lrust.lang.lib Require Import memcpy lock.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Export type.
From lrust.typing Require Import typing option.
From iris.prelude Require Import options.
Definition mutexN := lrustN .@ "mutex".
Section mutex.
Context `{!typeGS Σ}.
(*
pub struct Mutex<T: ?Sized> {
// Note that this mutex is in a *box*, not inlined into the struct itself.
// Once a native mutex has been used once, its address can never change (it
// can't be moved). This mutex type can be safely moved at any time, so to
// ensure that the native mutex is used correctly we box the inner lock to
// give it a constant address.
inner: Box<sys::Mutex>,
poison: poison::Flag,
data: UnsafeCell<T>,
}
*)
Program Definition mutex (ty : type) :=
{| ty_size := 1 + ty.(ty_size);
ty_own tid vl :=
match vl return _ with
| #(LitInt z) :: vl' =>
⌜∃ b, z = Z.b2z b ty.(ty_own) tid vl'
| _ => False end;
ty_shr κ tid l := κ', κ κ'
&at{κ, mutexN} (lock_proto l (&{κ'}((l + 1) ↦∗: ty.(ty_own) tid)))
|}%I.
Next Obligation.
iIntros (??[|[[]|]]); try iIntros "[]". rewrite ty_size_eq.
iIntros "[_ %] !% /=". congruence.
Qed.
Next Obligation.
iIntros (ty E κ l tid q ?) "#LFT Hbor Htok".
iMod (bor_acc_cons with "LFT Hbor Htok") as "[H Hclose]"; first done.
iDestruct "H" as ([|[[| |n]|]vl]) "[H↦ H]"; try iDestruct "H" as ">[]".
rewrite heap_pointsto_vec_cons. iDestruct "H↦" as ">[Hl H↦]".
iDestruct "H" as "[>EQ Hown]". iDestruct "EQ" as %[b ->].
(* We need to turn the ohne borrow into two, so we close it -- and then
we open one of them again. *)
iMod ("Hclose" $! (( b, l #(Z.b2z b)) ((l + 1) ↦∗: ty_own ty tid))%I
with "[] [Hl H↦ Hown]") as "[Hbor Htok]".
{ clear. iNext. iIntros "[Hl Hown] !>". iNext. iDestruct "Hl" as (b) "Hl".
iDestruct "Hown" as (vl) "[H↦ Hown]". iExists (#(Z.b2z b) :: vl).
rewrite heap_pointsto_vec_cons. iFrame. iPureIntro. eauto. }
{ iNext. iSplitL "Hl"; first by eauto.
iExists vl. iFrame. }
clear b vl. iMod (bor_sep with "LFT Hbor") as "[Hl Hbor]"; first done.
iMod (bor_acc_cons with "LFT Hl Htok") as "[>Hl Hclose]"; first done.
iDestruct "Hl" as (b) "Hl".
iMod (lock_proto_create with "Hl [Hbor]") as "Hproto".
{ destruct b; last by iExact "Hbor". done. }
iMod ("Hclose" with "[] Hproto") as "[Hl Htok]".
{ clear b. iIntros "!> Hproto !>".
iDestruct (lock_proto_destroy with "Hproto") as (b) "[Hl _]".
eauto 10 with iFrame. }
iFrame "Htok". iExists κ.
iMod (bor_share with "Hl") as "$"; [solve_ndisj..|].
iApply lft_incl_refl.
Qed.
Next Obligation.
iIntros (ty κ κ' tid l) "#Hincl H".
iDestruct "H" as (κ'') "(#Hlft & #Hlck)".
iExists _. by iSplit; [iApply lft_incl_trans|iApply at_bor_shorten].
Qed.
Global Instance mutex_wf ty `{!TyWf ty} : TyWf (mutex ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance mutex_type_ne : TypeNonExpansive mutex.
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance mutex_ne : NonExpansive mutex.
Proof.
constructor; solve_proper_core ltac:(fun _ => (eapply ty_size_ne; try reflexivity) || f_equiv).
Qed.
Global Instance mutex_mono E L : Proper (eqtype E L ==> subtype E L) mutex.
Proof.
move=>ty1 ty2 /eqtype_unfold EQ. iIntros (??) "HL".
iDestruct (EQ with "HL") as "#EQ". iIntros "!> #HE". clear EQ.
iDestruct ("EQ" with "HE") as "(% & #Howni & _) {EQ}".
iSplit; last iSplit.
- simpl. iPureIntro. f_equiv. done.
- iIntros "!> %tid %vl Hvl". destruct vl as [|[[| |n]|]vl]; try done.
simpl. iDestruct "Hvl" as "[$ Hvl]". by iApply "Howni".
- iIntros "!> %κ %tid %l Hshr". iDestruct "Hshr" as (κ') "[Hincl Hshr]".
iExists _. iFrame "Hincl". iApply (at_bor_iff with "[] Hshr"). iNext.
iApply lock_proto_iff_proper. iApply bor_iff_proper. iNext.
iApply heap_pointsto_pred_iff_proper.
iModIntro; iIntros; iSplit; iIntros; by iApply "Howni".
Qed.
Global Instance mutex_proper E L :
Proper (eqtype E L ==> eqtype E L) mutex.
Proof. by split; apply mutex_mono. Qed.
Global Instance mutex_send ty :
Send ty Send (mutex ty).
Proof.
iIntros (??? [|[[| |n]|]vl]); try done. iIntros "[$ Hvl]".
by iApply send_change_tid.
Qed.
Global Instance mutex_sync ty :
Send ty Sync (mutex ty).
Proof.
iIntros (???? l) "Hshr". iDestruct "Hshr" as (κ') "[Hincl Hshr]".
iExists _. iFrame "Hincl". iApply (at_bor_iff with "[] Hshr"). iNext.
iApply lock_proto_iff_proper. iApply bor_iff_proper. iNext.
iApply heap_pointsto_pred_iff_proper.
iModIntro; iIntros; iSplit; iIntros; by iApply send_change_tid.
Qed.
End mutex.
Section code.
Context `{!typeGS Σ}.
Definition mutex_new ty : val :=
fn: ["x"] :=
let: "m" := new [ #(mutex ty).(ty_size) ] in
"m" + #1 <-{ty.(ty_size)} !"x";;
mklock_unlocked ["m" + #0];;
delete [ #ty.(ty_size); "x"];; return: ["m"].
Lemma mutex_new_type ty `{!TyWf ty} :
typed_val (mutex_new ty) (fn(; ty) mutex ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
(* FIXME: It should be possible to infer the `S ty.(ty_size)` here.
This should be done in the @eq external hints added in lft_contexts.v. *)
iApply (type_new (S ty.(ty_size))); [solve_typing..|]; iIntros (m); simpl_subst.
(* FIXME: The following should work. We could then go into Iris later.
iApply (type_memcpy ty); [solve_typing..|]. *)
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hm [Hx _]]".
rewrite !tctx_hasty_val /=.
iDestruct (ownptr_uninit_own with "Hm") as (lm vlm) "(% & Hm & Hm†)".
subst m. inv_vec vlm=>m vlm. simpl. iDestruct (heap_pointsto_vec_cons with "Hm") as "[Hm0 Hm]".
destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]".
iDestruct (heap_pointsto_ty_own with "Hx") as (vl) "[>Hx Hxown]".
(* All right, we are done preparing our context. Let's get going. *)
wp_op. wp_apply (wp_memcpy with "[$Hm $Hx]"); [by rewrite length_vec_to_list..|].
iIntros "[Hm Hx]". wp_seq. wp_op. rewrite shift_loc_0. wp_lam.
wp_write.
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ #lx box (uninit ty.(ty_size)); #lm box (mutex ty)]
with "[] LFT HE Hna HL Hk"); last first.
(* TODO: It would be nice to say [{#}] as the last spec pattern to clear the context in there. *)
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val' // tctx_hasty_val' //.
iFrame. iSplitR.
- simpl. by rewrite length_vec_to_list.
- iExists (#false :: vl). rewrite heap_pointsto_vec_cons. iFrame; eauto. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition mutex_into_inner ty : val :=
fn: ["m"] :=
let: "x" := new [ #ty.(ty_size) ] in
"x" <-{ty.(ty_size)} !("m" + #1);;
delete [ #(mutex ty).(ty_size); "m"];; return: ["x"].
Lemma mutex_into_inner_type ty `{!TyWf ty} :
typed_val (mutex_into_inner ty) (fn(; mutex ty) ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>m. simpl_subst.
iApply (type_new ty.(ty_size)); [solve_typing..|]; iIntros (x); simpl_subst.
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hm _]]".
rewrite !tctx_hasty_val /=.
iDestruct (ownptr_uninit_own with "Hx") as (lx vlx) "(% & Hx & Hx†)".
subst x. simpl.
destruct m as [[|lm|]|]; try done. iDestruct "Hm" as "[Hm Hm†]".
iDestruct (heap_pointsto_ty_own with "Hm") as (vlm) "[>Hm Hvlm]".
inv_vec vlm=>m vlm. destruct m as [[|m|]|]; try by iDestruct "Hvlm" as ">[]".
simpl. iDestruct (heap_pointsto_vec_cons with "Hm") as "[Hm0 Hm]".
iDestruct "Hvlm" as "[_ Hvlm]".
(* All right, we are done preparing our context. Let's get going. *)
wp_op. wp_apply (wp_memcpy with "[$Hx $Hm]"); [by rewrite length_vec_to_list..|].
(* FIXME: Swapping the order of $Hx and $Hm breaks. *)
iIntros "[Hx Hm]". wp_seq.
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ #lx box ty; #lm box (uninit (mutex ty).(ty_size))]
with "[] LFT HE Hna HL Hk"); last first.
(* TODO: It would be nice to say [{#}] as the last spec pattern to clear the context in there. *)
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val' // tctx_hasty_val' //.
iFrame. iExists (_ :: _). rewrite heap_pointsto_vec_cons. iFrame.
rewrite uninit_own. rewrite /= length_vec_to_list. eauto. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition mutex_get_mut : val :=
fn: ["m"] :=
let: "m'" := !"m" in
"m" <- ("m'" + #1);;
return: ["m"].
Lemma mutex_get_mut_type ty `{!TyWf ty} :
typed_val mutex_get_mut (fn( α, ; &uniq{α}(mutex ty)) &uniq{α} ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg); inv_vec arg=>m; simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (m'); simpl_subst.
(* Go to Iris *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hm [Hm' _]]".
rewrite !tctx_hasty_val [[m]]lock.
destruct m' as [[|lm'|]|]; try done. simpl.
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose1)";
[solve_typing..|].
iMod (bor_acc_cons with "LFT Hm' Hα") as "[Hm' Hclose2]"; first done.
wp_op. iDestruct "Hm'" as (vl) "[H↦ Hm']".
destruct vl as [|[[|m'|]|] vl]; try done. simpl.
iDestruct (heap_pointsto_vec_cons with "H↦") as "[H↦1 H↦2]".
iDestruct "Hm'" as "[% Hvl]".
iMod ("Hclose2" $! ((lm' + 1) ↦∗: ty_own ty tid)%I with "[H↦1] [H↦2 Hvl]") as "[Hbor Hα]".
{ iIntros "!> Hlm' !>". iNext. clear vl. iDestruct "Hlm'" as (vl) "[H↦ Hlm']".
iExists (_ :: _). rewrite heap_pointsto_vec_cons. do 2 iFrame. done. }
{ iExists vl. iFrame. }
iMod ("Hclose1" with "Hα HL") as "HL".
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ m own_ptr _ _; #(lm' + 1) &uniq{α} ty]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End code.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree.
From lrust.lang.lib Require Import memcpy lock.
From lrust.lifetime Require Import na_borrow.
From lrust.typing Require Export type.
From lrust.typing Require Import typing util option mutex.
From iris.prelude Require Import options.
(* This type is an experiment in defining a Rust type on top of a non-typesysten-specific
interface, like the one provided by lang.lib.lock.
It turns out that we need an accessor-based spec for this purpose, so that
we can put the protocol into shared borrows. Cancellable invariants
don't work because their cancelling view shift has a non-empty mask,
and it would have to be executed in the consequence view shift of
a borrow.
*)
Section mguard.
Context `{!typeGS Σ}.
(*
pub struct MutexGuard<'a, T: ?Sized + 'a> {
// funny underscores due to how Deref/DerefMut currently work (they
// disregard field privacy).
__lock: &'a Mutex<T>,
__poison: poison::Guard,
}
*)
Program Definition mutexguard (α : lft) (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] =>
β, α β
&at{α, mutexN} (lock_proto l (&{β} ((l + 1) ↦∗: ty.(ty_own) tid)))
&{β} ((l + 1) ↦∗: ty.(ty_own) tid)
| _ => False end;
ty_shr κ tid l :=
(l':loc), &frac{κ}(λ q', l {q'} #l')
F q, ⌜↑shrN lftN F -∗ q.[ακ]
={F}[F∖↑shrN]▷=∗ ty.(ty_shr) (ακ) tid (l' + 1) q.[ακ]
|}%I.
Next Obligation. by iIntros (? ty tid [|[[]|][]]) "H". Qed.
(* This is to a large extend copy-pasted from RWLock's write guard. *)
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 "[>[] _]".
setoid_rewrite heap_pointsto_vec_singleton.
iMod (bor_exists with "LFT Hb") as (β) "Hb"; first done.
iMod (bor_sep with "LFT Hb") as "[Hαβ 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 mutexguard_wf α ty `{!TyWf ty} : TyWf (mutexguard α ty) :=
{ ty_lfts := [α]; ty_wf_E := ty_wf_E ty ++ ty_outlives_E ty α }.
Global Instance mutexguard_type_contractive α : TypeContractive (mutexguard α).
Proof.
constructor;
solve_proper_core ltac:(fun _ => f_type_equiv || f_contractive_fin || f_equiv
|| exact: type_dist2_S).
Qed.
Global Instance mutexguard_ne α : NonExpansive (mutexguard α).
Proof. apply type_contractive_ne, _. Qed.
Global Instance mutexguard_mono E L :
Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) mutexguard.
Proof.
intros α1 α2 ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold.
iIntros (Hty' qmax qL) "HL".
iDestruct (Hty' with "HL") as "#Hty". clear Hty'. iDestruct ( with "HL") as "#Hα".
iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα12.
iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs) {HE Hty}". iSplit; [done|iSplit; iModIntro].
- iIntros (tid [|[[]|][]]) "H"; try done. simpl.
iDestruct "H" as (β) "(#H⊑ & #Hinv & Hown)".
iExists β. iFrame. iSplit; last iSplit.
+ iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done.
+ iApply at_bor_shorten; first by iApply lft_incl_syn_sem.
iApply (at_bor_iff with "[] Hinv"). iNext.
iApply lock_proto_iff_proper. iApply bor_iff_proper. iNext.
iApply heap_pointsto_pred_iff_proper.
iModIntro; iIntros; iSplit; iIntros; by iApply "Ho".
+ iApply bor_iff; last done. iNext.
iApply heap_pointsto_pred_iff_proper.
iModIntro; iIntros; iSplit; iIntros; by iApply "Ho".
- 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 mutexguard_proper E L :
Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) mutexguard.
Proof. intros ??[]???. split; by apply mutexguard_mono. Qed.
Global Instance mutexguard_sync α ty :
Sync ty Sync (mutexguard α 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 MutexGuard. We can
prove this for our spinlock implementation, however. *)
Global Instance mutexguard_send α ty :
Send ty Send (mutexguard α ty).
Proof.
iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H".
iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_entails.
repeat match goal with
| |- (ty_own _ _ _) (ty_own _ _ _) => by apply send_change_tid'
| |- _ => f_equiv
end.
Qed.
End mguard.
Section code.
Context `{!typeGS Σ}.
Lemma mutex_acc E l ty tid q α κ :
lftN E mutexN E
let R := (&{κ}((l + 1) ↦∗: ty_own ty tid))%I in
(* FIXME: using ⊢ to work around https://gitlab.mpi-sws.org/iris/iris/-/issues/496 *)
lft_ctx &at{α,mutexN}(lock_proto l R) -∗ α κ -∗
((q).[α] ={E,}=∗ lock_proto l R ( lock_proto l R ={,E}=∗ (q).[α])).
Proof.
(* FIXME: This should work: iIntros (?? R). *) intros ?? R.
iIntros "#LFT #Hshr #Hlincl !> Htok".
iMod (at_bor_acc_tok with "LFT Hshr Htok") as "[Hproto Hclose1]"; [done..|].
iMod (fupd_mask_subseteq) as "Hclose2"; last iModIntro; first solve_ndisj.
iFrame. iIntros "Hproto". iMod "Hclose2" as "_".
iMod ("Hclose1" with "Hproto") as "$". done.
Qed.
Definition mutex_lock : val :=
fn: ["mutex"] :=
let: "m" := !"mutex" in
let: "guard" := new [ #1 ] in
acquire ["m"];;
"guard" + #0 <- "m";;
delete [ #1; "mutex" ];; return: ["guard"].
Lemma mutex_lock_type ty `{!TyWf ty} :
typed_val mutex_lock (fn( α, ; &shr{α}(mutex ty)) mutexguard α 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 (m); simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (g); simpl_subst.
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hg [Hx [Hm _]]]".
rewrite !tctx_hasty_val [[x]]lock /=.
destruct m as [[|lm|]|]; try done. iDestruct "Hm" as (κ') "[#Hακ' #Hshr]".
iDestruct (ownptr_uninit_own with "Hg") as (lg vlg) "(% & Hg & Hg†)".
subst g. inv_vec vlg=>g. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & HL & Hclose1)"; [solve_typing..|].
wp_apply (acquire_spec with "[] Hα"); first by iApply (mutex_acc with "LFT Hshr Hακ'").
iIntros "[Hcont Htok]". wp_seq. wp_op. rewrite shift_loc_0. wp_write.
iMod ("Hclose1" with "Htok HL") as "HL".
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ x own_ptr _ _; #lg box (mutexguard α ty)]
with "[] LFT HE Hna HL Hk"); last first.
(* TODO: It would be nice to say [{#}] as the last spec pattern to clear the context in there. *)
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame "Hg".
iExists _. iFrame "#∗". }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition mutexguard_derefmut : val :=
fn: ["g"] :=
let: "g'" := !"g" in
let: "m" := !"g'" in
letalloc: "r" <- ("m" + #1) in
delete [ #1; "g"];; return: ["r"].
Lemma mutexguard_derefmut_type ty `{!TyWf ty} :
typed_val mutexguard_derefmut
(fn( '(α, β), ; &uniq{α}(mutexguard β ty)) &uniq{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>g. simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (g'); simpl_subst.
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hg [Hg' _]]".
rewrite !tctx_hasty_val [[g]]lock /=.
destruct g' as [[|lg|]|]; try done. simpl.
iMod (bor_exists with "LFT Hg'") as (vl) "Hbor"; first done.
iMod (bor_sep with "LFT Hbor") as "[H↦ Hprot]"; first done.
iMod (lctx_lft_alive_tok α with "HE HL") as () "(Hα & HL & Hclose1)";
[solve_typing..|].
destruct vl as [|[[|lm|]|] [|]]; simpl;
try by iMod (bor_persistent with "LFT Hprot Hα") as "[>[] _]".
rewrite heap_pointsto_vec_singleton.
iMod (bor_exists with "LFT Hprot") as (κ) "Hprot"; first done.
iMod (bor_sep with "LFT Hprot") as "[Hβκ Hprot]"; first done.
iMod (bor_sep with "LFT Hprot") as "[_ Hlm]"; first done.
iMod (bor_persistent with "LFT Hβκ Hα") as "[#Hβκ Hα]"; first done.
iMod (bor_acc with "LFT H↦ Hα") as "[H↦ Hclose2]"; first done.
wp_bind (!_)%E. iMod (bor_unnest with "LFT Hlm") as "Hlm"; first done.
wp_read. wp_let. iMod "Hlm".
iDestruct (lctx_lft_incl_incl_noend α β with "HL HE") as "%"; [solve_typing..|].
iMod ("Hclose2" with "H↦") as "[_ Hα]".
iMod ("Hclose1" with "Hα HL") as "HL".
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ g own_ptr _ _; #lm + #1 &uniq{α} ty ]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iApply bor_shorten; last done.
iApply lft_incl_glb; last by iApply lft_incl_refl.
iApply lft_incl_trans; last done. by iApply lft_incl_syn_sem. }
iApply type_letalloc_1; [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition mutexguard_deref : val := mutexguard_derefmut.
Lemma mutexguard_deref_type ty `{!TyWf ty} :
typed_val mutexguard_derefmut
(fn( '(α, β), ; &shr{α}(mutexguard β ty)) &shr{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([α β] ϝ ret arg). inv_vec arg=>g. simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (g'); simpl_subst.
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hg [Hg' _]]".
rewrite !tctx_hasty_val [[g]]lock /=.
destruct g' as [[|lg|]|]; try done. simpl.
iDestruct "Hg'" as (lm) "[Hlg Hshr]".
iMod (lctx_lft_alive_tok α with "HE HL") as () "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|].
iMod (frac_bor_acc with "LFT Hlg Hα1") as (qlx') "[H↦ Hclose2]"; first done.
iMod (lctx_lft_alive_tok_noend β with "HE HL") as () "(Hβ & HL & Hclose3)";
[solve_typing..|].
iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hclose4]".
wp_bind (!_)%E. iApply (wp_step_fupd with "[Hshr Hα2β]");
[|by iApply ("Hshr" with "[] Hα2β")|]; first done.
wp_read. iIntros "[#Hshr Hα2β] !>". wp_let.
iDestruct ("Hclose4" with "Hα2β") as "[Hβ Hα2]".
iMod ("Hclose3" with "Hβ HL") as "HL".
iMod ("Hclose2" with "H↦") as "Hα1".
iMod ("Hclose1" with "[$] HL") as "HL".
iDestruct (lctx_lft_incl_incl α β with "HL HE") as "%"; [solve_typing..|].
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ g own_ptr _ _; #lm + #1 &shr{α} ty ]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iApply ty_shr_mono; last done.
iApply lft_incl_glb; last by iApply lft_incl_refl. by iApply lft_incl_syn_sem. }
iApply type_letalloc_1; [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition mutexguard_drop : val :=
fn: ["g"] :=
let: "m" := !"g" in
release ["m"];;
delete [ #1; "g" ];;
let: "r" := new [ #0 ] in return: ["r"].
Lemma mutexguard_drop_type ty `{!TyWf ty} :
typed_val mutexguard_drop (fn( α, ; mutexguard α ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>g. simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (m); simpl_subst.
(* Switch to Iris. *)
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hg [Hm _]]".
rewrite !tctx_hasty_val [[g]]lock /=.
destruct m as [[|lm|]|]; try done. iDestruct "Hm" as (β) "(#Hαβ & #Hshr & Hcnt)".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & HL & Hclose1)"; [solve_typing..|].
wp_apply (release_spec with "[] [Hα Hcnt]");
[by iApply (mutex_acc with "LFT Hshr Hαβ")|by iFrame|].
iIntros "Htok". wp_seq. iMod ("Hclose1" with "Htok HL") as "HL".
(* Switch back to typing mode. *)
iApply (type_type _ _ _ [ g own_ptr _ _ ]
with "[] LFT HE Hna HL Hk"); last first.
{ rewrite tctx_interp_singleton tctx_hasty_val. unlock. iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_new; [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_jump; solve_typing.
Qed.
(* TODO: Should we do try_lock? *)
End code.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import typing lib.panic.
From iris.prelude Require Import options.
Section option.
Context `{!typeGS Σ}.
Definition option (τ : type) := Σ[unit; τ]%T.
Global Instance option_ne : NonExpansive option.
Proof. solve_proper. Qed.
Global Instance option_type_ne : TypeNonExpansive option.
Proof. solve_proper. Qed.
(* Variant indices. *)
Definition none := 0%nat.
Definition some := 1%nat.
Definition option_as_mut : val :=
fn: ["o"] :=
let: "o'" := !"o" in
let: "r" := new [ #2 ] in
withcont: "k":
case: !"o'" of
[ "r" <-{Σ none} ();; "k" [];
"r" <-{Σ some} "o'" + #1;; "k" [] ]
cont: "k" [] :=
delete [ #1; "o"];; return: ["r"].
Lemma option_as_mut_type τ `{!TyWf τ} :
typed_val
option_as_mut (fn( α, ; &uniq{α} (option τ)) option (&uniq{α}τ)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret p). inv_vec p=>o. simpl_subst.
iApply type_deref; [solve_typing..|]. iIntros (o'). simpl_subst.
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply (type_cont [] [ϝ []] (λ _, [o _; r _])) ; [solve_typing..| |].
- iIntros (k). simpl_subst.
iApply type_case_uniq; [solve_typing..|].
constructor; last constructor; last constructor; left.
+ iApply (type_sum_unit (option $ &uniq{α}τ)%T); [solve_typing..|].
iApply type_jump; solve_typing.
+ iApply (type_sum_assign (option $ &uniq{α}τ)%T); [try solve_typing..|].
iApply type_jump; solve_typing.
- iIntros "/= !>". iIntros (k args). inv_vec args. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition option_unwrap_or τ : val :=
fn: ["o"; "def"] :=
case: !"o" of
[ delete [ #(S τ.(ty_size)); "o"];;
return: ["def"];
letalloc: "r" <-{τ.(ty_size)} !("o" + #1) in
delete [ #(S τ.(ty_size)); "o"];; delete [ #τ.(ty_size); "def"];;
return: ["r"]].
Lemma option_unwrap_or_type τ `{!TyWf τ} :
typed_val (option_unwrap_or τ) (fn(; option τ, τ) τ).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([] ϝ ret p). inv_vec p=>o def. simpl_subst.
iApply type_case_own; [solve_typing|]. constructor; last constructor; last constructor.
+ right. iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
+ left. iApply type_letalloc_n; [solve_typing..|]. iIntros (r). simpl_subst.
iApply (type_delete (Π[uninit _;uninit _;uninit _])); [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition option_unwrap τ : val :=
fn: ["o"] :=
case: !"o" of
[ let: "panic" := panic in
letcall: "emp" := "panic" [] in
case: !"emp" of [];
letalloc: "r" <-{τ.(ty_size)} !("o" + #1) in
delete [ #(S τ.(ty_size)); "o"];;
return: ["r"]].
Lemma option_unwrap_type τ `{!TyWf τ} :
typed_val (option_unwrap τ) (fn(; option τ) τ).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros ([] ϝ ret p). inv_vec p=>o. simpl_subst.
iApply type_case_own; [solve_typing|]. constructor; last constructor; last constructor.
+ right. iApply type_let; [iApply panic_type|solve_typing|].
iIntros (panic). simpl_subst.
iApply (type_letcall ()); [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_case_own; solve_typing.
+ left. iApply type_letalloc_n; [solve_typing..|]. iIntros (r). simpl_subst.
iApply (type_delete (Π[uninit _;uninit _;uninit _])); [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End option.
From iris.proofmode Require Import proofmode.
From lrust.typing Require Import typing.
From iris.prelude Require Import options.
(* Minimal support for panic. Lambdarust does not support unwinding the
stack. Instead, we just end the current thread by not calling any
continuation.
This properly models "panic=abort", but fails to take into account the
trouble caused by unwinding. This is why we do not get into trouble with
[take_mut]. *)
Section panic.
Context `{!typeGS Σ}.
Definition panic : val :=
fn: [] := #☠.
Lemma panic_type : typed_val panic (fn() ).
Proof.
intros E L. iApply type_fn; [done|]. iIntros "!> _ %κ %ret %args".
inv_vec args. iIntros (tid qmax) "LFT HE Hna HL Hk HT /=". simpl_subst.
by iApply wp_value.
Qed.
End panic.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree excl numbers.
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 option.
From iris.prelude Require Import options.
Definition rc_stR :=
prodUR (optionUR (csumR (prodR fracR positiveR) (exclR unitO))) natUR.
Class rcG Σ :=
RcG :: inG Σ (authR rc_stR).
Definition rcΣ : gFunctors := #[GFunctor (authR rc_stR)].
Global Instance subG_rcΣ {Σ} : subG rcΣ Σ rcG Σ.
Proof. solve_inG. Qed.
Definition rc_tok q : authR rc_stR :=
(Some $ Cinl (q, 1%positive), 0%nat).
Definition rc_dropping_tok : authR rc_stR :=
(Some $ Cinr $ Excl (), 0%nat).
Definition weak_tok : authR rc_stR := (None, 1%nat).
Definition rcN := lrustN .@ "rc".
Definition rc_invN := rcN .@ "inv".
Definition rc_shrN := rcN .@ "shr".
Section rc.
Context `{!typeGS Σ, !rcG Σ}.
(* The RC can be in four different states :
- The living state, meaning that some strong reference exists. The
authoritative state is something like (Some (Cinl (q, strong)), weak),
where q is the total fraction owned by strong references, strong is
the number of strong references and weak is the number of weak
references.
- The "dropping" state, meaning that the last strong reference has been
dropped, and that the content in being dropped. The authoritative
state is something like (Some (Cinr (Excl ())), weak), where weak is
the number of weak references. The client owning the Excl also owns
the content of the box.
In our case, this state is not really necesary, because we do not
properly support dropping the content, but just copy it out of the RC
box. However, including it is more realistic, and this state is
still necessary for Arc anyway.
- The weak state, meaning that there only exists weak references. The
authoritative state is something like (None, weak), where weak is the
number of weak references.
- The dead state, meaning that no reference exist anymore. The
authoritative state is something like (None, 0).
Note that when we are in the living or dropping states, the weak reference
count stored in the heap is actually one plus the actual number of weak
references. This hack (which also exists in Rust's implementation) makes the
implementation of weak_drop easier, because it does not have to check the
strong count. *)
Definition rc_inv tid ν (γ : gname) (l : loc) (ty : type) : iProp Σ :=
( st : rc_stR, own γ ( st)
match st with
| (Some (Cinl (q, strong)), weak) => q',
l #(Zpos strong) (l + 1) #(S weak) l(2 + ty.(ty_size))
(q + q')%Qp = 1%Qp q'.[ν]
(* We keep this view shift decomposed because we store the persistent
part in ty_own, to make it available with one step less. *)
([ν] ={lftN}=∗ (l + 2) ↦∗: ty.(ty_own) tid)
| (Some (Cinr _), weak) =>
l #0 (l + 1) #(S weak)
| (None, (S _) as weak) =>
l #0 (l + 1) #weak l(2 + ty.(ty_size))
(l + 2) ↦∗: λ vl, length vl = ty.(ty_size)
| _ => True
end)%I.
Global Instance rc_inv_ne tid ν γ l n :
Proper (type_dist2 n ==> dist n) (rc_inv tid ν γ l).
Proof.
solve_proper_core ltac:(fun _ => f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Definition rc_persist tid ν (γ : gname) (l : loc) (ty : type) : iProp Σ :=
tc_opaque ( ty', type_incl ty' ty
na_inv tid rc_invN (rc_inv tid ν γ l ty')
(* We use this disjunction, and not simply [ty_shr] here,
because [weak_new] cannot prove ty_shr, even for a dead
lifetime. *)
(ty.(ty_shr) ν tid (l + 2) [ν])
(1.[ν] ={lftN lft_userE}[lft_userE]▷=∗ [ν]))%I.
Global Instance rc_persist_persistent tid ν γ l ty : Persistent (rc_persist tid ν γ l ty).
Proof. unfold rc_persist, tc_opaque. apply _. Qed.
Global Instance rc_persist_ne tid ν γ l n :
Proper (type_dist2 n ==> dist n) (rc_persist tid ν γ l).
Proof.
solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Lemma rc_persist_type_incl tid ν γ l ty1 ty2:
type_incl ty1 ty2 -∗ rc_persist tid ν γ l ty1 -∗ rc_persist tid ν γ l ty2.
Proof.
iIntros "#Hincl H". iDestruct "H" as (ty) "#(?&?& Hs &?)". iExists ty.
iFrame "#". iSplit.
- iNext. by iApply type_incl_trans.
- iDestruct "Hs" as "[?|?]"; last auto.
iLeft. iDestruct "Hincl" as "(_&_&Hincls)". by iApply "Hincls".
Qed.
Program Definition rc (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] =>
(* The ty_own part of the rc type cointains either the
shared protocol or the full ownership of the
content. The reason is that get_mut does not have the
masks to rebuild the invariants. *)
(l #1 (l + 1) #1 l(2 + ty.(ty_size))
(l + 2) ↦∗: ty.(ty_own) tid)
( γ ν q, rc_persist tid ν γ l ty own γ (rc_tok q) q.[ν])
| _ => False end;
ty_shr κ tid l :=
(l' : loc), &frac{κ} (λ q, l {q} #l')
F q, ⌜↑shrN lftN F -∗ q.[κ]
={F}[F∖↑shrN]▷=∗ q.[κ] γ ν q', κ ν
rc_persist tid ν γ l' ty
&na{κ, tid, rc_shrN}(own γ (rc_tok q'))
|}%I.
Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". 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 "[>[] _]".
setoid_rewrite heap_pointsto_vec_singleton.
iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx.
iDestruct "Hb" as (i) "(#Hpb&Hpbown)".
set (C := ( _ _ _, _ _ &na{_,_,_} _)%I).
iMod (inv_alloc shrN _ (idx_bor_own 1 i C)%I
with "[Hpbown]") as "#Hinv"; first by iLeft.
iIntros "!> !> * % Htok".
iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iAssert (&{κ} _)%I with "[Hbtok]" as "Hb".
{ rewrite bor_unfold_idx. iExists _. by iFrame. }
iClear "H↦ Hinv Hpb".
iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose2]"; first solve_ndisj.
set (X := ( _ _ _, _)%I).
iModIntro. iNext. iAssert X with "[>HP]" as "HX".
{ iDestruct "HP" as "[(Hl'1 & Hl'2 & H† & Hown)|$]"; last done.
iMod (lft_create with "LFT") as (ν) "[[Hν1 Hν2] #Hν†]"; first solve_ndisj.
(* TODO: We should consider changing the statement of
bor_create to dis-entangle the two masks such that this is no
longer necessary. *)
iApply (fupd_mask_mono (lftN)); first solve_ndisj.
iMod (bor_create with "LFT Hown") as "[HP HPend]"; first solve_ndisj.
iMod (own_alloc ( (Some $ Cinl ((1/2)%Qp, 1%positive), 0%nat)
(Some $ Cinl ((1/2)%Qp, 1%positive), 0%nat))) as (γ) "[Ha Hf]".
{ by apply auth_both_valid_discrete. }
iMod (na_inv_alloc tid _ _ (rc_inv tid ν γ l' ty)
with "[Ha Hν2 Hl'1 Hl'2 H† HPend]") as "#?".
{ rewrite /rc_inv. iExists (Some $ Cinl (_, _), _). iFrame. rewrite Qp.div_2; auto. }
iMod (ty_share with "LFT HP Hν1") as "[??]"; first solve_ndisj.
iExists _, _, _. iFrame. iExists ty. iFrame "#". iSplitR; last by auto.
by iApply type_incl_refl. }
iDestruct "HX" as (γ ν q') "(#Hpersist & Hrctok)".
iMod ("Hclose2" with "[] Hrctok") as "[HX $]".
{ unfold X. iIntros "!> [??] !>". iNext. iRight. do 3 iExists _.
iFrame "#∗". }
iAssert C with "[>HX]" as "#$".
{ iExists _, _, _. iFrame "Hpersist".
iMod (bor_sep with "LFT HX") as "[Hrc Hlft]"; first solve_ndisj.
iDestruct (frac_bor_lft_incl with "LFT [> Hlft]") as "$".
{ iApply (bor_fracture with "LFT"); first solve_ndisj. by rewrite Qp.mul_1_r. }
iApply (bor_na with "Hrc"); solve_ndisj. }
iApply ("Hclose1" with "[]"). by auto.
- iMod ("Hclose1" with "[]") as "_"; first by auto.
iApply step_fupd_intro; first solve_ndisj. by iFrame.
Qed.
Next Obligation.
iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]".
iExists _. iSplit; first by iApply frac_bor_shorten.
iModIntro. iIntros (F q) "% Htok".
iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hshr" with "[] Htok") as "Hshr2"; first done.
iModIntro. iNext. iMod "Hshr2" as "[Htok HX]".
iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν ?) "(? & ? & ?)".
iExists _, _, _. iModIntro. iFrame. iSplit.
- by iApply lft_incl_trans.
- by iApply na_bor_shorten.
Qed.
Global Instance rc_wf ty `{!TyWf ty} : TyWf (rc ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance rc_type_contractive : TypeContractive rc.
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance rc_ne : NonExpansive rc.
Proof. apply type_contractive_ne, _. Qed.
Lemma rc_subtype ty1 ty2 :
type_incl ty1 ty2 -∗ type_incl (rc ty1) (rc ty2).
Proof.
iIntros "#Hincl". iPoseProof "Hincl" as "(#Hsz & #Hoincl & #Hsincl)".
iSplit; first done. iSplit; iModIntro.
- iIntros "%tid %vl Hvl". destruct vl as [|[[|vl|]|] [|]]; try done.
iDestruct "Hvl" as "[(Hl1 & Hl2 & H† & Hc) | Hvl]".
{ iLeft. iFrame. iDestruct "Hsz" as %->.
iFrame. iApply (heap_pointsto_pred_wand with "Hc"). iApply "Hoincl". }
iDestruct "Hvl" as (γ ν q) "(#Hpersist & Htk & Hν)".
iRight. iExists _, _, _. iFrame "#∗". by iApply rc_persist_type_incl.
- iIntros "%κ %tid %l #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'.
iIntros "{$Hfrac} !> * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H".
iModIntro. iNext. iMod "H" as "[$ H]".
iDestruct "H" as (γ ν q') "(Hlft & Hpersist & Hna)".
iExists _, _, _. iFrame. by iApply rc_persist_type_incl.
Qed.
Global Instance rc_mono E L :
Proper (subtype E L ==> subtype E L) rc.
Proof.
iIntros (ty1 ty2 Hsub ??) "HL". iDestruct (Hsub with "HL") as "#Hsub".
iIntros "!> #HE". iApply rc_subtype. by iApply "Hsub".
Qed.
Global Instance rc_proper E L :
Proper (eqtype E L ==> eqtype E L) rc.
Proof. intros ??[]. by split; apply rc_mono. Qed.
End rc.
Section code.
Context `{!typeGS Σ, !rcG Σ}.
Lemma rc_check_unique ty F tid (l : loc) :
rc_invN F
{{{ na_own tid F ty_own (rc ty) tid [ #l ] }}}
!#l
{{{ strong, RET #(Zpos strong); l #(Zpos strong)
(((strong = 1%positive
( weak : Z, (l + 1) #weak
((weak = 1
|={}[lft_userE]▷=> l(2 + ty.(ty_size))
(l + 2) ↦∗: ty.(ty_own) tid na_own tid F)
(weak > 1
((l #1 -∗ (l + 1) #weak
={}=∗ na_own tid F ty_own (rc ty) tid [ #l ])
(l #0 -∗ (l + 1) #(weak - 1)
={}[lft_userE]▷=∗ (l + 2) ↦∗: ty.(ty_own) tid
((l + 2) ↦∗: (λ vl, length vl = ty.(ty_size))
={}=∗ na_own tid F))))))
(l #0 ={}[lft_userE]▷=∗
(l + 2) ↦∗: ty.(ty_own) tid l(2 + ty.(ty_size)) na_own tid F
(na_own tid F ={}=∗ weak : Z,
(l + 1) #weak
((weak = 1 l #0 na_own tid F)
(weak > 1
( l(2 + ty.(ty_size)) -∗ (l + 1) #(weak-1) -∗
(l + 2) ↦∗: (λ vl, length vl = ty.(ty_size))
={}=∗ na_own tid F))))))
)
((1 < strong)%positive
((l #(Zpos strong) ={}=∗ na_own tid F ty_own (rc ty) tid [ #l ])
(l #(Zpos strong - 1) ={}=∗ na_own tid F))))
}}}.
Proof.
iIntros (? Φ) "[Hna [(Hl1 & Hl2 & H† & Hown)|Hown]] HΦ".
- wp_read. iApply "HΦ". iFrame "Hl1". iLeft. iSplit; first done. iSplit.
+ iExists _. iFrame "Hl2". iLeft. iFrame. iSplit; first done.
iApply step_fupd_intro; auto.
+ iIntros "Hl1". iFrame. iApply step_fupd_intro; first done.
auto 10 with iFrame.
- iDestruct "Hown" as (γ ν q) "(#Hpersist & Htok & Hν1)".
iPoseProof "Hpersist" as (ty') "(Hincl & Hinv & _ & #Hνend)".
iMod (na_inv_acc with "Hinv Hna") as "(Hproto & Hna & Hclose)"; [solve_ndisj..|].
iDestruct "Hproto" as ([st weak]) "[>Hst Hproto]".
iCombine "Hst Htok" gives %[[[[=]|(?&st'&[=<-]&EQst'&Hincl)]
%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete.
simpl in EQst'. subst st. destruct Hincl as [Hincl|Hincl].
+ destruct st' as [[]| |]; try by inversion Hincl. apply (inj Cinl) in Hincl.
apply (inj2 pair) in Hincl. destruct Hincl as [<-%leibniz_equiv <-%leibniz_equiv].
iDestruct "Hproto" as (q') "(Hl1 & Hl2 & Hl† & >Hqq' & Hν & Hν†)".
iDestruct "Hqq'" as %Hqq'. iPoseProof "Hincl" as "(>Hincls & Hinclo & _)".
iDestruct "Hincls" as %Hincls.
wp_read. iApply "HΦ". iFrame "Hl1". iLeft. iSplit; first done. iSplit.
* iExists _. iFrame "Hl2". destruct weak.
-- iLeft. iSplit; first done. rewrite Hincls. iFrame "Hl†".
iMod (own_update_2 with "Hst Htok") as "Hst".
{ apply auth_update_dealloc.
rewrite -{1}(right_id ((None, 0%nat) : prodUR _ _) _ (_, _)).
apply cancel_local_update_unit, _. }
rewrite -[in (1).[ν]%I]Hqq' -[(|={lft_userE,}=>_)%I]fupd_trans.
iApply step_fupd_mask_mono;
last iMod ("Hνend" with "[$Hν $Hν1]") as "H†"; try done.
iModIntro. iNext. iMod "H†".
iMod fupd_mask_subseteq as "Hclose2"; last iMod ("Hν†" with "H†") as "Hty".
{ set_solver+. }
iMod "Hclose2" as "_". iModIntro.
iMod ("Hclose" with "[Hst $Hna]") as "$"; first by iExists _; iFrame.
iModIntro. iNext. iDestruct "Hty" as (vl) "[??]". iExists _. iFrame.
by iApply "Hinclo".
-- iRight. iSplit; first by auto with lia. iSplit.
++ iIntros "Hl1 Hl2".
iMod ("Hclose" with "[-Htok Hν1]") as "$"; last by auto 10 with iFrame.
iFrame. iExists _. auto with iFrame.
++ iIntros "Hl1 Hl2".
rewrite -[in (1).[ν]%I]Hqq' -[(|={lft_userE,}=>_)%I]fupd_trans.
iApply step_fupd_mask_mono;
last iMod ("Hνend" with "[$Hν $Hν1]") as "H†"; try done.
iModIntro. iNext. iMod "H†".
iMod fupd_mask_subseteq as "Hclose2"; last iMod ("Hν†" with "H†") as "Hty".
{ set_solver+. }
iMod "Hclose2" as "_". iModIntro.
iMod (own_update_2 with "Hst Htok") as "Hst".
{ apply auth_update_dealloc, prod_local_update_1,
(cancel_local_update_unit (Some _) None). }
iSplitL "Hty".
{ iDestruct "Hty" as (vy) "[H Hty]". iExists vy. iFrame.
by iApply "Hinclo". }
iIntros "!> H". iApply ("Hclose" with "[>-]"). iFrame.
rewrite Hincls /= !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l.
by iFrame.
* iIntros "Hl1".
iMod (own_update_2 with "Hst Htok") as "[Hst Htok]".
{ apply auth_update. etrans.
- rewrite [(Some _, weak)]pair_split. apply cancel_local_update_unit, _.
- apply (op_local_update _ _ (Some (Cinr (Excl tt)), 0%nat)). auto. }
rewrite -[(|={lft_userE,}=>_)%I]fupd_trans -[in (1).[ν]%I]Hqq'.
iApply step_fupd_mask_mono;
last iMod ("Hνend" with "[$Hν $Hν1]") as "H†"; try done.
iModIntro. iNext. iMod "H†".
iMod fupd_mask_subseteq as "Hclose2"; last iMod ("Hν†" with "H†") as "Hty".
{ set_solver+. }
iMod "Hclose2" as "_". iModIntro.
iMod ("Hclose" with "[Hst $Hna Hl1 Hl2]") as "$";
first by iExists _; iFrame; iFrame.
rewrite Hincls. iFrame. iSplitL "Hty".
{ iDestruct "Hty" as (vl) "[??]". iExists _. iFrame. by iApply "Hinclo". }
iIntros "!> Hna". iClear "Hνend". clear q' Hqq' weak Hval.
iMod (na_inv_acc with "Hinv Hna") as "(Hproto & Hna & Hclose)"; [solve_ndisj..|].
iDestruct "Hproto" as ([st weak]) "[>Hst Hproto]".
iCombine "Hst Htok" gives %[[[[=]|(?&st'&[=<-]&EQst'&Hincl)]
%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete.
simpl in EQst'. subst st. destruct Hincl as [Hincl|Hincl]; first last.
{ apply csum_included in Hincl. destruct Hincl as
[->|[(?&?&[=]&?)|(?&?&[=<-]&->&?%exclusive_included)]]; try done. apply _. }
setoid_subst. iDestruct "Hproto" as ">(Hl1 & Hl2)". iExists _. iFrame.
rewrite right_id. destruct weak as [|weak].
-- iLeft. iFrame. iSplitR; first done.
iApply ("Hclose" with "[>- $Hna]"). iExists (None, 0%nat).
iMod (own_update_2 with "Hst Htok") as "$"; last done.
apply auth_update_dealloc.
rewrite -{1}(right_id ((None, 0%nat) : prodUR _ _) _ (_, _)).
apply cancel_local_update_unit, _.
-- iRight. iSplitR; first by auto with lia. iIntros "!> Hl† Hl2 Hvl".
iApply ("Hclose" with "[>- $Hna]"). iExists (None, S weak).
rewrite Hincls. iFrame. iSplitR "Hl2"; last first.
{ by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. }
iMod (own_update_2 with "Hst Htok") as "$"; last done.
apply auth_update_dealloc, prod_local_update', reflexivity.
rewrite -{1}(right_id None _ (Some _)). apply: cancel_local_update_unit.
+ apply csum_included in Hincl.
destruct Hincl as [->|[(?&(?,?)&[=<-]&->&(q',strong)&[]%(inj2 pair))|
(?&?&[=]&?)]]; first done. setoid_subst.
iDestruct "Hproto" as (q'') "(Hl1 & Hl2 & Hl† & >Hqq' & Hν & Hν†)".
iDestruct "Hqq'" as %Hqq'. wp_read. iApply "HΦ". iFrame "Hl1". iRight.
iSplit; first by rewrite !pos_op_add; auto with lia. iSplit; iIntros "H↦".
* iMod ("Hclose" with "[- Htok Hν1]") as "$"; last by auto 10 with iFrame.
iFrame. iExists _. iFrame. auto with iFrame.
* iMod (own_update_2 with "Hst Htok") as "Hst".
{ apply auth_update_dealloc.
rewrite pair_op Cinl_op Some_op -{1}(left_id (0%nat : natUR) _ weak) pair_op.
apply (cancel_local_update_unit _ (_, _)). }
iApply "Hclose". iFrame "Hst Hl2 Hl† Hna Hν†". iExists (q+q'')%Qp. iFrame.
iSplitL; first last.
{ rewrite [(_+_)%Qp]assoc [(q'+_)%Qp]comm. auto. }
rewrite pos_op_add Z.sub_1_r -Pos2Z.inj_pred; last lia.
by rewrite Pos.add_1_l Pos.pred_succ.
Qed.
Definition rc_strong_count : val :=
fn: ["rc"] :=
let: "r" := new [ #1 ] in
let: "rc'" := !"rc" in
let: "rc''" := !"rc'" in
let: "strong" := !("rc''" + #0) in
"r" <- "strong";;
delete [ #1; "rc" ];; return: ["r"].
Lemma rc_strong_count_type ty `{!TyWf ty} :
typed_val rc_strong_count (fn( α, ; &shr{α}(rc ty)) int).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)".
iModIntro. wp_let. wp_op. rewrite shift_loc_0.
(* Finally, finally... opening the thread-local Rc protocol. *)
iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)".
iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|].
iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weak]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hrctok" gives %[[[[=]|(?&[[q0 s0]| |]&[=<-]&?&Hincl)]
%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete;
setoid_subst; try done; last first.
{ exfalso. destruct Hincl as [Hincl|Hincl].
- by inversion Hincl.
- apply csum_included in Hincl. naive_solver. }
iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >% & Hνtok & Hν†)".
wp_read. wp_let.
(* And closing it again. *)
iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hνtok Hν† $Hna]") as "Hna".
{ iExists _. iFrame "Hrc●". iExists _. auto with iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(rc ty)); r box (uninit 1);
#(Z.pos s0) int ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock.
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_weak_count : val :=
fn: ["rc"] :=
let: "r" := new [ #1 ] in
let: "rc'" := !"rc" in
let: "rc''" := !"rc'" in
let: "weak" := !("rc''" + #1) in
let: "one" := #1 in
let: "weak" := "weak" - "one" in
"r" <- "weak";;
delete [ #1; "rc" ];; return: ["r"].
Lemma rc_weak_count_type ty `{!TyWf ty} :
typed_val rc_weak_count (fn( α, ; &shr{α}(rc ty)) int).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock [[r]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)".
iModIntro. wp_let. wp_op.
(* Finally, finally... opening the thread-local Rc protocol. *)
iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)".
iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|].
iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weak]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hrctok" gives %[[[[=]|(?&[[q0 weak0]| |]&[=<-]&?&Hincl)]
%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete;
setoid_subst; try done; last first.
{ exfalso. destruct Hincl as [Hincl|Hincl].
- by inversion Hincl.
- apply csum_included in Hincl. naive_solver. }
iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >% & Hνtok & Hν†)".
wp_read. wp_let.
(* And closing it again. *)
iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hνtok Hν† $Hna]") as "Hna".
{ iExists _. iFrame "Hrc●". iExists _. auto with iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(rc ty)); r box (uninit 1);
#(S weak) int ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock.
rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
iFrame. }
iApply type_int. iIntros (?). simpl_subst.
iApply type_minus; [solve_typing..|]. iIntros (?). simpl_subst.
iApply type_assign; [solve_typing..|].
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_new ty : val :=
fn: ["x"] :=
let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in
let: "rc" := new [ #1 ] in
"rcbox" + #0 <- #1;;
"rcbox" + #1 <- #1;;
"rcbox" + #2 <-{ty.(ty_size)} !"x";;
"rc" <- "rcbox";;
delete [ #ty.(ty_size); "x"];; return: ["rc"].
Lemma rc_new_type ty `{!TyWf ty} :
typed_val (rc_new ty) (fn(; ty) rc ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (rc); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]".
rewrite !tctx_hasty_val.
iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x.
iDestruct (ownptr_uninit_own with "Hrcbox")
as (lrcbox vlrcbox) "(% & Hrcbox↦ & Hrcbox†)". subst rcbox. inv_vec vlrcbox=>???.
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]".
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]".
rewrite shift_loc_assoc.
iDestruct (ownptr_uninit_own with "Hrc") as (lrc vlrc) "(% & Hrc↦ & Hrc†)". subst rc.
inv_vec vlrc=>?. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write.
wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz.
wp_apply (wp_memcpy with "[$Hrcbox↦2 $Hx↦]"); [by auto using length_vec_to_list..|].
iIntros "[Hrcbox↦2 Hx↦]". wp_seq. wp_write.
iApply (type_type _ _ _ [ #lx box (uninit (ty_size ty)); #lrc box (rc ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=. iFrame.
iSplitR; first done.
iNext. iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame. iLeft.
rewrite freeable_sz_full_S. iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_clone : val :=
fn: ["rc"] :=
let: "r" := new [ #1 ] in
let: "rc'" := !"rc" in
let: "rc''" := !"rc'" in
let: "strong" := !("rc''" + #0) in
"rc''" + #0 <- "strong" + #1;;
"r" <- "rc''";;
delete [ #1; "rc" ];; return: ["r"].
Lemma rc_clone_type ty `{!TyWf ty} :
typed_val rc_clone (fn( α, ; &shr{α}(rc ty)) rc ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)".
subst r. inv_vec vlr=>r. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)".
iModIntro. wp_let. wp_op. rewrite shift_loc_0.
(* Finally, finally... opening the thread-local Rc protocol. *)
iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)".
iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|].
iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weak]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hrctok" gives %[[[[=]|(?&[[q0 s0]| |]&[=<-]&?&Hincl)]
%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete;
setoid_subst; try done; last first.
{ exfalso. destruct Hincl as [Hincl|Hincl].
- by inversion Hincl.
- apply csum_included in Hincl. naive_solver. }
iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >%Hq & Hνtok & Hν†)".
wp_read. wp_let. wp_op. rewrite shift_loc_0. wp_op. wp_write. wp_write.
(* And closing it again. *)
iMod (own_update with "Hrc●") as "[Hrc● Hrctok2]".
{ apply auth_update_alloc, prod_local_update_1,
(op_local_update_discrete _ _ (Some (Cinl ((qb/2)%Qp, 1%positive))))=>-[/= Hqa _].
split; simpl; last done. apply frac_valid. rewrite /= -Hq comm_L.
by apply Qp.add_le_mono_l, Qp.div_le. }
rewrite right_id -Some_op -Cinl_op -pair_op. iDestruct "Hνtok" as "[Hνtok1 Hνtok2]".
iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hνtok2 Hν† $Hna]") as "Hna".
{ iExists _. iFrame "Hrc●". iExists _. rewrite Z.add_comm. iFrame.
rewrite [_ _]comm frac_op -[(_ + _)%Qp]assoc Qp.div_2. auto. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(rc ty)); #lr box (rc ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame "Hx". iFrame "Hr†". iExists [# #l'].
rewrite heap_pointsto_vec_singleton /=. simpl. eauto 10 with iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_deref : val :=
fn: ["rc"] :=
let: "x" := new [ #1 ] in
let: "rc'" := !"rc" in
"x" <- (!"rc'" + #2);;
delete [ #1; "rc" ];; return: ["x"].
Lemma rc_deref_type ty `{!TyWf ty} :
typed_val rc_deref (fn( α, ; &shr{α}(rc ty)) &shr{α}ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (x); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [#Hrc' [Hx _]]]".
rewrite !tctx_hasty_val [[rcx]]lock.
iDestruct (ownptr_uninit_own with "Hx") as (lrc2 vlrc2) "(% & Hx & Hx†)".
subst x. inv_vec vlrc2=>?. rewrite heap_pointsto_vec_singleton.
destruct rc' as [[|rc'|]|]; try done. simpl of_val.
iDestruct "Hrc'" as (l') "[Hrc' #Hdelay]".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hdelay" with "[] Hα1"); last iApply (wp_step_fupd with "Hdelay"); [done..|].
iMod (frac_bor_acc with "LFT Hrc' Hα2") as (q') "[Hrc'↦ Hclose2]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose2" with "[$Hrc'↦]") as "Hα2". iIntros "!> [Hα1 #Hproto] !>".
iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & _)".
iDestruct "Hpersist" as (ty') "(_ & _ & [Hshr|Hν†] & _)"; last first.
{ iMod (lft_incl_dead with "Hαν Hν†") as "Hα†"; first done.
iDestruct (lft_tok_dead with "Hα1 Hα†") as "[]". }
wp_op. wp_write. iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ rcx box (&shr{α}(rc ty)); #lrc2 box (&shr{α}ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame "Hrcx". iFrame "Hx†". iExists [_]. rewrite heap_pointsto_vec_singleton.
iFrame "Hx". by iApply ty_shr_mono. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_try_unwrap ty : val :=
fn: ["rc"] :=
let: "r" := new [ #(ty_size Σ[ ty; rc ty ]) ] in
withcont: "k":
let: "rc'" := !"rc" in
let: "strong" := !("rc'" + #0) in
if: "strong" = #1 then
(* Decrement strong *)
"rc'" + #0 <- #0;;
Skip;;
(* Return content *)
"r" <-{ty.(ty_size),Σ 0} !("rc'" + #2);;
(* Decrement weak (removing the one weak ref collectively held by all
strong refs), and deallocate if weak count reached 0. *)
let: "weak" := !("rc'" + #1) in
if: "weak" = #1 then
delete [ #(2 + ty.(ty_size)); "rc'" ];;
"k" []
else
"rc'" + #1 <- "weak" - #1;;
"k" []
else
"r" <-{Σ 1} "rc'";;
"k" []
cont: "k" [] :=
delete [ #1; "rc"];; return: ["r"].
Lemma rc_try_unwrap_type ty `{!TyWf ty} :
typed_val (rc_try_unwrap ty) (fn(; rc ty) Σ[ ty; rc ty ]).
Proof.
(* TODO: There is a *lot* of duplication between this proof and the one for drop. *)
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new (Σ[ ty; rc ty ]).(ty_size)); [solve_typing..|]; iIntros (r); simpl_subst.
iApply (type_cont [] [ϝ []]
(λ _, [rcx box (uninit 1); r box (Σ[ ty; rc ty ])])) ;
[solve_typing..| |]; last first.
{ simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iIntros (k). simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock.
destruct rc' as [[|rc'|]|]; try done.
rewrite [[ #rc' ]]lock. wp_op. rewrite shift_loc_0 -(lock [ #rc' ]).
wp_apply (rc_check_unique with "[$Hna $Hrc']"); first solve_ndisj.
iIntros (strong) "[Hrc Hc]". wp_let.
iDestruct "Hc" as "[[% [_ Hproto]]|[% [Hproto _]]]".
- subst strong. wp_op. wp_if. wp_op.
rewrite shift_loc_0. wp_write. wp_bind (#☠;;#☠)%E.
iApply (wp_step_fupd with "[Hproto Hrc]");
[|by iApply ("Hproto" with "Hrc")|];
[done..|wp_seq; iIntros "(Hty&H†&Hna&Hproto) !>"].
rewrite <-freeable_sz_full_S, <-(freeable_sz_split _ 2 ty.(ty_size)).
iDestruct "H†" as "[H†1 H†2]". wp_seq. wp_bind (_<-_;;_)%E.
iApply (wp_wand with "[Hna HL Hty Hr H†2]").
{ iApply (type_sum_memcpy_instr 0 [_; (rc ty)] with "LFT HE Hna HL"); first done.
{ by eapply (write_own _ (uninit _)). } { apply read_own_move. }
iSplitL "Hr"; first by unlock; rewrite tctx_hasty_val. iSplitL; last done.
rewrite tctx_hasty_val'; last done. iFrame. }
iIntros (?) "(Hna & HL & [Hr [Hrc _]])". wp_seq.
iMod ("Hproto" with "Hna") as (weak) "[H↦weak Hproto]". wp_op. wp_read. wp_let.
iDestruct "Hproto" as "[(% & H↦strong & Hna)|[% Hproto]]".
+ subst. wp_op. wp_if.
iApply (type_type _ _ _
[ r own_ptr (ty_size Σ[ ty; rc ty ]) (Σ[ ty; rc ty]);
rcx box (uninit 1);
#rc' + #2 own_ptr (2 + ty.(ty_size)) (uninit (ty.(ty_size)));
#rc' + #0 own_ptr (2 + ty.(ty_size)) (uninit 2)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock. rewrite 3!tctx_interp_cons tctx_interp_singleton. iFrame "Hr Hrc".
rewrite 1!tctx_hasty_val tctx_hasty_val' //. iFrame "Hrcx".
rewrite shift_loc_0 /=. iFrame. iExists [_; _].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
auto with iFrame. }
iApply (type_delete (Π[uninit 2;uninit ty.(ty_size)])); [solve_typing..|].
iApply type_jump; solve_typing.
+ rewrite (tctx_hasty_val' _ (#rc' + #2)); last done.
iDestruct "Hrc" as "[Hrc H†2]". wp_op; case_bool_decide; first lia. wp_if. wp_op. wp_op. wp_write.
iMod ("Hproto" with "[H†1 H†2] H↦weak Hrc") as "Hna".
{ rewrite -freeable_sz_full_S -(freeable_sz_split _ 2 ty.(ty_size)). iFrame. }
iApply (type_type _ _ _
[ r own_ptr (ty_size Σ[ ty; rc ty ]) (Σ[ ty; rc ty]);
rcx box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply type_jump; solve_typing.
- wp_op; case_bool_decide as Hc; first lia. wp_if. iMod ("Hproto" with "Hrc") as "[Hna Hrc]".
iApply (type_type _ _ _ [ r own_ptr (ty_size Σ[ ty; rc ty ]) (uninit _);
rcx box (uninit 1); #rc' rc ty ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton.
rewrite !tctx_hasty_val tctx_hasty_val' //. unlock. iFrame. }
iApply (type_sum_assign Σ[ ty; rc ty ]); [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_drop ty : val :=
fn: ["rc"] :=
let: "r" := new [ #(option ty).(ty_size) ] in
withcont: "k":
let: "rc'" := !"rc" in
let: "strong" := !("rc'" + #0) in
"rc'" + #0 <- "strong" - #1;;
if: "strong" = #1 then
(* Return content. *)
"r" <-{ty.(ty_size),Σ some} !("rc'" + #2);;
(* Decrement weak (removing the one weak ref collectively held by all
strong refs), and deallocate if weak count reached 0. *)
let: "weak" := !("rc'" + #1) in
if: "weak" = #1 then
delete [ #(2 + ty.(ty_size)); "rc'" ];;
"k" []
else
"rc'" + #1 <- "weak" - #1;;
"k" []
else
"r" <-{Σ none} ();;
"k" []
cont: "k" [] :=
delete [ #1; "rc"];; return: ["r"].
Lemma rc_drop_type ty `{!TyWf ty} :
typed_val (rc_drop ty) (fn(; rc ty) option ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new (option ty).(ty_size)); [solve_typing..|]; iIntros (r); simpl_subst.
iApply (type_cont [] [ϝ []]
(λ _, [rcx box (uninit 1); r box (option ty)]));
[solve_typing..| |]; last first.
{ simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iIntros (k). simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock.
destruct rc' as [[|rc'|]|]; try done. rewrite [[ #rc' ]]lock.
wp_op. rewrite shift_loc_0. rewrite -(lock [ #rc' ]).
wp_apply (rc_check_unique with "[$Hna $Hrc']"); first solve_ndisj.
iIntros (strong) "[Hrc Hc]". wp_let. wp_op. rewrite shift_loc_0.
rewrite {3}[Z.pos strong]lock. wp_op. unlock. wp_write.
iDestruct "Hc" as "[[% [_ Hproto]]|[% [_ Hproto]]]".
- subst strong. wp_bind (#1 = #1)%E.
iApply (wp_step_fupd with "[Hproto Hrc]");
[|by iApply ("Hproto" with "Hrc")|];
[done..|wp_op; iIntros "(Hty&H†&Hna&Hproto) !>"; wp_if].
rewrite <-freeable_sz_full_S, <-(freeable_sz_split _ 2 ty.(ty_size)).
iDestruct "H†" as "[H†1 H†2]". wp_bind (_<-_;;_)%E.
iApply (wp_wand with "[Hna HL Hty Hr H†2]").
{ iApply (type_sum_memcpy_instr 1 [unit; _] with "LFT HE Hna HL"); first done.
{ by eapply (write_own _ (uninit _)). } { apply read_own_move. }
iSplitL "Hr"; first by unlock; rewrite tctx_hasty_val. iSplitL; last done.
rewrite tctx_hasty_val'; last done. iFrame. }
iIntros (?) "(Hna & HL & [Hr [Hrc _]])". wp_seq.
iMod ("Hproto" with "Hna") as (weak) "[H↦weak Hproto]". wp_op. wp_read. wp_let.
iDestruct "Hproto" as "[(% & H↦strong & Hna)|[% Hproto]]".
+ subst. wp_op. wp_if.
iApply (type_type _ _ _
[ r own_ptr (ty_size (option ty)) (option ty);
rcx box (uninit 1);
#rc' + #2 own_ptr (2 + ty.(ty_size)) (uninit (ty.(ty_size)));
#rc' + #0 own_ptr (2 + ty.(ty_size)) (uninit 2)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock. rewrite 3!tctx_interp_cons tctx_interp_singleton. iFrame "Hr Hrc".
rewrite 1!tctx_hasty_val tctx_hasty_val' //. iFrame "Hrcx".
rewrite shift_loc_0 /=. iFrame. iExists [_; _].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton.
auto with iFrame. }
iApply (type_delete (Π[uninit 2;uninit ty.(ty_size)])); [solve_typing..|].
iApply type_jump; solve_typing.
+ rewrite (tctx_hasty_val' _ (#rc' + #2)); last done.
iDestruct "Hrc" as "[Hrc H†2]". wp_op. case_bool_decide; first lia. wp_if. wp_op. wp_op. wp_write.
iMod ("Hproto" with "[H†1 H†2] H↦weak Hrc") as "Hna".
{ rewrite -freeable_sz_full_S -(freeable_sz_split _ 2 ty.(ty_size)). iFrame. }
iApply (type_type _ _ _
[ r own_ptr (ty_size (option ty)) (option ty); rcx box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
iApply type_jump; solve_typing.
- wp_op; case_bool_decide as Hc; first lia. wp_if. iMod ("Hproto" with "Hrc") as "Hna".
iApply (type_type _ _ _ [ r own_ptr (ty_size (option ty)) (uninit _);
rcx box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton.
rewrite !tctx_hasty_val. iFrame. }
iApply (type_sum_unit (option ty)); [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_get_mut : val :=
fn: ["rc"] :=
let: "r" := new [ #2 ] in
withcont: "k":
let: "rc'" := !"rc" in
let: "rc''" := !"rc'" in
let: "strong" := !("rc''" + #0) in
if: "strong" = #1 then
let: "weak" := !("rc''" + #1) in
if: "weak" = #1 then
"r" <-{Σ some} ("rc''" + #2);;
"k" []
else
"r" <-{Σ none} ();;
"k" []
else
"r" <-{Σ none} ();;
"k" []
cont: "k" [] :=
delete [ #1; "rc"];; return: ["r"].
Lemma rc_get_mut_type ty `{!TyWf ty} :
typed_val rc_get_mut (fn( α, ; &uniq{α}(rc ty)) option (&uniq{α}ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst.
iApply (type_cont [] [ϝ []]
(λ _, [rcx box (uninit 1); r box (option $ &uniq{α}ty)]));
[solve_typing..| |]; last first.
{ simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iIntros (k). simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock.
destruct rc' as [[|rc'|]|]; try done.
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & HL & Hclose1)"; [solve_typing..|].
iMod (bor_acc_cons with "LFT Hrc' Hα") as "[Hrc' Hclose2]"; first solve_ndisj.
iDestruct (heap_pointsto_ty_own with "Hrc'") as (vl) "[>Hrc' Hrcown]".
inv_vec vl=>l. rewrite heap_pointsto_vec_singleton.
wp_read. destruct l as [[|l|]|]; try done.
wp_let. wp_op. rewrite shift_loc_0.
wp_apply (rc_check_unique with "[$Hna Hrcown]"); first done.
{ (* Boy this is silly... *) iDestruct "Hrcown" as "[(?&?&?&?)|?]"; last by iRight.
iLeft. by iFrame. }
iIntros (c) "(Hl1 & Hc)". wp_let. wp_op; case_bool_decide as Hc.
- wp_if. iDestruct "Hc" as "[[% [Hc _]]|[% _]]"; last lia. subst.
iDestruct "Hc" as (weak) "[Hl2 Hweak]". wp_op. wp_read. wp_let.
iDestruct "Hweak" as "[[% Hrc]|[% [Hrc _]]]".
+ subst. wp_bind (#1 = #1)%E. iApply (wp_step_fupd with "Hrc"); [done..|].
wp_op. iIntros "(Hl† & Hty & Hna)!>". wp_if.
iMod ("Hclose2" with "[Hrc' Hl1 Hl2 Hl†] [Hty]") as "[Hty Hα]"; [|iNext; iExact "Hty"|].
{ iIntros "!> Hty". iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame "Hrc'".
iLeft. by iFrame. }
iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 2); #l + #2 &uniq{α}ty;
rcx box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. }
iApply (type_sum_assign (option (&uniq{_}_))); [solve_typing..|].
iApply type_jump; solve_typing.
+ wp_op; case_bool_decide; first lia. wp_if. iMod ("Hrc" with "Hl1 Hl2") as "[Hna Hrc]".
iMod ("Hclose2" with "[] [Hrc' Hrc]") as "[Hrc' Hα]".
{ iIntros "!> HX". iModIntro. iExact "HX". }
{ iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame. auto. }
iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 2); rcx box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
unlock. iFrame. }
iApply (type_sum_unit (option (&uniq{_}_))); [solve_typing..|].
iApply type_jump; solve_typing.
- wp_if. iDestruct "Hc" as "[[% _]|[% [Hproto _]]]"; first lia.
iMod ("Hproto" with "Hl1") as "[Hna Hrc]".
iMod ("Hclose2" with "[] [Hrc' Hrc]") as "[Hrc' Hα]".
{ iIntros "!> HX". iModIntro. iExact "HX". }
{ iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame. auto. }
iMod ("Hclose1" with "Hα HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 2); rcx box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
unlock. iFrame. }
iApply (type_sum_unit (option (&uniq{_}_))); [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* TODO : it is not nice that we have to inline the definition of
rc_new and of rc_deref. *)
Definition rc_make_mut (ty : type) (clone : val) : val :=
fn: ["rc"] :=
let: "r" := new [ #1 ] in
withcont: "k":
let: "rc'" := !"rc" in
let: "rc''" := !"rc'" in
let: "strong" := !("rc''" + #0) in
if: "strong" = #1 then
let: "weak" := !("rc''" + #1) in
if: "weak" = #1 then
(* This is the last strong ref, and there is no weak ref.
We just return a deep ptr into the Rc. *)
"r" <- "rc''" + #2;;
"k" []
else
(* This is the last strong ref, but there are weak refs.
We make ourselves a new Rc, move the content, and mark the old one killed
(strong count becomes 0, strong idx removed from weak cnt).
We store the new Rc in our argument (which is a &uniq rc),
and return a deep ptr into it. *)
"rc''" + #0 <- #0;;
"rc''" + #1 <- "weak" - #1;;
(* Inlined rc_new("rc''" +ₗ #2) begins. *)
let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in
"rcbox" + #0 <- #1;;
"rcbox" + #1 <- #1;;
"rcbox" + #2 <-{ty.(ty_size)} !"rc''" + #2;;
"rc'" <- "rcbox";;
(* Inlined rc_new ends. *)
"r" <- "rcbox" + #2;;
"k" []
else
(* There are other strong refs, we have to make a copy and clone the content. *)
let: "x" := new [ #1 ] in
"x" <- "rc''" + #2;;
let: "clone" := clone in
letcall: "c" := "clone" ["x"]%E in (* FIXME : why do I need %E here ? *)
(* Inlined rc_new("c") begins. *)
let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in
"rcbox" + #0 <- #1;;
"rcbox" + #1 <- #1;;
"rcbox" + #2 <-{ty.(ty_size)} !"c";;
delete [ #ty.(ty_size); "c"];;
"rc'" <- "rcbox";;
(* Inlined rc_new ends. *)
letalloc: "rcold" <- "rc''" in
(* FIXME : here, we are dropping the old rc pointer. In the
case another strong reference has been dropped while
cloning, it is possible that we are actually dropping the
last reference here. This means that we may have to drop an
instance of [ty] here. Instead, we simply forget it. *)
let: "drop" := rc_drop ty in
letcall: "content" := "drop" ["rcold"]%E in
delete [ #(option ty).(ty_size); "content"];;
"r" <- "rcbox" + #2;;
"k" []
cont: "k" [] :=
delete [ #1; "rc"];; return: ["r"].
Lemma rc_make_mut_type ty `{!TyWf ty} clone :
(* ty : Clone, as witnessed by the impl clone *)
typed_val clone (fn( α, ; &shr{α}ty) ty)
typed_val (rc_make_mut ty clone) (fn( α, ; &uniq{α}(rc ty)) &uniq{α}ty).
Proof.
intros Hclone E L. iApply type_fn; [solve_typing..|].
rewrite [(2 + ty_size ty)%nat]lock.
iIntros "/= !>". iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply (type_cont [] [ϝ []]
(λ _, [rcx box (uninit 1); r box (&uniq{α}ty)]));
[solve_typing..| |]; last first.
{ simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iIntros (k). simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock.
destruct rc' as [[|rc'|]|]; try done.
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)";
[solve_typing..|].
iMod (bor_acc_cons with "LFT Hrc' Hα1") as "[Hrc' Hclose2]"; first solve_ndisj.
iDestruct (heap_pointsto_ty_own with "Hrc'") as (vl) "[>Hrc' Hrcown]".
inv_vec vl=>l. rewrite heap_pointsto_vec_singleton.
wp_read. destruct l as [[|l|]|]; try done. wp_let. wp_op. rewrite shift_loc_0.
wp_apply (rc_check_unique with "[$Hna Hrcown]"); first done.
{ (* Boy this is silly... *) iDestruct "Hrcown" as "[(?&?&?&?)|?]"; last by iRight.
iLeft. by iFrame. }
iIntros (c) "(Hl1 & Hc)". wp_let. wp_op; case_bool_decide as Hc; wp_if.
- iDestruct "Hc" as "[[% [Hc _]]|[% _]]"; last lia. subst.
iDestruct "Hc" as (weak) "[Hl2 Hweak]". wp_op. wp_read. wp_let.
iDestruct "Hweak" as "[[% Hrc]|[% [_ Hrc]]]".
+ subst. wp_bind (#1 = #1)%E. iApply (wp_step_fupd with "Hrc"); [done..|].
wp_op. iIntros "(Hl† & Hty & Hna)!>". wp_if.
iMod ("Hclose2" with "[Hrc' Hl1 Hl2 Hl†] [Hty]") as "[Hty Hα1]"; [|iNext; iExact "Hty"|].
{ iIntros "!> Hty". iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame "Hrc'".
iLeft. by iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 1); #l + #2 &uniq{α}ty;
rcx box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
+ wp_op; case_bool_decide; first lia. wp_if. wp_op. rewrite shift_loc_0. wp_write. wp_op.
wp_op. wp_write. wp_bind (new _). iSpecialize ("Hrc" with "Hl1 Hl2").
iApply (wp_step_fupd with "Hrc"); [done..|]. iApply wp_new; first lia; first done.
rewrite -!lock Nat2Z.id.
iNext. iIntros (lr) "/= ([H†|%] & H↦lr) [Hty Hproto] !>"; last lia.
rewrite 2!heap_pointsto_vec_cons. iDestruct "H↦lr" as "(Hlr1 & Hlr2 & Hlr3)".
wp_let. wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write. wp_op. wp_op.
iDestruct "Hty" as (vlr) "[H↦vlr Hty]". rewrite shift_loc_assoc.
iDestruct (ty_size_eq with "Hty") as %Hsz.
wp_apply (wp_memcpy with "[$Hlr3 $H↦vlr]").
{ by rewrite repeat_length. } { by rewrite Hsz. }
iIntros "[Hlr3 Hvlr]". wp_seq. wp_write. wp_op.
iMod ("Hproto" with "[Hvlr]") as "Hna"; first by eauto.
iMod ("Hclose2" $! ((lr + 2) ↦∗: ty_own ty tid)%I
with "[Hrc' Hlr1 Hlr2 H†] [Hlr3 Hty]") as "[Hb Hα1]".
{ iIntros "!> H !>". iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iLeft. iFrame. }
{ iExists _. iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
iApply (type_type _ _ _ [ r box (uninit 1); #(lr + 2) &uniq{α}ty;
rcx box (uninit 1)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
tctx_hasty_val' //. unlock. iFrame. }
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
- wp_apply wp_new; first lia; first done.
iIntros (lr) "/= ([H†|%] & H↦lr)"; last lia.
iDestruct "Hc" as "[[% ?]|[% [Hproto _]]]"; first lia.
iMod ("Hproto" with "Hl1") as "[Hna Hty]". wp_let. wp_op.
rewrite heap_pointsto_vec_singleton. wp_write.
iAssert ( γ ν q, rc_persist tid ν γ l ty own γ (rc_tok q) q.[ν])%I
with "[>Hty]" as (γ ν q') "(Hty & Htok & Hν)".
{ iDestruct "Hty" as "[(Hl1 & Hl2 & Hl† & Hl3)|Hty]"; last done.
iMod (own_alloc ( (Some $ Cinl ((1/2)%Qp, 1%positive), 0%nat)
rc_tok (1/2)%Qp)) as (γ) "[Ha Hf]";
first by apply auth_both_valid_discrete.
iMod (lft_create with "LFT") as (ν) "[[Hν1 Hν2] #Hν†]"=>//.
iApply (fupd_mask_mono (lftN))=>//. iExists γ, ν, (1/2)%Qp. iFrame.
iMod (bor_create _ ν with "LFT Hl3") as "[Hb Hh]"=>//. iExists _.
iSplitR; first by iApply type_incl_refl.
iMod (ty_share with "LFT Hb Hν1") as "[Hty Hν]"=>//.
iSplitR "Hty"; last by auto. iApply na_inv_alloc. iExists _.
do 2 iFrame. by rewrite Qp.div_2. }
iDestruct "Hty" as (ty') "#(Hty'ty & Hinv & Hs & Hν†)".
iDestruct "Hs" as "[Hs|Hν']"; last by iDestruct (lft_tok_dead with "Hν Hν'") as "[]".
wp_bind (of_val clone). iApply (wp_wand with "[Hna]").
{ iApply (Hclone _ [] $! _ 1%Qp with "LFT HE Hna").
- by rewrite /llctx_interp.
- by rewrite /tctx_interp. }
clear clone Hclone. iIntros (clone) "(Hna & _ & Hclone)".
wp_let. wp_let. rewrite tctx_interp_singleton tctx_hasty_val.
iDestruct (lft_intersect_acc with "Hα2 Hν") as (q'') "[Hαν Hclose3]".
rewrite -[α ν](right_id_L).
iApply (type_call_iris _ [α ν] (α ν) [_]
with "LFT HE Hna Hαν Hclone [H† H↦lr]"); [solve_typing| |].
{ rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full_S.
iFrame. iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iApply ty_shr_mono; last done. iApply lft_intersect_incl_r. }
iIntros ([[|cl|]|]) "Hna Hαν Hcl //". wp_rec.
iDestruct "Hcl" as "[Hcl Hcl†]". iDestruct "Hcl" as (vl) "[Hcl↦ Hown]".
iDestruct (ty_size_eq with "Hown") as %Hsz.
iDestruct ("Hclose3" with "Hαν") as "[Hα2 Hν]".
wp_apply wp_new=>//; first lia. iIntros (l') "(Hl'† & Hl')". wp_let. wp_op.
rewrite shift_loc_0. rewrite -!lock Nat2Z.id.
rewrite !heap_pointsto_vec_cons shift_loc_assoc.
iDestruct "Hl'" as "(Hl' & Hl'1 & Hl'2)". wp_write. wp_op. wp_write. wp_op.
wp_apply (wp_memcpy with "[$Hl'2 $Hcl↦]").
{ by rewrite repeat_length. } { by rewrite Hsz. }
iIntros "[Hl'2 Hcl↦]". wp_seq. rewrite freeable_sz_full.
wp_apply (wp_delete with "[$Hcl↦ Hcl†]");
[lia|by replace (length vl) with (ty.(ty_size))|].
iIntros "_". wp_seq. wp_write.
iMod ("Hclose2" $! ((l' + 2) ↦∗: ty_own ty tid)%I with
"[Hrc' Hl' Hl'1 Hl'†] [Hl'2 Hown]") as "[Hl' Hα1]".
{ iIntros "!> H". iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iLeft. iFrame. iDestruct "Hl'†" as "[?|%]"=>//. }
{ iExists _. iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
iApply (type_type _ _ _ [ #l rc ty; #l' + #2 &uniq{α}ty;
r box (uninit 1); rcx box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 3!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val
!tctx_hasty_val' //. unlock. iFrame. iRight. iExists γ, ν, _.
unfold rc_persist, tc_opaque. iFrame "∗#". }
iApply type_letalloc_1; [solve_typing..|]. iIntros (rcold). simpl_subst.
iApply type_let.
{ apply rc_drop_type. }
{ solve_typing. }
iIntros (drop). simpl_subst.
iApply (type_letcall ()); [solve_typing..|]. iIntros (content). simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_assign; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
End code.
From iris.proofmode Require Import proofmode.
From iris.algebra Require Import auth csum frac agree numbers.
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 option.
From lrust.typing.lib Require Export rc.
From iris.prelude Require Import options.
Section weak.
Context `{!typeGS Σ, !rcG Σ}.
Program Definition weak (ty : type) :=
{| ty_size := 1;
ty_own tid vl :=
match vl return _ with
| [ #(LitLoc l) ] => γ ν, rc_persist tid ν γ l ty own γ weak_tok
| _ => False end;
ty_shr κ tid l :=
(l' : loc), &frac{κ} (λ q, l {q} #l')
F q, ⌜↑shrN lftN F -∗ q.[κ]
={F}[F∖↑shrN]▷=∗ q.[κ] γ ν, rc_persist tid ν γ l' ty
&na{κ, tid, rc_shrN}(own γ weak_tok)
|}%I.
Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". 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.
(* Ideally, we'd change ty_shr to say "l ↦{q} #l" in the fractured borrow,
but that would be additional work here... *)
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 "[>[] _]".
setoid_rewrite heap_pointsto_vec_singleton.
iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx.
iDestruct "Hb" as (i) "(#Hpb&Hpbown)".
set (C := ( _ _, _ &na{_,_,_} _)%I).
iMod (inv_alloc shrN _ (idx_bor_own 1 i C)%I with "[Hpbown]") as "#Hinv";
first by iLeft.
iIntros "!> !> * % Htok".
iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj.
iDestruct "INV" as "[>Hbtok|#Hshr]".
- iAssert (&{κ} _)%I with "[Hbtok]" as "Hb".
{ rewrite bor_unfold_idx. iExists _. by iFrame. }
iClear "H↦ Hinv Hpb".
iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose2]"; first solve_ndisj.
iDestruct "HP" as (γ ν) "(#Hpersist & Hweak)".
iModIntro. iNext. iMod ("Hclose2" with "[] Hweak") as "[Hb $]"; first by auto 10.
iAssert C with "[>Hb]" as "#HC".
{ iExists _, _. iFrame "Hpersist". iApply (bor_na with "Hb"). solve_ndisj. }
iMod ("Hclose1" with "[]") as "_"; by auto.
- iMod ("Hclose1" with "[]") as "_"; first by auto.
iApply step_fupd_intro; first solve_ndisj. auto.
Qed.
Next Obligation.
iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]".
iExists _. iSplit; first by iApply frac_bor_shorten.
iModIntro. iIntros (F q) "% Htok".
iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj.
iMod ("Hshr" with "[] Htok") as "Hshr2"; first done.
iModIntro. iNext. iMod "Hshr2" as "[Htok HX]".
iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν) "(? & ?)".
iExists _, _. iFrame. by iApply na_bor_shorten.
Qed.
Global Instance weak_wf ty `{!TyWf ty} : TyWf (weak ty) :=
{ ty_lfts := ty_lfts ty; ty_wf_E := ty_wf_E ty }.
Global Instance weak_type_contractive : TypeContractive weak.
Proof.
constructor;
solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
f_type_equiv || f_contractive_fin || f_equiv).
Qed.
Global Instance weak_ne : NonExpansive weak.
Proof. apply type_contractive_ne, _. Qed.
Lemma weak_subtype ty1 ty2 :
type_incl ty1 ty2 -∗ type_incl (weak ty1) (weak ty2).
Proof.
iIntros "#Hincl". iPoseProof "Hincl" as "(Hsz & #Hoincl & #Hsincl)".
iSplit; first done. iSplit; iModIntro.
- iIntros "%tid %vl Hvl". destruct vl as [|[[|vl|]|] [|]]; try done.
iDestruct "Hvl" as (γ ν) "(#Hpersist & Htok)".
iExists _, _. iFrame "Htok". by iApply rc_persist_type_incl.
- iIntros "%κ %tid %l #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'.
iIntros "{$Hfrac} !> * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H".
iModIntro. iNext. iMod "H" as "[$ H]". iDestruct "H" as (γ ν) "(Hpersist & Hshr)".
iExists _, _. iFrame "Hshr". by iApply rc_persist_type_incl.
Qed.
Global Instance weak_mono E L :
Proper (subtype E L ==> subtype E L) weak.
Proof.
iIntros (ty1 ty2 Hsub ??) "HL". iDestruct (Hsub with "HL") as "#Hsub".
iIntros "!> #HE". iApply weak_subtype. iApply "Hsub"; done.
Qed.
Global Instance weak_proper E L :
Proper (eqtype E L ==> eqtype E L) weak.
Proof. intros ??[]. by split; apply weak_mono. Qed.
End weak.
Section code.
Context `{!typeGS Σ, !rcG Σ}.
Definition rc_upgrade : val :=
fn: ["w"] :=
let: "r" := new [ #2 ] in
withcont: "k":
let: "w'" := !"w" in
let: "w''" := !"w'" in
let: "strong" := !("w''" + #0) in
if: "strong" = #0 then
"r" <-{Σ none} ();;
"k" []
else
"w''" + #0 <- "strong" + #1;;
"r" <-{Σ some} "w''";;
"k" []
cont: "k" [] :=
delete [ #1; "w" ];; return: ["r"].
Lemma rc_upgrade_type ty `{!TyWf ty} :
typed_val rc_upgrade (fn( α, ; &shr{α}(weak ty)) option (rc ty)).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>w. simpl_subst.
iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst.
iApply (type_cont [] [ϝ []]
(λ _, [w box (&shr{α}(weak ty)); r box (option (rc ty))])) ;
[solve_typing..| |]; last first.
{ simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iIntros (k). simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (w'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hw [Hw' [Hr _]]]".
rewrite !tctx_hasty_val [[w]]lock.
destruct w' as [[|lw|]|]; try done. iDestruct "Hw'" as (l') "[#Hlw #Hshr]".
iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)".
subst r. inv_vec vlr=>r0 r1. rewrite !heap_pointsto_vec_cons.
iDestruct "Hr" as "(Hr1 & Hr2 & _)".
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlw Hα2") as (q') "[Hlw↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlw↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν) "#(Hpersist & Hwtokb)".
iModIntro. wp_let. wp_op. rewrite shift_loc_0.
(* Finally, finally... opening the thread-local Rc protocol. *)
iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)".
iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|].
iMod (na_bor_acc with "LFT Hwtokb Hα1 Hna") as "(>Hwtok & Hna & Hclose3)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hwtok" gives
%[[_ Hweak%nat_included]%prod_included [Hval _]]%auth_both_valid_discrete.
destruct st as [[[q'' strong]| |]|]; try done.
- (* Success case. *)
iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >Hq''q0 & [Hν1 Hν2] & Hν†)".
iDestruct "Hq''q0" as %Hq''q0.
wp_read. wp_let. wp_op. wp_if. wp_op. rewrite shift_loc_0. wp_op. wp_write.
(* Closing the invariant. *)
iMod (own_update with "Hrc●") as "[Hrc● Hrctok2]".
{ apply auth_update_alloc, prod_local_update_1,
(op_local_update_discrete _ _ (Some (Cinl ((qb/2)%Qp, 1%positive))))=>-[/= Hqa _].
split; simpl; last done. apply frac_valid.
rewrite /= -Hq''q0 comm_L. by apply Qp.add_le_mono_l, Qp.div_le. }
rewrite right_id -Some_op -Cinl_op -pair_op.
iMod ("Hclose3" with "[$Hwtok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hν2 Hν† $Hna]") as "Hna".
{ iExists _. iFrame "Hrc●". iExists _. rewrite Z.add_comm. iFrame.
rewrite [_ _]comm frac_op -[(_ + _)%Qp]assoc Qp.div_2. auto. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ w box (&shr{α}(weak ty)); #lr box (uninit 2);
#l' rc ty ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite 2!tctx_interp_cons tctx_interp_singleton tctx_hasty_val !tctx_hasty_val' //.
unlock. iFrame "Hr† Hw". iSplitL "Hr1 Hr2".
- iExists [_;_].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton. auto with iFrame.
- iRight. auto with iFrame. }
iApply (type_sum_assign (option (rc ty))); [solve_typing..|].
iApply type_jump; solve_typing.
- (* Failure : dropping *)
(* TODO : The two failure cases are almost identical. *)
iDestruct "Hrcst" as "[Hl'1 Hl'2]". wp_read. wp_let. wp_op. wp_if.
(* Closing the invariant. *)
iMod ("Hclose3" with "[$Hwtok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 $Hna]") as "Hna".
{ iExists _. auto with iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ w box (&shr{α}(weak ty)); #lr box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iExists [_;_].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton. auto with iFrame. }
iApply (type_sum_unit (option (rc ty))); [solve_typing..|].
iApply type_jump; solve_typing.
- (* Failure : general case *)
destruct weakc as [|weakc]; first by simpl in *; lia.
iDestruct "Hrcst" as "[Hl'1 Hrcst]". wp_read. wp_let. wp_op. wp_if.
(* Closing the invariant. *)
iMod ("Hclose3" with "[$Hwtok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hrcst $Hna]") as "Hna".
{ iExists _. auto with iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ w box (&shr{α}(weak ty)); #lr box (uninit 2) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iExists [_;_].
rewrite heap_pointsto_vec_cons heap_pointsto_vec_singleton. auto with iFrame. }
iApply (type_sum_unit (option (rc ty))); [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition rc_downgrade : val :=
fn: ["rc"] :=
let: "r" := new [ #1 ] in
let: "rc'" := !"rc" in
let: "rc''" := !"rc'" in
let: "weak" := !("rc''" + #1) in
"rc''" + #1 <- "weak" + #1;;
"r" <- "rc''";;
delete [ #1; "rc" ];; return: ["r"].
Lemma rc_downgrade_type ty `{!TyWf ty} :
typed_val rc_downgrade (fn( α, ; &shr{α}(rc ty)) weak ty).
Proof.
(* TODO : this is almost identical to rc_clone *)
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)".
subst r. inv_vec vlr=>r. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)".
iModIntro. wp_let. wp_op.
(* Finally, finally... opening the thread-local Rc protocol. *)
iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)".
iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|].
iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hrctok" gives %[[[[=]|(?&[[q0 weak0]| |]&[=<-]&?&Hincl)]
%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete;
setoid_subst; try done; last first.
{ exfalso. destruct Hincl as [Hincl|Hincl].
- by inversion Hincl.
- apply csum_included in Hincl. naive_solver. }
iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hrcst)".
wp_read. wp_let. wp_op. wp_op. wp_write. wp_write.
(* And closing it again. *)
iMod (own_update with "Hrc●") as "[Hrc● Hrctok2]".
{ by apply auth_update_alloc, prod_local_update_2, (op_local_update_discrete _ _ 1%nat). }
iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]".
iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hrcst $Hna]") as "Hna".
{ iExists _. iFrame "Hrc●". iExists _.
rewrite !Nat2Z.inj_succ Z.add_1_r. iFrame. }
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL".
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(rc ty)); #lr box (weak ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame "Hx". iFrame "Hr†". iExists [# #l'].
rewrite heap_pointsto_vec_singleton /=. eauto 10 with iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
(* Exact same code as downgrade *)
Definition weak_clone : val :=
fn: ["w"] :=
let: "r" := new [ #1 ] in
let: "w'" := !"w" in
let: "w''" := !"w'" in
let: "weak" := !("w''" + #1) in
"w''" + #1 <- "weak" + #1;;
"r" <- "w''";;
delete [ #1; "w" ];; return: ["r"].
Lemma weak_clone_type ty `{!TyWf ty} :
typed_val weak_clone (fn( α, ; &shr{α}(weak ty)) weak ty).
Proof.
(* TODO : this is almost identical to rc_clone *)
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]".
rewrite !tctx_hasty_val [[x]]lock.
destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)".
subst r. inv_vec vlr=>r. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|].
wp_bind (!_)%E.
iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|].
iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj.
iApply wp_fupd. wp_read.
iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
iDestruct "Hproto" as (γ ν) "#[Hpersist Hwtokb]".
iModIntro. wp_let. wp_op.
(* Finally, finally... opening the thread-local Rc protocol. *)
iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)".
iAssert ( wv : Z, (l' + 1) #wv ((l' + 1) #(wv + 1) ={}=∗
na_own tid (q / 2).[α] own γ weak_tok))%I
with "[> Hna Hα1]" as (wv) "[Hwv Hclose2]".
{ iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|].
iMod (na_bor_acc with "LFT Hwtokb Hα1 Hna") as "(>Hwtok & Hna & Hclose3)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hwtok" gives
%[[_ Hweak%nat_included]%prod_included [Hval _]]%auth_both_valid_discrete.
iMod (own_update with "Hrc●") as "[Hrc● $]".
{ by apply auth_update_alloc, prod_local_update_2,
(op_local_update_discrete _ _ 1%nat). }
destruct st as [[[q'' strong]| |]|]; try done.
- iExists _. iDestruct "Hrcst" as (q0) "(Hl'1 & >$ & Hrcst)".
iIntros "!> Hl'2". iMod ("Hclose3" with "[$Hwtok] Hna") as "[$ Hna]".
iApply ("Hclose2" with "[- $Hna]"). iExists _. iFrame "Hrc●".
iExists _. rewrite /Z.add /= Pos.add_1_r. iFrame.
- iExists _. iDestruct "Hrcst" as "[Hl'1 >$]". iIntros "!> Hl'2".
iMod ("Hclose3" with "[$Hwtok] Hna") as "[$ Hna]".
iApply ("Hclose2" with "[- $Hna]"). iExists _. iFrame "Hrc●".
rewrite /Z.add /= Pos.add_1_r. iFrame.
- destruct weakc as [|weakc]; first by simpl in Hweak; lia.
iExists _. iDestruct "Hrcst" as "(Hl'1 & >$ & Hrcst)". iIntros "!> Hl'2".
iMod ("Hclose3" with "[$Hwtok] Hna") as "[$ Hna]".
iApply ("Hclose2" with "[- $Hna]"). iExists _. iFrame "Hrc●".
rewrite /Z.add /= Pos.add_1_r. iFrame. }
wp_read. wp_let. wp_op. wp_op. wp_write.
iMod ("Hclose2" with "Hwv") as "(Hna & Hα1 & Hwtok)".
iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". wp_write.
(* Finish up the proof. *)
iApply (type_type _ _ _ [ x box (&shr{α}(weak ty)); #lr box (weak ty)]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
unlock. iFrame. iExists [# #l']. rewrite heap_pointsto_vec_singleton /=.
auto 10 with iFrame. }
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing.
Qed.
Definition weak_drop ty : val :=
fn: ["w"] :=
withcont: "k":
let: "w'" := !"w" in
let: "weak" := !("w'" + #1) in
if: "weak" = #1 then
delete [ #(2 + ty.(ty_size)); "w'" ];;
"k" []
else
"w'" + #1 <- "weak" - #1;;
"k" []
cont: "k" [] :=
let: "r" := new [ #0] in
delete [ #1; "w"];; return: ["r"].
Lemma weak_drop_type ty `{!TyWf ty} :
typed_val (weak_drop ty) (fn(; weak ty) unit).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg=>w. simpl_subst.
iApply (type_cont [] [ϝ []] (λ _, [w box (uninit 1)]));
[solve_typing..| |]; last first.
{ simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst.
iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
iApply type_delete; [solve_typing..|].
iApply type_jump; solve_typing. }
iIntros (k). simpl_subst.
iApply type_deref; [solve_typing..|]; iIntros (w'); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hw [Hw' _]]".
rewrite !tctx_hasty_val [[w]]lock. destruct w' as [[|lw|]|]; try done. wp_op.
iDestruct "Hw'" as (γ ν) "[#Hpersist Hwtok]".
iAssert ( wv : Z, (lw + 1) #wv
((wv = 1 na_own tid
s, lw s ((lw + 2) ↦∗: λ vl, length vl = ty.(ty_size))
lw (2 + ty_size ty))
(wv > 1 ((lw + 1) #(wv - 1) ={}=∗ na_own tid ))))%I
with "[>Hna Hwtok]" as (wv) "[Hlw [(% & Hna & H)|[% Hclose]]]".
{ iPoseProof "Hpersist" as (ty') "([>Hszeq _] & Hinv & _ & _)".
iDestruct "Hszeq" as %Hszeq.
iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose)"; [solve_ndisj..|].
iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]".
iCombine "Hrc● Hwtok" gives
%[[_ Hweak%nat_included]%prod_included [Hval _]]%auth_both_valid_discrete.
destruct weakc as [|weakc]; first by simpl in *; lia.
iMod (own_update_2 with "Hrc● Hwtok") as "Hrc●".
{ apply auth_update_dealloc, prod_local_update_2,
(cancel_local_update_unit 1%nat), _. }
destruct st as [[[q'' strong]| |]|]; [..|done|].
- iExists _. iDestruct "Hrcst" as (q0) "(Hlw & >$ & Hrcst)". iRight.
iSplitR; first by auto with lia. iIntros "!>?". iApply "Hclose". iFrame.
by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l.
- iExists _. iDestruct "Hrcst" as "[Hlw >$]". iRight.
iSplitR; first by auto with lia. iIntros "!>?". iApply "Hclose". iFrame.
by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l.
- iExists _. iDestruct "Hrcst" as "(>Hlw & >$ & >H† & >H)". destruct weakc.
+ iLeft. iSplitR; first done. iMod ("Hclose" with "[$Hna Hrc●]") as "$".
{ iExists _. iFrame. }
rewrite Hszeq. auto with iFrame.
+ iRight. iSplitR; first by auto with lia. iIntros "!>?". iApply "Hclose".
iFrame.
by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. }
- subst. wp_read. wp_let. wp_op. wp_if.
iApply (type_type _ _ _ [ w box (uninit 1); #lw box (uninit (2 + 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. iDestruct "H" as (?) "(? & H↦ & ?)". iDestruct "H↦" as (?) "[? %]".
rewrite /= freeable_sz_full_S. iFrame. iExists (_::_::_).
rewrite 2!heap_pointsto_vec_cons shift_loc_assoc. iFrame.
iIntros "!> !%". simpl. congruence. }
iApply type_delete; [try solve_typing..|].
iApply type_jump; solve_typing.
- wp_read. wp_let. wp_op; case_bool_decide; first lia. wp_if. wp_op. wp_op. wp_write.
iMod ("Hclose" with "Hlw") as "Hna".
iApply (type_type _ _ _ [ w box (uninit 1) ]
with "[] LFT HE Hna HL Hk [-]"); last first.
{ unlock. by rewrite tctx_interp_singleton tctx_hasty_val. }
iApply type_jump; solve_typing.
Qed.
Definition weak_new ty : val :=
fn: [] :=
let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in
let: "w" := new [ #1 ] in
"rcbox" + #0 <- #0;;
"rcbox" + #1 <- #1;;
"w" <- "rcbox";;
return: ["w"].
Lemma weak_new_type ty `{!TyWf ty} :
typed_val (weak_new ty) (fn() weak ty).
Proof.
intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>".
iIntros (_ ϝ ret arg). inv_vec arg. simpl_subst.
iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst.
iApply (type_new 1); [solve_typing..|]; iIntros (w); simpl_subst.
iIntros (tid qmax) "#LFT #HE Hna HL Hk [Hw [Hrcbox _]]". rewrite !tctx_hasty_val.
iDestruct (ownptr_uninit_own with "Hrcbox") as (lrcbox vlrcbox)
"(% & Hrcbox↦ & Hrcbox†)". subst rcbox. inv_vec vlrcbox=>??? /=.
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]".
iDestruct (heap_pointsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]".
iDestruct (ownptr_uninit_own with "Hw") as (lw vlw) "(% & Hw↦ & Hw†)". subst w.
inv_vec vlw=>?. rewrite heap_pointsto_vec_singleton.
(* All right, we are done preparing our context. Let's get going. *)
wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write.
iMod (lft_create with "LFT") as (ν) "[Hν #Hν†]"; first done.
iPoseProof ("Hν†" with "Hν") as "H†". wp_bind (_ <- _)%E.
iApply wp_mask_mono; last iApply (wp_step_fupd with "H†"); [set_solver+..|].
wp_write. iIntros "#Hν !>". wp_seq.
iApply (type_type _ _ _ [ #lw box (weak ty)]
with "[] LFT HE Hna HL Hk [> -]"); last first.
{ rewrite tctx_interp_singleton tctx_hasty_val' //=. iFrame.
iExists [_]. rewrite heap_pointsto_vec_singleton. iFrame.
iMod (own_alloc ( _ _)) as (γ) "[??]"; last (iExists _, _; iFrame).
{ apply auth_both_valid_discrete. by split. }
iExists ty. iFrame "Hν†". iSplitR; first by iApply type_incl_refl.
iSplitL; last by iRight. iMod (na_inv_alloc with "[-]") as "$"; last done.
iExists _. iFrame. rewrite freeable_sz_full_S shift_loc_assoc. iFrame.
rewrite length_vec_to_list. auto. }
iApply type_jump; solve_typing.
Qed.
End code.