diff --git a/theories/derived.v b/theories/derived.v index 8d4507f1f60c108a1b884ece8d18a08803a6c9e0..a7c7ad94ce5ee4ff8392f8947580535d00056f9d 100644 --- a/theories/derived.v +++ b/theories/derived.v @@ -12,8 +12,8 @@ Notation SeqCtx e2 := (LetCtx BAnon e2). Notation Skip := (Seq (Lit LitUnit) (Lit LitUnit)). Coercion lit_of_bool : bool >-> base_lit. Notation If e0 e1 e2 := (Case e0 [e2;e1]). -Notation Newlft := (Lit LitUnit) (only parsing). -Notation Endlft := (Seq Skip (Lit LitUnit)) (only parsing). +Definition Newlft := Lit LitUnit. +Definition Endlft := Skip. Section derived. Context `{irisG lrust_lang Σ}. diff --git a/theories/lifetime.v b/theories/lifetime.v index 097ceae61cc8e90dd9d5ecb05ab01d636f81ff1d..aa17d03c8b2f97345be9914f26928a597e39ee16 100644 --- a/theories/lifetime.v +++ b/theories/lifetime.v @@ -130,7 +130,7 @@ Section lft. Qed. Axiom lft_create : - ∀ `(nclose lftN ⊆ E), True ={E}=∗ ∃ κ, 1.[κ] ∗ â–¡ (1.[κ] ={⊤,∅}â–·=∗ [†κ]). + ∀ `(nclose lftN ⊆ E), True ={E}=∗ ∃ κ, 1.[κ] ∗ â–¡ (1.[κ] ={⊤,⊤∖nclose lftN}â–·=∗ [†κ]). Axiom idx_borrow_acc : ∀ `(nclose lftN ⊆ E) q κ i P, diff --git a/theories/perm.v b/theories/perm.v index b925525ff10339e4b778d372ab47ff9c89dae9cf..c6cc6caa36a3962f3fe484ca31ad6a476bc78aa6 100644 --- a/theories/perm.v +++ b/theories/perm.v @@ -24,11 +24,15 @@ Section perm. from_option (λ v, ty.(ty_own) tid [v]) False%I (eval_expr ν). Definition extract (κ : lifetime) (Ï : perm) : perm := - λ tid, ([†κ] ={lftN}=∗ â–· Ï tid)%I. + λ tid, ([†κ] ={lftN}=∗ Ï tid)%I. Definition tok (κ : lifetime) (q : Qp) : perm := λ _, q.[κ]%I. + + Definition killable (κ : lifetime) : perm := + λ _, (â–¡ (1.[κ] ={⊤,⊤∖nclose lftN}â–·=∗ [†κ]))%I. + Definition incl (κ κ' : lifetime) : perm := λ _, (κ ⊑ κ')%I. @@ -55,6 +59,8 @@ Notation "κ ∋ Ï" := (extract κ Ï) Notation "q .[ κ ]" := (tok κ q) (format "q .[ κ ]", at level 0) : perm_scope. +Notation "†κ" := (killable κ) (format "†κ", at level 75). + Infix "⊑" := incl (at level 60) : perm_scope. Notation "∃ x .. y , P" := diff --git a/theories/perm_incl.v b/theories/perm_incl.v index 4ecad12ac60590be2040d3a0d8b084e8c43dfcd8..93de38c14250b758264aea1ac7a08c0b8b45ba77 100644 --- a/theories/perm_incl.v +++ b/theories/perm_incl.v @@ -211,6 +211,17 @@ Section props. (* FIXME : namespaces problem. *) Admitted. + Lemma reborrow_shr_perm_incl κ κ' ν ty : + κ ⊑ κ' ∗ ν â— &shr{κ'}ty ⇒ ν â— &shr{κ}ty. + Proof. + iIntros (tid) "[#Hord #Hκ']". unfold has_type. + destruct (eval_expr ν) as [[[|l|]|]|]; + try by (iDestruct "Hκ'" as "[]" || iDestruct "Hκ'" as (l) "[% _]"). + iDestruct "Hκ'" as (l') "[EQ Hκ']". iDestruct "EQ" as %[=]. subst l'. + iModIntro. iExists _. iSplit. done. + by iApply (ty_shr_mono with "Hord Hκ'"). + Qed. + Lemma borrowing_perm_incl κ Ï Ï1 Ï2 θ : borrowing κ Ï Ï1 Ï2 → Ï âˆ— κ ∋ θ ∗ Ï1 ⇒ Ï2 ∗ κ ∋ (θ ∗ Ï1). Proof. @@ -247,24 +258,14 @@ Section props. iIntros "H†". iMod ("Hextr" with "H†"). simpl. auto. Qed. - Lemma reborrow_shr_perm_incl κ κ' ν ty : - κ ⊑ κ' ∗ ν â— &shr{κ'}ty ⇒ ν â— &shr{κ}ty. - Proof. - iIntros (tid) "[#Hord #Hκ']". unfold has_type. - destruct (eval_expr ν) as [[[|l|]|]|]; - try by (iDestruct "Hκ'" as "[]" || iDestruct "Hκ'" as (l) "[% _]"). - iDestruct "Hκ'" as (l') "[EQ Hκ']". iDestruct "EQ" as %[=]. subst l'. - iModIntro. iExists _. iSplit. done. - by iApply (ty_shr_mono with "Hord Hκ'"). - Qed. - Lemma lftincl_borrowing κ κ' q : borrowing κ ⊤ q.[κ'] (κ ⊑ κ'). Proof. iIntros (tid) "_ Htok". iApply fupd_mask_mono. done. iMod (borrow_create with "[$Htok]") as "[Hbor Hclose]". reflexivity. iMod (borrow_fracture (λ q', (q * q').[κ'])%I with "[Hbor]") as "Hbor". done. { by rewrite Qp_mult_1_r. } - iSplitL "Hbor". iApply frac_borrow_incl. done. eauto. + iSplitL "Hbor". iApply frac_borrow_incl. done. + iIntros "!>H". by iMod ("Hclose" with "H") as ">$". Qed. End props. diff --git a/theories/typing.v b/theories/typing.v index 0c4ee46a848a6ec570987c6bdca13e8a05010041..b8249aa1221b75ccb064d0d6c7483bfb9fada68c 100644 --- a/theories/typing.v +++ b/theories/typing.v @@ -132,21 +132,18 @@ Section typing. typed_step Ï Newlft (λ _, ∃ α, 1.[α] ∗ α ∋ top)%P. Proof. iIntros (tid) "!#(_&_&$)". wp_value. iMod lft_create as (α) "[?#?]". done. - iExists α. iFrame. iApply fupd_mask_mono. done. - iMod (borrow_create with "[]") as "[_?]". reflexivity. 2:by iIntros "!>". eauto. + iExists α. iFrame. iIntros "!>_!>". done. Qed. - (* TODO : lifetime ending permission. *) - (* Lemma typed_endlft κ Ï: *) - (* typed_step (κ ∋ Ï âˆ— 1.[κ]) Endlft (λ _, Ï)%P. *) - (* Proof. *) - (* iIntros (tid) "!#(_&[Hextr Htok]&$)". wp_bind (#() ;; #())%E. *) - (* iApply (wp_wand_r _ _ (λ _, _ ∗ True)%I). iSplitR "Hextr". *) - (* iApply (wp_frame_step_l with "[-]"); try done. *) - (* iDestruct (lft_end with "Hlft Htok") as "$". by wp_seq. *) - (* iIntros (v) "[#Hκ _]". iMod (lft_extract_out with "Hκ Hextr"). done. *) - (* by wp_seq. *) - (* Qed. *) + Lemma typed_endlft κ Ï: + typed_step (κ ∋ Ï âˆ— 1.[κ] ∗ †κ) Endlft (λ _, Ï)%P. + Proof. + iIntros (tid) "!#(_&(Hextr&Htok&Hend)&$)". + iApply wp_fupd. iApply (wp_wand_r _ _ (λ _, _ ∗ True)%I). iSplitR "Hextr". + iApply (wp_frame_step_l with "[-]"); try done. + iDestruct ("Hend" with "Htok") as "$". by wp_seq. + iIntros (v) "[#Hκ _]". by iApply fupd_mask_mono; last iApply "Hextr". + Qed. Lemma typed_alloc Ï (n : nat): 0 < n → typed_step_ty Ï (Alloc #n) (own 1 (uninit n)).