diff --git a/.gitattributes b/.gitattributes new file mode 100644 index 0000000000000000000000000000000000000000..c925e3166ee66e64206fc6ef05efd4bda052ded0 --- /dev/null +++ b/.gitattributes @@ -0,0 +1 @@ +*.v gitlab-language=coq diff --git a/.gitignore b/.gitignore index d51191e19c49af8645e0384ee5f6b1044178492d..43ed78ba1c6e0fa0d16f1ee1df9c105b9ccefe52 100644 --- a/.gitignore +++ b/.gitignore @@ -1,6 +1,10 @@ *.vo *.vio *.v.d +*.vos +*.vok +.coqdeps.d +.Makefile.coq.d *.glob *.cache *.aux diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index e5f276478b9d2fba35b0b05042142c7cdba5e4b8..47783892f4add9ac78367422301089faf0a5bc63 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -31,6 +31,7 @@ build-coq.8.13.2: <<: *template variables: OPAM_PINS: "coq version 8.13.2" + DENY_WARNINGS: "1" tags: - fp-timing diff --git a/Missing.md b/Missing.md index 06065e1b3d84106655881537092d961ae2004132..76971ab8accc7aa8713ce91a7692959fffecc03d 100644 --- a/Missing.md +++ b/Missing.md @@ -3,5 +3,11 @@ Missing APIs from the types we cover (APIs have been added after this formalizat # Cell * Structural conversion for slices. The matching operations in our model would be - `&mut Cell<(A, B)>` -> `(&mut Cell<A>, &mut Cell<B>)` and - `&Cell<(A, B)>` -> `(&Cell<A>, &Cell<B>)`. + `&mut Cell<(A, B)>` -> `&mut (Cell<A>, Cell<B>)` and + `&Cell<(A, B)>` -> `&(Cell<A>, Cell<B>)` (both being NOPs). + * Turns out to be very hard! The way we currently associate NA-masks with locations is in conflict with this. + The invariant for the entire cell gets allocated "on" the first location of the cell, so when we do splitting the 2nd projection has no way to access it... + +# ZST + +* Something like the example from <https://github.com/rust-lang/unsafe-code-guidelines/issues/168#issuecomment-512528361> diff --git a/README.md b/README.md index b70356bfc9550fcfadef073a08bbc7094abd4463..d5dfa6cb3d54bda00d593eae79c12205bf16ae65 100644 --- a/README.md +++ b/README.md @@ -9,8 +9,10 @@ This version is known to compile with: - Coq 8.13.2 - A development version of [GPFSL]. -The easiest way to install the correct versions of the dependencies is through -opam (2.0.0 or newer). You will need the Coq and Iris opam repositories: +## Building from source + +When building from source, we recommend to use opam (2.0.0 or newer) for +installing the dependencies. This requires the following two repositories: opam repo add coq-released https://coq.inria.fr/opam/released opam repo add iris-dev https://gitlab.mpi-sws.org/iris/opam.git @@ -64,6 +66,26 @@ CPU cores. `thread::spawn`, `take_mut::take`, `alias::once` as well as converting `&&T` to `&Box<T>`. +## Changes since original RustBelt publication + +In this section we list fundamental changes to the model that were done since +the publication of the +[original RustBelt paper](https://plv.mpi-sws.org/rustbelt/popl18/). + +### Support for branding + +As part of the [GhostCell paper](http://plv.mpi-sws.org/rustbelt/ghostcell/), +the model was adjusted to support branding. + +* The semantic interpretation of external lifetime contexts had to be changed to use a *syntactic* form of lifetime inclusion. +* This changed interpretation broke the proof of "lifetime equalization". + Instead we prove a weaker rule that only substitutes lifetimes on positions that are compatible with *semantic* lifetime inclusion. + This is good enough for [the example](theories/typing/examples/nonlexical.v). +* Furthermore, we had to redo the proof of `type_call_iris'`, a key lemma involved in calling functions and ensuring that their assumptions about lifetime parameters do indeed hold. + The old proof exploited *semantic* lifetime inclusion in external lifetime contexts in a crucial step. + The proof was fixed by adjusting the semantic interpretation of the local lifetime context. + In particular there is a new parameter `qmax` here that has to be threaded through everywhere. + ## Where to Find the Proof Rules From the Paper ### Type System Rules @@ -103,10 +125,7 @@ borrows are called "unique borrows" in the Coq development. | F-endlft | programs.v | type_endlft | | F-call | function.v | type_call' | -Some of these lemmas are called `something'` because the version -without the `'` is a derived, more specialized form used together with -our eauto-based `solve_typing` tactic. You can see this tactic in -action in the [examples](theories/typing/examples) subfolder. +Some of these lemmas are called `something'` because the version without the `'` is a derived, more specialized form used together with our eauto-based `solve_typing` tactic. You can see this tactic in action in the [examples](theories/typing/examples) subfolder. ### Lifetime Logic Rules diff --git a/_CoqProject b/_CoqProject index 149aadc71298d21346a88155364e6d71646cbd4f..e973f1f0967bf0154d8556ba4de5f9fc5ecab41d 100644 --- a/_CoqProject +++ b/_CoqProject @@ -5,21 +5,22 @@ # (https://github.com/coq/coq/issues/6294). -arg -w -arg -redundant-canonical-projection +theories/lifetime/lifetime_sig.v theories/lifetime/model/boxes.v theories/lifetime/model/definitions.v +theories/lifetime/model/primitive.v theories/lifetime/model/faking.v theories/lifetime/model/creation.v -theories/lifetime/model/primitive.v theories/lifetime/model/accessors.v theories/lifetime/model/borrow.v theories/lifetime/model/borrow_sep.v theories/lifetime/model/reborrow.v theories/lifetime/model/in_at_borrow.v -theories/lifetime/lifetime_sig.v theories/lifetime/lifetime.v theories/lifetime/at_borrow.v theories/lifetime/na_borrow.v theories/lifetime/frac_borrow.v +theories/lifetime/meta.v theories/logic/gps.v theories/lang/notation.v theories/lang/memcpy.v @@ -58,11 +59,13 @@ theories/typing/lib/fake_shared.v theories/typing/lib/cell.v theories/typing/lib/spawn.v theories/typing/lib/join.v -theories/typing/lib/diverging_static.v theories/typing/lib/take_mut.v theories/typing/lib/rc/rc.v theories/typing/lib/rc/weak.v theories/typing/lib/arc.v +theories/typing/lib/swap.v +theories/typing/lib/diverging_static.v +theories/typing/lib/brandedvec.v theories/typing/lib/mutex/mutex.v theories/typing/lib/mutex/mutexguard.v theories/typing/lib/refcell/refcell.v diff --git a/theories/lang/lock.v b/theories/lang/lock.v index b3076cfc543c54b48c0d755bb77b10c965b70953..137f34050a09203623a58e088f1f2c16d429e9e5 100644 --- a/theories/lang/lock.v +++ b/theories/lang/lock.v @@ -27,7 +27,7 @@ Instance subG_lockΣ {Σ} : subG lockΣ Σ → lockG Σ. Proof. solve_inG. Qed. Section proof. - Context `{noprolG Σ, !lftG view_Lat (↑histN) Σ, lockG Σ}. + Context `{noprolG Σ, !lftG Σ view_Lat (↑histN), lockG Σ}. Implicit Types (l : loc). Local Notation vProp := (vProp Σ). diff --git a/theories/lang/notation.v b/theories/lang/notation.v index bab80a461290f7aee648f5607d64799acd46d78f..21e5bb5d10f6ab0c208517d4089b55358ccf4612 100644 --- a/theories/lang/notation.v +++ b/theories/lang/notation.v @@ -1,5 +1,8 @@ From gpfsl.lang Require Export notation. +Notation "e1 < e2" := (e1+#1 ≤ e2)%E + (at level 70) : expr_scope. + Notation Newlft := (Lit LitPoison) (only parsing). Notation Endlft := Skip (only parsing). @@ -16,7 +19,7 @@ Notation "'withcont:' k1 ':' e1 'cont:' k2 xl := e2" := ((Lam (@cons binder k1%binder nil) e1%E) [Rec k2%binder ((fun _ : eq k1%binder k2%binder => xl%binder) eq_refl) e2%E]) (only parsing, at level 151, k1, k2, xl at level 1, e2 at level 150) : expr_scope. -Notation "'call:' f args → k" := (f (@cons expr (λ: ["_r"], Endlft ;; k ["_r"]) args))%E +Notation "'call:' f args → k" := (f (@cons expr (λ: ["_r"], k ["_r"]) args))%E (only parsing, at level 102, f, args, k at level 1) : expr_scope. Notation "'letcall:' x := f args 'in' e" := (letcont: "_k" [ x ] := e in call: f args → "_k")%E diff --git a/theories/lifetime/at_borrow.v b/theories/lifetime/at_borrow.v index bc049b54a374a8cb3f043c3a75c7efc1158079d7..3513d9943a31a3e1f66cf4b52b301275dd94f4a4 100644 --- a/theories/lifetime/at_borrow.v +++ b/theories/lifetime/at_borrow.v @@ -9,20 +9,20 @@ Set Default Proof Using "Type". created, and the payload could be weakened without synchronizing with the end of the lifetime (remember that the underlying payload is not necessarily independent of the view for views older than V). *) -Definition at_bor `{!invG Σ, !lftG Lat E0 Σ} κ N (P : monPred _ _) : monPred _ _ := +Definition at_bor `{!invG Σ, !lftG Σ Lat userE} κ N (P : monPred _ _) : monPred _ _ := (∃ i, ⌜N ## lftN⌠∗ &{κ,i}P ∗ ⎡inv N (∃ V', idx_bor_own i V')⎤)%I. Notation "&at{ κ , N }" := (at_bor κ N) (format "&at{ κ , N }") : bi_scope. Section atomic_bors. - Context `{!invG Σ, !lftG Lat E0 Σ}. + Context `{!invG Σ, !lftG Σ Lat userE}. Global Instance at_bor_ne κ n N : Proper (dist n ==> dist n) (at_bor κ N). Proof. solve_proper. Qed. Global Instance at_bor_contractive κ N : Contractive (at_bor κ N). Proof. solve_contractive. Qed. - Global Instance at_bor_proper κ N : Proper ((≡) ==> (≡)) (at_bor κ N). + Global Instance at_bor_proper κ N : Proper ((⊣⊢) ==> (⊣⊢)) (at_bor κ N). Proof. solve_proper. Qed. - Global Instance at_bor_persistent κ N P : Persistent (&at{κ, N}P) := _. + Global Instance at_bor_persistent κ N P : Persistent (&at{κ, N} P) := _. Lemma at_bor_iff N P P' κ : â–· â–¡ (P ↔ P') -∗ &at{κ, N}P -∗ &at{κ, N}P'. @@ -31,7 +31,8 @@ Section atomic_bors. iExists i. iFrame "#%". iApply (idx_bor_iff with "[HPP'] HP")=>/=. auto. Qed. - Lemma bor_share N P E κ : N ## lftN → &{κ}P ={E}=∗ &at{κ, N}P. + Lemma bor_share N P E κ : + N ## lftN → &{κ}P ={E}=∗ &at{κ, N}P. Proof. iIntros (HN) "HP". rewrite bor_unfold_idx. iDestruct "HP" as (i) "(#?&Hown)". iExists i. iFrame "#%". iMod (inv_alloc with "[Hown]") as "$"; [|done]. iNext. diff --git a/theories/lifetime/frac_borrow.v b/theories/lifetime/frac_borrow.v index e852ec895585f275608d0e6bea3cd9671a07c462..6c112d31666674d9122bac43bf9ddf78fd43a411 100644 --- a/theories/lifetime/frac_borrow.v +++ b/theories/lifetime/frac_borrow.v @@ -6,7 +6,25 @@ Set Default Proof Using "Type". Class frac_borG Σ := frac_borG_inG :> inG Σ fracR. -Definition frac_bor `{!invG Σ, !lftG Lat E0 Σ, !frac_borG Σ} κ +Local Definition frac_bor_inv `{!invG Σ, !lftG Σ Lat userE, !frac_borG Σ} + (φ : Qp → monPred _ _) (V0 : Lat ) γ κ : monPred _ _:= + (∃ (qfresh : Qp) (qdep qused : option Qp), + (* We divide the 1 fraction into 3 parts... *) + ⌜(qfresh â‹…? qdep â‹…? qused = 1)%Qp⌠∗ + (* First part: the piece of φ' which is still fresh. *) + ⎡φ qfresh V0⎤ ∗ + (* Second part: some tokens have been left in deposit while the + fractured borrow is accessed. *) + ⎡from_option (λ qdep', qdep'.[κ]) True qdep V0⎤ ∗ + (* Third part: the piece of φ' that is "already used". It can be given + back when the lifetime ends, but cannot be used when accessing the + fractured borrow. *) + from_option φ True%I qused ∗ + (* We keep a set of "receipts", to prove that we will always have + enough fractions of the lifetime in deposit. *) + ⎡own γ (qfresh â‹…? qused)⎤)%I. + +Definition frac_bor `{!invG Σ, !lftG Σ Lat userE, !frac_borG Σ} κ (φ : Qp → monPred _ _) : monPred _ _ := (* First, we close up to a bunch of things: iff, lifetime inclusion and view inclusion. We also existentially bind γ, the ghost name of receipts. *) @@ -15,39 +33,25 @@ Definition frac_bor `{!invG Σ, !lftG Lat E0 Σ, !frac_borG Σ} κ â–· â–¡ ⎡(∀ q1 q2, φ' (q1+q2)%Qp ↔ φ' q1 ∗ φ' q2) V0⎤ ∗ (* Then we use an (internal) atomic persistent borrow. *) - &in_at{κ'} (∃ (qfresh : Qp) (qdep qused : option Qp), - (* We divide the 1 fraction into 3 parts... *) - ⌜(qfresh â‹…? qdep â‹…? qused = 1)%Qp⌠∗ - (* First part: the piece of φ' which is still fresh. *) - ⎡φ' qfresh V0⎤ ∗ - (* Second part: some tokens have been left in deposit while the - fractured borrow is accessed. *) - ⎡from_option (λ qdep', qdep'.[κ']) True qdep V0⎤ ∗ - (* Third part: the piece of φ' that is "already used". It can be given - back when the lifetime ends, but cannot be used when accessing the - fractured borrow. *) - from_option φ' True%I qused ∗ - (* We keep a set of "receipts", to prove that we will always have - enough fractions of the lifetime in deposit. *) - ⎡own γ (qfresh â‹…? qused)⎤))%I. + &in_at{κ'} (frac_bor_inv φ' V0 γ κ'))%I. Notation "&frac{ κ }" := (frac_bor κ) (format "&frac{ κ }") : bi_scope. Section frac_bor. - Context `{invG Σ, lftG Lat E0 Σ, frac_borG Σ}. + Context `{!invG Σ, !lftG Σ Lat userE, !frac_borG Σ}. Implicit Types E : coPset. Global Instance frac_bor_contractive κ n : Proper (pointwise_relation _ (dist_later n) ==> dist n) (frac_bor κ). - Proof. solve_contractive. Qed. + Proof. rewrite /frac_bor /frac_bor_inv. solve_contractive. Qed. Global Instance frac_bor_ne κ n : Proper (pointwise_relation _ (dist n) ==> dist n) (frac_bor κ). - Proof. solve_proper. Qed. + Proof. rewrite /frac_bor /frac_bor_inv. solve_proper. Qed. Global Instance frac_bor_proper κ : - Proper (pointwise_relation _ (≡) ==> (≡)) (frac_bor κ). - Proof. solve_proper. Qed. + Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (frac_bor κ). + Proof. rewrite /frac_bor /frac_bor_inv. solve_proper. Qed. - Lemma frac_bor_iff κ φ φ' : + Lemma frac_bor_iff κ φ φ' : â–· â–¡ (∀ q, φ q ↔ φ' q) -∗ &frac{κ}φ -∗ &frac{κ}φ'. Proof. iIntros "#Hφφ' H". iDestruct "H" as (γ κ' V0 φ0) "(? & ? & #Hφ0φ & ?)". @@ -69,7 +73,7 @@ Section frac_bor. iDestruct (monPred_in_intro with "HH") as (V) "(#HV & #Hfrac' & Hφ)". iMod ("Hclose" with "[] [-]") as "Hφ"; last first. { iExists γ, κ', V, φ. iFrame "#". iSplitR; [|iSplitR]. - - iIntros "!>!>!#*". by iApply (bi.iff_refl True%I). + - iIntros "!>!>!>*". by iApply (bi.iff_refl True%I). - iModIntro. iNext. iModIntro. iIntros (??). rewrite !bi.intuitionistically_elim. iApply ("Hfrac'" $! _ _). - by iApply (in_at_bor_share with "[Hφ]"). } @@ -88,7 +92,7 @@ Section frac_bor. - iDestruct (monPred_in_intro with "Hfrac") as (V) "[HV Hfrac']". iDestruct "H" as (κ') "(#? & #? & >_)". iExists γ, κ', V, φ. iFrame "#". iSplitR; [|iSplitR]. - + iIntros "!> !> !# *". by iApply (bi.iff_refl True%I). + + iIntros "!> !> !> *". by iApply (bi.iff_refl True%I). + iModIntro. iNext. iModIntro. iIntros (??). rewrite !bi.intuitionistically_elim. iApply ("Hfrac'" $! _ _). + by iApply in_at_bor_fake. @@ -97,7 +101,7 @@ Section frac_bor. ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}P ={E}=∗ &frac{κ}φ. Proof. iIntros (?) "#? ?". rewrite [P]as_fractional. iApply bor_fracture'=>//. - iIntros "!> !# %%". by rewrite fractional -(bi.iff_refl True%I). + iIntros "!> !> %%". by rewrite fractional -(bi.iff_refl True%I). Qed. Lemma frac_bor_acc `{LatBottom Lat} φ E q κ : diff --git a/theories/lifetime/lifetime.v b/theories/lifetime/lifetime.v index e87fcc146f97e63f6b52ccba65dbb34d6bc4bff9..3ca7be4bc2a6c61bb71e91742fdb6161344accb5 100644 --- a/theories/lifetime/lifetime.v +++ b/theories/lifetime/lifetime.v @@ -25,10 +25,23 @@ Instance lft_inhabited : Inhabited lft := populate static. Canonical lftO := leibnizO lft. +Definition lft_incl_syntactic (κ κ' : lft) : Prop := ∃ κ'', κ'' ⊓ κ' = κ. +Infix "⊑ˢʸâ¿" := lft_incl_syntactic (at level 40) : stdpp_scope. + Section derived. -Context `{!invG Σ, !lftG Lat E0 Σ}. +Context `{!invG Σ, !lftG Σ Lat userE}. Implicit Types κ : lft. +Lemma lft_create E `{LatBottom Lat} : + ↑lftN ⊆ E → + ⎡lft_ctx⎤ ={E}=∗ ∃ κ, 1.[κ] ∗ â–¡ (1.[κ] ={↑lftN ∪ userE}[userE]â–·=∗ [†κ]). +Proof. + iIntros (?) "#LFT". + iMod (lft_create_strong (λ _, True) with "LFT") as (κ _) "H"=>//; + [by apply pred_infinite_True|]. + by auto. +Qed. + Lemma lft_tok_dead_subj q κ : q.[κ] -∗ <subj> [†κ] -∗ False. Proof. @@ -48,7 +61,7 @@ Lemma bor_acc_atomic_cons E κ P : ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}P ={E,E∖↑lftN}=∗ (∃ P', â–· (⎡P'⎤ ∧ P) ∗ ∀ Q, - â–· (â–· Q ={E0}=∗ â–· P) -∗ â–· (⎡P'⎤ ∧ Q) ={E∖↑lftN,E}=∗ &{κ}Q) ∨ + â–· (â–· Q ={userE}=∗ â–· P) -∗ â–· (⎡P'⎤ ∧ Q) ={E∖↑lftN,E}=∗ &{κ}Q) ∨ (∃ κ', κ ⊑ κ' ∗ <subj> [†κ'] ∗ |={E∖↑lftN,E}=> True). Proof. iIntros (?) "#LFT HP". @@ -80,7 +93,7 @@ Qed. Lemma bor_acc_cons E q κ P : ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}P -∗ q.[κ] ={E}=∗ â–· P ∗ - ∀ Q, â–· (â–· Q ={E0}=∗ â–· P) -∗ â–· Q ={E}=∗ &{κ}Q ∗ q.[κ]. + ∀ Q, â–· (â–· Q ={userE}=∗ â–· P) -∗ â–· Q ={E}=∗ &{κ}Q ∗ q.[κ]. Proof. iIntros (?) "#LFT HP Htok". iMod (bor_acc_strong with "LFT HP Htok") as (κ') "(#Hκκ' & $ & Hclose)"; first done. @@ -134,22 +147,22 @@ Proof. iMod (bor_exists (A:=bool) with "LFT H") as ([]) "H"; auto. Qed. -Lemma bor_iff κ P P' : - â–· â–¡ (P ↔ P') -∗ &{κ}P -∗ &{κ}P'. +Lemma bor_iff κ P P' : â–· â–¡ (P ↔ P') -∗ &{κ}P -∗ &{κ}P'. Proof. - rewrite !bor_unfold_idx. iIntros "#?". - iDestruct 1 as (i) "[??]". iExists i. iFrame. by iApply idx_bor_iff. + rewrite !bor_unfold_idx. iIntros "#HP Hbor". + iDestruct "Hbor" as (i) "[Hbor Htok]". iExists i. iFrame "Htok". + iApply idx_bor_iff; done. Qed. -Lemma bor_iff_proper κ P P' : - â–· â–¡ (P ↔ P') -∗ â–¡ (&{κ}P ↔ &{κ}P'). +Lemma bor_iff_proper κ P P': â–· â–¡ (P ↔ P') -∗ â–¡ (&{κ}P ↔ &{κ}P'). Proof. - iIntros "#HP !#". iSplit; iIntros "?"; iApply bor_iff; try done. - iNext. iIntros "!#". iSplit; iIntros "?"; iApply "HP"; done. + iIntros "#HP !>". iSplit; iIntros "?"; iApply bor_iff; try done. + iNext. iModIntro. iSplit; iIntros "?"; iApply "HP"; done. Qed. Lemma bor_later `{LatBottom Lat} E κ P : - ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}(â–· P) ={E}[E∖↑lftN]â–·=∗ &{κ}P. + ↑lftN ⊆ E → + ⎡lft_ctx⎤ -∗ &{κ}(â–· P) ={E}[E∖↑lftN]â–·=∗ &{κ}P. Proof. iIntros (?) "#LFT Hb". iMod (bor_acc_atomic_cons with "LFT Hb") as "[H|H]"; first done. @@ -225,7 +238,7 @@ Qed. Lemma lft_incl_static κ : ⊢ κ ⊑ static. Proof. - iApply lft_incl_intro. iIntros "!#" (q) "$". iExists 1%Qp. iSplitR; [|auto]. + iApply lft_incl_intro. iIntros "!>" (q) "$". iExists 1%Qp. iSplitR; [|auto]. by iApply lft_tok_static. Qed. @@ -244,7 +257,7 @@ Proof. iIntros "Hκ". iDestruct (monPred_in_intro with "Hκ") as (V) "[#HV Hκ]". iMod (inv_alloc lftN _ (∃ q, q.[κ] V)%I with "[Hκ]") as "#Hnv". { iExists _. by iNext. } - iApply lft_incl_intro. iIntros "!> !#" (q') "_". + iApply lft_incl_intro. iIntros "!> !>" (q') "_". iInv lftN as ">Hκ" "Hclose". iDestruct "Hκ" as (q'') "[Hκ1 Hκ2]". iMod ("Hclose" with "[Hκ2]") as "_". { iExists _. by iNext. } iModIntro. iExists _. iSplitL; [by iApply monPred_in_elim; auto|]. @@ -257,4 +270,62 @@ Proof. iIntros (?) "#LFT#H†". iApply (in_at_bor_share with "[>]"); try done. by iApply (bor_fake with "LFT H†"). Qed. + +(** Syntactic lifetime inclusion *) +Lemma lft_incl_syn_sem κ κ' : + κ ⊑ˢʸ⿠κ' → ⊢ κ ⊑ κ'. +Proof. intros [z Hxy]. rewrite -Hxy. by apply lft_intersect_incl_r. Qed. + +Lemma lft_intersect_incl_syn_l κ κ' : κ ⊓ κ' ⊑ˢʸ⿠κ. +Proof. by exists κ'; rewrite (comm _ κ κ'). Qed. + +Lemma lft_intersect_incl_syn_r κ κ' : κ ⊓ κ' ⊑ˢʸ⿠κ'. +Proof. by exists κ. Qed. + +Lemma lft_incl_syn_refl κ : κ ⊑ˢʸ⿠κ. +Proof. exists static; by rewrite left_id. Qed. + +Lemma lft_incl_syn_trans κ κ' κ'' : + κ ⊑ˢʸ⿠κ' → κ' ⊑ˢʸ⿠κ'' → κ ⊑ˢʸ⿠κ''. +Proof. + intros [κ0 Hκ0] [κ'0 Hκ'0]. + rewrite -Hκ0 -Hκ'0 assoc. + by apply lft_intersect_incl_syn_r. +Qed. + +Lemma lft_intersect_mono_syn κ1 κ1' κ2 κ2' : + κ1 ⊑ˢʸ⿠κ1' → κ2 ⊑ˢʸ⿠κ2' → (κ1 ⊓ κ2) ⊑ˢʸ⿠(κ1' ⊓ κ2'). +Proof. + intros [κ1'' Hκ1] [κ2'' Hκ2]. + exists (κ1'' ⊓ κ2''). + rewrite -!assoc (comm _ κ2'' _). + rewrite -assoc assoc (comm _ κ2' _). + by f_equal. +Qed. + +Lemma lft_incl_syn_static κ : κ ⊑ˢʸ⿠static. +Proof. + exists κ; by apply lft_intersect_right_id. +Qed. + +Lemma lft_intersect_list_elem_of_incl_syn κs κ : + κ ∈ κs → lft_intersect_list κs ⊑ˢʸ⿠κ. +Proof. + induction 1 as [κ|κ κ' κs Hin IH]. + - apply lft_intersect_incl_syn_l. + - eapply lft_incl_syn_trans; last done. + apply lft_intersect_incl_syn_r. +Qed. + +Global Instance lft_incl_syn_anti_symm : AntiSymm eq (λ x y, x ⊑ˢʸ⿠y). +Proof. + intros κ1 κ2 [κ12 Hκ12] [κ21 Hκ21]. + assert (κ21 = static) as Hstatic. + { apply (lft_intersect_static_cancel_l _ κ12). + eapply (inj (κ1 ⊓.)). + rewrite assoc right_id. + eapply (inj (.⊓ κ2)). + rewrite (comm _ κ1 κ21) -assoc Hκ21 Hκ12 (comm _ κ2). done. } + by rewrite Hstatic left_id in Hκ21. +Qed. End derived. diff --git a/theories/lifetime/lifetime_sig.v b/theories/lifetime/lifetime_sig.v index 765bb4b575ab403ff2b8442234554e60b748181d..7ac073086782b7a45fb7e034e714610c5565f516 100644 --- a/theories/lifetime/lifetime_sig.v +++ b/theories/lifetime/lifetime_sig.v @@ -14,39 +14,49 @@ Module Type lifetime_sig. (** CMRAs needed by the lifetime logic *) (* We can't instantie the module parameters with inductive types, so we have aliases here. *) - Parameter lftG' : ∀ Lat E0, gFunctors → Set. + (** userE is a (disjoint) mask that is available in the "consequence" view + shift of borrow accessors. *) + Parameter lftG' : ∀ (Σ : gFunctors) (Lat : latticeT) (userE : coPset), Set. Global Notation lftG := lftG'. Existing Class lftG'. - Parameter lftPreG' : ∀ Lat, gFunctors → Set. + Parameter lftPreG' : ∀ (Σ : gFunctors) (Lat : latticeT), Set. Global Notation lftPreG := lftPreG'. Existing Class lftPreG'. - (** Definitions *) + (** * Definitions *) Parameter lft : Type. Parameter static : lft. Declare Instance lft_intersect : Meet lft. Section defs. - Context {Σ : gFunctors} {Lat}. - Notation monPred := (monPredI (lat_bi_index Lat) (uPredI (iResUR Σ))). + Notation monPred Σ Lat := (monPredI (lat_bi_index Lat) (uPredI (iResUR Σ))). - Parameter lft_ctx : ∀ `{!invG Σ, !lftG Lat E0 Σ}, iProp Σ. + Parameter lft_ctx : ∀ `{!invG Σ, !lftG Σ Lat userE}, iProp Σ. - Parameter lft_tok : ∀ `{!lftG Lat E0 Σ} (q : Qp) (κ : lft), monPred. - Parameter lft_dead : ∀ `{!lftG Lat E0 Σ} (κ : lft), monPred. + Parameter lft_tok : ∀ `{!lftG Σ Lat userE} (q : Qp) (κ : lft), monPred Σ Lat. + Parameter lft_dead : ∀ `{!lftG Σ Lat userE} (κ : lft), monPred Σ Lat. - Parameter lft_incl : ∀ `{!invG Σ, !lftG Lat E0 Σ} (κ κ' : lft), monPred. - Parameter bor : ∀ `{!invG Σ, !lftG Lat E0 Σ} (κ : lft) (P : monPred), monPred. + Parameter lft_incl : + ∀ `{!invG Σ, !lftG Σ Lat userE} (κ κ' : lft), monPred Σ Lat. + Parameter bor : + ∀ `{!invG Σ, !lftG Σ Lat userE} (κ : lft) (P : monPred Σ Lat), monPred Σ Lat. Parameter bor_idx : Type. - Parameter idx_bor_own : ∀ `{!lftG Lat E0 Σ} (i : bor_idx), monPred. + Parameter idx_bor_own : ∀ `{!lftG Σ Lat userE} (i : bor_idx), monPred Σ Lat. Parameter idx_bor : - ∀ `{!invG Σ, !lftG Lat E0 Σ} (κ : lft) (i : bor_idx) (P : monPred), monPred. + ∀ `{!invG Σ, !lftG Σ Lat userE} (κ : lft) (i : bor_idx) (P : monPred Σ Lat), + monPred Σ Lat. - Parameter in_at_bor : ∀ `{!invG Σ, !lftG Lat E0 Σ} (κ : lft) (P : monPred), monPred. + Parameter in_at_bor : + ∀ `{!invG Σ, !lftG Σ Lat userE} (κ : lft) (P : monPred Σ Lat), monPred Σ Lat. End defs. - (** Notation *) + (** Our lifetime creation lemma offers allocating a lifetime that is defined + by a [positive] in some given infinite set. This operation converts the + [positive] to a lifetime. *) + Parameter positive_to_lft : positive → lft. + + (** * Notation *) Notation "q .[ κ ]" := (lft_tok q κ) (format "q .[ κ ]", at level 0) : bi_scope. Notation "[†κ ]" := (lft_dead κ) (format "[†κ ]"): bi_scope. @@ -58,9 +68,9 @@ Module Type lifetime_sig. Infix "⊑" := lft_incl (at level 70) : bi_scope. Section properties. - Context `{!invG Σ, !lftG Lat E0 Σ}. + Context `{!invG Σ, !lftG Σ Lat userE}. - (** Instances *) + (** * Instances *) Global Declare Instance lft_inhabited : Inhabited lft. Global Declare Instance bor_idx_inhabited : Inhabited bor_idx. @@ -99,7 +109,9 @@ Module Type lifetime_sig. Global Declare Instance lft_tok_as_fractional κ q : AsFractional q.[κ] (λ q, q.[κ])%I q. - (** Laws *) + Global Declare Instance positive_to_lft_inj : Inj eq eq positive_to_lft. + + (** * Laws *) Parameter lft_tok_sep : ∀ q κ1 κ2, q.[κ1] ∗ q.[κ2] ⊣⊢ q.[κ1 ⊓ κ2]. Parameter lft_tok_split_obj : ∀ `{LatBottom Lat} q1 q2 κ, (q1 + q2).[κ] -∗ q1.[κ] ∗ <obj> q2.[κ]. @@ -107,13 +119,19 @@ Module Type lifetime_sig. Parameter lft_tok_dead : ∀ q κ, q.[κ] -∗ [†κ] -∗ False. Parameter lft_tok_static : ∀ q, ⊢ q.[static]. Parameter lft_dead_static : [†static] -∗ False. - - Parameter lft_create : ∀ `{LatBottom Lat} E, ↑lftN ⊆ E → - ⎡lft_ctx⎤ ={E}=∗ ∃ κ, 1.[κ] ∗ â–¡ (1.[κ] ={↑lftN ∪ E0}[E0]â–·=∗ [†κ]). + Parameter lft_intersect_static_cancel_l : ∀ κ κ', κ ⊓ κ' = static → κ = static. + + (** Create a lifetime in some given set of names [P]. This lemma statement + requires exposing [atomic_lft], because [P] restricted to the image of + [atomic_to_lft] might well not be infinite. *) + Parameter lft_create_strong : ∀ `{LatBottom Lat} P E, pred_infinite P → ↑lftN ⊆ E → + ⎡lft_ctx⎤ ={E}=∗ + ∃ p : positive, let κ := positive_to_lft p in ⌜P p⌠∗ + (1).[κ] ∗ â–¡ ((1).[κ] ={↑lftN ∪ userE}[userE]â–·=∗ [†κ]). Parameter bor_create : ∀ `{LatBottom Lat} E κ P, - ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ â–· P ={E}=∗ &{κ}P ∗ ([†κ] ={E}=∗ â–· P). + ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ â–· P ={E}=∗ &{κ} P ∗ ([†κ] ={E}=∗ â–· P). Parameter bor_fake : ∀ `{LatBottom Lat} E κ P, - ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ <subj>[†κ] ={E}=∗ &{κ}P. + ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ <subj>[†κ] ={E}=∗ &{κ}P. Parameter bor_sep : ∀ `{LatBottom Lat} E κ P Q, ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}(P ∗ Q) ={E}=∗ &{κ}P ∗ &{κ}Q. @@ -138,9 +156,11 @@ Module Type lifetime_sig. ∃ V, ⌜V ⊑ Vbor ⊔ Vtok⌠∗ â–· P V ∗ (∀ Vtok', ⌜Vtok ⊑ Vtok'⌠-∗ â–· P (Vtok' ⊔ V) ={E}=∗ idx_bor_own i (Vtok' ⊔ V) ∗ q.[κ] Vtok'). + (* We have to expose κ' here as without [lftN] in the mask of the Q-to-P view + shift, we cannot turn [†κ'] into [†κ]. *) Parameter bor_acc_strong : ∀ E q κ P, ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}P -∗ q.[κ] ={E}=∗ ∃ κ', (κ ⊑ κ') ∗ â–· P ∗ - ∀ Q, â–· (â–· Q -∗ <subj>[†κ'] ={E0}=∗ â–· P) -∗ â–· Q ={E}=∗ &{κ'}Q ∗ q.[κ]. + ∀ Q, â–· (â–· Q -∗ <subj>[†κ'] ={userE}=∗ â–· P) -∗ â–· Q ={E}=∗ &{κ'}Q ∗ q.[κ]. (* Atomic accessor for full borrows. In the context of view predicates, this accessor is a bit @@ -161,7 +181,7 @@ Module Type lifetime_sig. ⎡lft_ctx⎤ -∗ &{κ}P ={E,E∖↑lftN}=∗ (∃ P' κ', κ ⊑ κ' ∗ â–· (⎡P'⎤ ∧ P) ∗ ∀ Q, - â–· (â–· Q -∗ <subj>[†κ'] ={E0}=∗ â–· P) -∗ + â–· (â–· Q -∗ <subj>[†κ'] ={userE}=∗ â–· P) -∗ â–· (⎡P'⎤ ∧ Q) ={E∖↑lftN,E}=∗ &{κ'}Q) ∨ (∃ κ', (κ ⊑ κ') ∗ <subj>[†κ'] ∗ |={E∖↑lftN,E}=> True). @@ -169,9 +189,8 @@ Module Type lifetime_sig. ⎡lft_ctx⎤ -∗ &in_at{κ}P -∗ q.[κ] ={E,E∖↑lftN}=∗ ∃ Vb, (monPred_in Vb → â–· P) ∗ ((monPred_in Vb → â–· P) ={E∖↑lftN,E}=∗ q.[κ]). - (* Because Coq's module system is horrible, we have to repeat - properties of lft_incl here unless we want to prove them twice - (both here and in model.primitive) *) + (* Because Coq's module system is horrible, we have to repeat properties of lft_incl here + unless we want to prove them twice (both here and in model.primitive) *) Parameter lft_intersect_acc : ∀ κ κ' q q', q.[κ] -∗ q'.[κ'] -∗ ∃ q'', q''.[κ ⊓ κ'] ∗ (q''.[κ ⊓ κ'] -∗ q.[κ] ∗ q'.[κ']). Parameter lft_intersect_incl_l : ∀ κ κ', ⊢ κ ⊓ κ' ⊑ κ. @@ -189,9 +208,9 @@ Module Type lifetime_sig. End properties. Parameter lftΣ : latticeT → gFunctors. - Global Declare Instance subG_lftPreG Lat Σ : subG (lftΣ Lat) Σ → lftPreG Lat Σ. + Global Declare Instance subG_lftPreG Lat Σ : subG (lftΣ Lat) Σ → lftPreG Σ Lat. - Parameter lft_init : ∀ `{!invG Σ, !lftPreG Lat Σ} E E0, - ↑lftN ⊆ E → ↑lftN ## E0 → - ⊢ |={E}=> ∃ _ : lftG Lat E0 Σ, lft_ctx. + Parameter lft_init : ∀ `{!invG Σ, !lftPreG Σ Lat} E userE, + ↑lftN ⊆ E → ↑lftN ## userE → + ⊢ |={E}=> ∃ _ : lftG Σ Lat userE, lft_ctx. End lifetime_sig. diff --git a/theories/lifetime/meta.v b/theories/lifetime/meta.v new file mode 100644 index 0000000000000000000000000000000000000000..f6d677a67fdfc3b6a1426dc1f6b2eb29c955c749 --- /dev/null +++ b/theories/lifetime/meta.v @@ -0,0 +1,67 @@ +From iris.algebra Require Import dyn_reservation_map agree. +From iris.proofmode Require Import tactics. +From lrust.lifetime Require Export lifetime. +Set Default Proof Using "Type". + +(** This module provides support for attaching metadata (specifically, a +[gname]) to a lifetime (as is required for types using branding). *) + +Class lft_metaG Σ := LftMetaG { + lft_meta_inG :> inG Σ (dyn_reservation_mapR (agreeR gnameO)); +}. +Definition lft_metaΣ : gFunctors := + #[GFunctor (dyn_reservation_mapR (agreeR gnameO))]. +Instance subG_lft_meta Σ : + subG (lft_metaΣ) Σ → lft_metaG Σ. +Proof. solve_inG. Qed. + +(** We need some global ghost state, but we do not actually care about the name: +we always use a frame-preserving update starting from ε to obtain the ownership +we need. In other words, we use [own_unit] instead of [own_alloc]. As a result +we can just hard-code an arbitrary name here. *) +Local Definition lft_meta_gname : gname := 42%positive. + +Definition lft_meta `{!lftG Σ Lat userE, lft_metaG Σ} (κ : lft) (γ : gname) : iProp Σ := + ∃ p : positive, ⌜κ = positive_to_lft p⌠∗ + own lft_meta_gname (dyn_reservation_map_data p (to_agree γ)). + +Section lft_meta. + Context `{!invG Σ, !lftG Σ Lat userE, lft_metaG Σ}. + + Global Instance lft_meta_timeless κ γ : Timeless (lft_meta κ γ). + Proof. apply _. Qed. + Global Instance lft_meta_persistent κ γ : Persistent (lft_meta κ γ). + Proof. apply _. Qed. + + Lemma lft_create_meta `{LatBottom Lat} {E : coPset} (γ : gname) : + ↑lftN ⊆ E → + ⎡lft_ctx⎤ ={E}=∗ + ∃ κ, ⎡lft_meta κ γ⎤ ∗ (1).[κ] ∗ â–¡ ((1).[κ] ={↑lftN ∪ userE}[userE]â–·=∗ [†κ]). + Proof. + iIntros (HE) "#LFT". + iMod (own_unit (dyn_reservation_mapUR (agreeR gnameO)) lft_meta_gname) as "Hown". + iMod (own_updateP _ _ _ dyn_reservation_map_reserve' with "Hown") + as (? [Etok [Hinf ->]]) "Hown". + iMod (lft_create_strong (.∈ Etok) with "LFT") as (p HEtok) "Hκ"; [done..|]. + iExists (positive_to_lft p). iFrame "Hκ". + iMod (own_update with "Hown") as "Hown". + { eapply (dyn_reservation_map_alloc _ p (to_agree γ)); done. } + iModIntro. iExists p. iModIntro. eauto. + Qed. + + Lemma lft_meta_agree (κ : lft) (γ1 γ2 : gname) : + lft_meta κ γ1 -∗ lft_meta κ γ2 -∗ ⌜γ1 = γ2âŒ. + Proof. + iIntros "Hidx1 Hidx2". + iDestruct "Hidx1" as (p1) "(% & Hidx1)". subst κ. + iDestruct "Hidx2" as (p2) "(Hlft & Hidx2)". + iDestruct "Hlft" as %<-%(inj positive_to_lft). + iCombine "Hidx1 Hidx2" as "Hidx". + iDestruct (own_valid with "Hidx") as %Hval. + rewrite ->(dyn_reservation_map_data_valid (A:=agreeR gnameO)) in Hval. + apply to_agree_op_inv_L in Hval. + done. + Qed. +End lft_meta. + +Typeclasses Opaque lft_meta. diff --git a/theories/lifetime/model/accessors.v b/theories/lifetime/model/accessors.v index ce7933d4d773a855639641f29e49652c2a668541..8f60ff1943ee82df27ce2b0dd92e832c09ff7e26 100644 --- a/theories/lifetime/model/accessors.v +++ b/theories/lifetime/model/accessors.v @@ -4,7 +4,7 @@ From iris.algebra Require Import csum auth excl frac gmap agree gset. Set Default Proof Using "Type". Section accessors. -Context `{!invG Σ, !lftG Lat E0 Σ}. +Context `{!invG Σ, !lftG Σ Lat userE}. Implicit Types κ : lft. (* Internal helper lemmas *) @@ -185,7 +185,7 @@ Proof. - unfold lft_inv. iExists (Vtok' ⊔ Vκ). iSplit; [by eauto using ame_hv_κ|]. iLeft. iSplit. + rewrite lft_inv_alive_unfold. iExists _, _. iFrame. + eauto using ame_lft_alive_in. - - iApply (@big_sepS_impl with "[$Hclose']"). iIntros "!# * % ?". by iApply ame_lft_inv. + - iApply (@big_sepS_impl with "[$Hclose']"). iIntros "!> * % ?". by iApply ame_lft_inv. Qed. (** Indexed borrow *) @@ -199,8 +199,8 @@ Proof. unfold idx_bor. iIntros (HE) "#LFT [#Hle #Hs] Hbor Htok". iDestruct "Hs" as (P') "[#HPP' Hs]". destruct i as [κ' s]. iMod (lft_incl_acc with "Hle Htok") as (q') "[Htok Hclose]". solve_ndisj. - iMod (idx_bor_acc_internal with "LFT Hs Hbor Htok") as (V) "(% & HP & Hclose')". - solve_ndisj. solve_ndisj. + iMod (idx_bor_acc_internal with "LFT Hs Hbor Htok") as (V) "(% & HP & Hclose')"; + [solve_ndisj|solve_ndisj|]. iExists (V ⊔ Vtok). iSplitR; [by iPureIntro; f_equiv|]. iModIntro. iSplitL "HP". - iNext. by iApply "HPP'". - iIntros (Vtok' ?) "HP". iMod ("Hclose'" with "[//] [HP]") as "[$ Htok]". @@ -210,11 +210,11 @@ Qed. (** Basic borrows *) Lemma add_vs Pb Pb' P Q Pi κ : - â–· (Pb ≡ (P ∗ Pb')) -∗ lft_vs κ Pb Pi -∗ (â–· Q -∗ <subj> [†κ] ={E0}=∗ â–· P) -∗ + â–· (Pb ≡ (P ∗ Pb')) -∗ lft_vs κ Pb Pi -∗ (â–· Q -∗ <subj> [†κ] ={userE}=∗ â–· P) -∗ lft_vs κ (Q ∗ Pb') Pi. Proof. iIntros "#HEQ Hvs HvsQ". iApply (lft_vs_cons with "[-Hvs] Hvs"). - iIntros "$ [HQ HPb'] #H†". + iIntros "[HQ HPb'] #H†". iApply fupd_mask_mono; [|iMod ("HvsQ" with "HQ H†") as "HP"]. set_solver. iModIntro. iNext. iRewrite "HEQ". iFrame. @@ -235,7 +235,7 @@ Qed. Lemma bor_acc_strong E q κ P : ↑lftN ⊆ E → ⎡lft_ctx⎤ -∗ &{κ}P -∗ q.[κ] ={E}=∗ ∃ κ', κ ⊑ κ' ∗ â–· P ∗ - ∀ Q, â–· (â–· Q -∗ <subj> [†κ'] ={E0}=∗ â–· P) -∗ â–· Q ={E}=∗ &{κ'}Q ∗ q.[κ]. + ∀ Q, â–· (â–· Q -∗ <subj> [†κ'] ={userE}=∗ â–· P) -∗ â–· Q ={E}=∗ &{κ'}Q ∗ q.[κ]. Proof. unfold bor, raw_bor. iIntros (HE) "#LFT Hbor Htok". iDestruct "Hbor" as (κ'') "(#Hle & Htmp)". iDestruct "Htmp" as (s'') "(Hbor & #Hs)". @@ -292,7 +292,7 @@ Proof. rewrite -fmap_delete lookup_fmap fmap_None -fmap_None -lookup_fmap fmap_delete //. } iDestruct "Hown" as "[Hown Hbor]". - iAssert (â–· (â–· Q -∗ <subj> [†κ''] ={E0}=∗ â–· P'))%I with "[HvsQ]" as "HvsQ". + iAssert (â–· (â–· Q -∗ <subj> [†κ''] ={userE}=∗ â–· P'))%I with "[HvsQ]" as "HvsQ". { iNext. iIntros "HQ H†". iApply "HPP'". iApply ("HvsQ" with "HQ H†"). } iCombine "HV HQ HvsQ" as "HH". iDestruct (monPred_in_intro with "HH") as (V') "(#HV' & % & HQ & HvsQ)". @@ -312,7 +312,7 @@ Proof. iFrame "Hinh". iSplitL "Halive"; [iDestruct "Halive" as (?) "?"; by iExists _|]. iApply (add_vs with "EQ Hvs"). iFrame. + eauto using ame_lft_alive_in. - - iApply (@big_sepS_impl (uPredI _) with "[$Hclose'']"). iIntros "!# % % ?". + - iApply (@big_sepS_impl (uPredI _) with "[$Hclose'']"). iIntros "!> % % ?". by iApply ame_lft_inv. } iDestruct (monPred_in_elim with "HV' Htok") as "Htok". iMod ("Hclose" with "Htok") as "$". iExists κ''. iModIntro. @@ -327,7 +327,7 @@ Lemma bor_acc_atomic_strong E κ P: ⎡lft_ctx⎤ -∗ &{κ}P ={E,E∖↑lftN}=∗ (∃ P' κ', κ ⊑ κ' ∗ â–· (⎡P'⎤ ∧ P) ∗ ∀ Q, - â–· (â–· Q -∗ <subj> [†κ'] ={E0}=∗ â–· P) -∗ + â–· (â–· Q -∗ <subj> [†κ'] ={userE}=∗ â–· P) -∗ â–· (⎡P'⎤ ∧ Q) ={E∖↑lftN,E}=∗ &{κ'}Q) ∨ (∃ κ', κ ⊑ κ' ∗ <subj> [†κ'] ∗ |={E∖↑lftN,E}=> True). @@ -383,7 +383,7 @@ Proof. { iModIntro. unfold bor, raw_bor, idx_bor_own. iExists (i.1). iSplit; first by iApply lft_incl_refl. iExists j. iSplitL "Hâ—¯"; [by iExists _; iFrame|]. - iExists _. iFrame "#". iIntros "!>!#". iSplit; iIntros "H /=". + iExists _. iFrame "#". iIntros "!>!>". iSplit; iIntros "H /=". - by iApply "H". - rewrite /Q'. auto. } iModIntro ⎡_⎤%I. iDestruct (add_vs _ _ _ Q' _ _ with "EQ Hvs [HvsQ]") as "Hvs". diff --git a/theories/lifetime/model/borrow.v b/theories/lifetime/model/borrow.v index 427bd7559ca69bc0f484ba3864d94a4f08120537..9b10f6b51c84d852ae6558790319aa9fdbd8982b 100644 --- a/theories/lifetime/model/borrow.v +++ b/theories/lifetime/model/borrow.v @@ -5,7 +5,7 @@ From iris.proofmode Require Import tactics. Set Default Proof Using "Type*". Section borrow. -Context `{!invG Σ, !lftG Lat E0 Σ, LatBottom Lat}. +Context `{!invG Σ, !lftG Σ Lat userE, LatBottom Lat}. Implicit Types κ : lft. Lemma raw_bor_create E κ P : @@ -52,7 +52,7 @@ Proof. * iExists bot. iFrame. iDestruct (monPred_in_intro True%I with "[//]") as (V') "[H _]". by rewrite -(lat_bottom_sqsubseteq V'). - * iExists _. iIntros "{$HsliceB} !> !#". rewrite /P'. + * iExists _. iIntros "{$HsliceB} !> !>". rewrite /P'. by iSplit; iIntros "H **"; iApply "H". + clear -HE. iIntros "!> H†". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". diff --git a/theories/lifetime/model/borrow_sep.v b/theories/lifetime/model/borrow_sep.v index dfab5d68d7134f6303556c15950348e8bb99bf9d..ef7f0cd1ed6e42b50eccc719c7e5018907f0caab 100644 --- a/theories/lifetime/model/borrow_sep.v +++ b/theories/lifetime/model/borrow_sep.v @@ -24,11 +24,11 @@ Lemma monPred_and_split {M I} (P Q1 Q2 : monPred I (uPredI M)) : Proof. iStartProof (uPred _). iIntros (i). rewrite uPred_and_split. iDestruct 1 as (P1 P2) "(#? & H1 & H2)". iExists P1, P2. - iSplit; [by iIntros (? ->) "!#"|iFrame]. + iSplit; [by iIntros (? ->) "!>"|iFrame]. Qed. Section borrow. -Context `{!invG Σ, !lftG Lat E0 Σ, LatBottom Lat}. +Context `{!invG Σ, !lftG Σ Lat userE, LatBottom Lat}. Implicit Types κ : lft. Lemma bor_sep E κ P Q : @@ -107,7 +107,7 @@ Proof. + by rewrite -fmap_None -lookup_fmap fmap_delete. + by rewrite -(fmap_None bor_filled) -lookup_fmap fmap_insert fmap_delete. - iFrame "Hinh". iNext. iApply (lft_vs_cons with "[] Hvs"). - iIntros (V'' ?) "HD %% (HQ & HP & ?) %% ?". iFrame "HD". iModIntro. iNext. iRewrite "HPb0". + iIntros "%% (HQ & HP & ?) %% ?". iModIntro. iNext. iRewrite "HPb0". iFrame. rewrite !bi.and_or_r. iDestruct "HPP'" as "[_ HPP']". iDestruct "HP" as "[[% HP]|HP]"; iDestruct "HQ" as "[[% HQ]|HQ]". + iApply "HPP'". iSplitL "HP"; [iApply "HP"|iApply "HQ"]; solve_lat. diff --git a/theories/lifetime/model/creation.v b/theories/lifetime/model/creation.v index 1bcedb891ba0f6bf9000e3bab4bf7e8c5c1b8910..e9389716b2a0e8863679e33c262930d99806253b 100644 --- a/theories/lifetime/model/creation.v +++ b/theories/lifetime/model/creation.v @@ -1,11 +1,11 @@ From lrust.lifetime Require Export primitive. From lrust.lifetime Require Import faking. -From iris.algebra Require Import csum auth frac gmap agree gset numbers. +From iris.algebra Require Import csum auth frac gmap agree gset numbers excl. From iris.proofmode Require Import tactics. Set Default Proof Using "Type*". Section creation. -Context `{!invG Σ, !lftG Lat E0 Σ, LatBottom Lat}. +Context `{!invG Σ, !lftG Σ Lat userE, LatBottom Lat}. Implicit Types κ : lft. Lemma lft_kill (I : gmap lft lft_names) (K K' : gset lft) (κ : lft) (Vs : lft → Lat) : @@ -17,7 +17,7 @@ Lemma lft_kill (I : gmap lft lft_names) (K K' : gset lft) (κ : lft) (Vs : lft (∀ κ', κ' ∈ dom (gset _) I → κ ⊂ κ' → κ' ∈ K) → (∀ κ', κ' ∈ dom (gset _) I → κ' ⊂ κ → κ' ∈ K') → Iinv -∗ lft_inv_alive κ (Vs κ) -∗ (∃ V, [†κ] V) - ={↑borN ∪ ↑inhN ∪ E0}=∗ Iinv ∗ lft_inv_dead κ (Vs κ). + ={userE ∪ ↑borN ∪ ↑inhN}=∗ Iinv ∗ lft_inv_dead κ (Vs κ). Proof. iIntros (HVs Iinv HK HK') "(HI & Hdead & Halive) Hlalive Hκ". rewrite lft_inv_alive_unfold; @@ -39,20 +39,20 @@ Proof. rewrite lft_vs_unfold; iDestruct "Hvs" as (Vκ n) "(% & Hcnt & Hvs)". iDestruct (big_sepS_filter_acc (.⊂ κ) _ _ (dom _ I) with "Halive") as "[Halive Halive']"; [by eauto|]. - iApply fupd_trans. iApply (fupd_mask_mono (↑borN ∪ E0)); [set_solver+|]. - iMod ("Hvs" $! I Vs with "[//] [//] [HI Halive Hbox Hbor] HP Hκ") - as "(Hinv & HQ & Hcnt')". - { rewrite lft_vs_inv_unfold. iFrame. rewrite /lft_bor_dead. - iExists ((λ s, match s with Bor_in V => V | _ => bot end) <$> B), P. - rewrite -!map_fmap_compose. iFrame. erewrite map_fmap_ext; [done|]. - by move=>i x /HB [?[-> _]]. } - rewrite lft_vs_inv_unfold; iDestruct "Hinv" as "(?&HI&Halive)". + iApply fupd_trans. iApply (fupd_mask_mono (userE ∪ ↑borN)); [set_solver+|]. + iMod ("Hvs" $! I Vs with "[//] [//] [HI Halive] HP Hκ") as "(Hinv & HQ & Hcnt')". + { rewrite lft_vs_inv_unfold. iFrame. } + rewrite lft_vs_inv_unfold; iDestruct "Hinv" as "(HI&Halive)". iSpecialize ("Halive'" with "Halive"). iMod (own_cnt_update_2 with "Hcnt Hcnt'") as "?". { apply auth_update_dealloc, (nat_local_update _ _ 0 0); lia. } rewrite /Iinv. iFrame "Hdead Halive' HI". iModIntro. iMod (lft_inh_kill with "[$Hinh $HQ]"); first set_solver+. - iModIntro. rewrite /lft_inv_dead. iExists Q, (Vs κ). by iFrame. + iModIntro. rewrite /lft_inv_dead. iExists Q, (Vs κ). iFrame. iSplitR; [done|]. + rewrite /lft_bor_dead. + iExists ((λ st, match st with Bor_in V => V | _ => inhabitant end) <$> B), P. + rewrite -!map_fmap_compose. + unfold to_borUR. rewrite (map_fmap_ext _ Excl); [by iFrame|naive_solver]. Qed. Lemma lfts_kill A (I : gmap lft lft_names) (K K' : gset lft) (Vs : lft → Lat) : @@ -63,7 +63,7 @@ Lemma lfts_kill A (I : gmap lft lft_names) (K K' : gset lft) (Vs : lft → Lat) Iinv K' -∗ ([∗ set] κ ∈ K, (∃ V, [†κ] V) ∗ lft_inv_alive κ (Vs κ) ∗ ⌜@lft_alive_in Lat A κ⌠∨ lft_inv_dead κ (Vs κ) ∗ ⌜lft_dead_in A κâŒ) - ={↑borN ∪ ↑inhN ∪ E0}=∗ Iinv K' ∗ [∗ set] κ ∈ K, lft_inv_dead κ (Vs κ). + ={userE ∪ ↑borN ∪ ↑inhN}=∗ Iinv K' ∗ [∗ set] κ ∈ K, lft_inv_dead κ (Vs κ). Proof. intros Iinv HVs. revert K'. induction (set_wf K) as [K _ IH]=> K' HKK' HKK'I HK. @@ -151,12 +151,16 @@ Proof. rewrite functions.fn_lookup_insert_ne //. Qed. -Lemma lft_create E : - ↑lftN ⊆ E → ⎡lft_ctx⎤ ={E}=∗ ∃ κ, 1.[κ] ∗ â–¡ (1.[κ] ={↑lftN ∪ E0}[E0]â–·=∗ [†κ]). +Lemma lft_create_strong P E : + pred_infinite P → ↑lftN ⊆ E → + ⎡lft_ctx⎤ ={E}=∗ + ∃ p : positive, let κ := positive_to_lft p in ⌜P p⌠∗ + (1).[κ] ∗ â–¡ ((1).[κ] ={↑lftN ∪ userE}[userE]â–·=∗ [†κ]). Proof. - iIntros (?) "#LFT". + assert (userE_lftN_disj:=userE_lftN_disj). iIntros (HP ?) "#LFT". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". - destruct (exist_fresh (dom (gset atomic_lft) A)) as [Λ HΛ%not_elem_of_dom]. + rewrite ->pred_infinite_set in HP. + destruct (HP (dom (gset _) A)) as [Λ [HPx HΛ%not_elem_of_dom]]. iMod (own_update with "HA") as "[HA HΛ]". { apply auth_update_alloc, (alloc_singleton_local_update _ Λ (Cinl (1%Qp, to_latT bot)))=>//. by rewrite lookup_fmap HΛ. } @@ -164,20 +168,20 @@ Proof. { iNext. rewrite /lfts_inv /own_alft_auth. iExists (<[Λ:=(true, bot)]>A), I. rewrite /to_alftUR fmap_insert; iFrame. rewrite big_sepS_impl. iApply "Hinv". - rewrite /lft_inv. iIntros (κ ?) "!# !# H". iDestruct "H" as (Vκ) "[Hhv H]". + rewrite /lft_inv. iIntros (κ ?) "!> !> H". iDestruct "H" as (Vκ) "[Hhv H]". iDestruct "Hhv" as %Hhv. iExists Vκ. iSplit. - iPureIntro. intros Λ' HΛ'. specialize (Hhv Λ' HΛ'). rewrite lookup_insert_ne //. intros <-. by rewrite HΛ in Hhv. - iDestruct "H" as "[[Hκ %]|[Hκ %]]". + iLeft. iFrame "Hκ". iPureIntro. by apply lft_alive_in_insert. + iRight. iFrame "Hκ". iPureIntro. by apply lft_dead_in_insert. } - iModIntro; iExists {[+ Λ +]}. iSplitL. + iModIntro; iExists Λ. iSplit; first done. iSplitL. { rewrite /lft_tok big_sepMS_singleton. iExists _. iFrame. iDestruct (monPred_in_intro True%I with "[//]") as (V) "[H _]". by rewrite -(lat_bottom_sqsubseteq V). } - clear I A HΛ. iIntros "!# HΛ". - iApply (step_fupd_mask_mono (↑lftN ∪ E0) _ ((↑lftN ∪ E0)∖↑mgmtN)); [|done|]. - { assert (↑mgmtN ## E0) by (pose proof E0_lftN_disj; solve_ndisj). set_solver. } + clear I A HΛ. iIntros "!> HΛ". + iApply (step_fupd_mask_mono (↑lftN ∪ userE) _ ((↑lftN ∪ userE)∖↑mgmtN)); [|done|]. + { assert (↑mgmtN ## userE) by (pose proof userE_lftN_disj; solve_ndisj). set_solver. } iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". { rewrite <-union_subseteq_l. solve_ndisj. } iMod (ilft_create _ _ {[+ Λ +]} with "HA HI Hinv") as "H". clear A I. @@ -217,11 +221,15 @@ Proof. + edestruct HK=>//. set_solver + Hal HΛκ HκK HI. - move=>-[[[[??]?]|?][?|?]]//. by destruct (lft_alive_and_dead_in A κ). } rewrite HI !big_sepS_union //. iDestruct "Hinv" as "[HinvD HinvK]". - iApply fupd_trans. iApply (fupd_mask_mono (↑borN ∪ ↑inhN ∪ E0)). - { repeat apply union_least. - - trans (difference (A:=coPset) (↑lftN) (↑mgmtN)); [solve_ndisj|set_solver+]. - - trans (difference (A:=coPset) (↑lftN) (↑mgmtN)); [solve_ndisj|set_solver+]. - - assert (HE0:↑mgmtN ## E0) by (pose proof E0_lftN_disj; solve_ndisj). set_solver+HE0. } + iApply fupd_trans. iApply (fupd_mask_mono (userE ∪ ↑borN ∪ ↑inhN)). + { (* FIXME can we make solve_ndisj handle this? *) + clear -userE_lftN_disj. rewrite -assoc. apply union_least. + - assert (userE ##@{coPset} ↑mgmtN) by solve_ndisj. set_solver. + - assert (↑inhN ##@{coPset} ↑mgmtN) by solve_ndisj. + assert (↑inhN ⊆@{coPset} ↑lftN) by solve_ndisj. + assert (↑borN ##@{coPset} ↑mgmtN) by solve_ndisj. + assert (↑borN ⊆@{coPset} ↑lftN) by solve_ndisj. + set_solver. } iMod (lfts_kill A I K K' Vs with "[$HI HinvD] [HinvK]") as "[[HI HinvD] HinvK]"=>//. { intros κ κ' Hκ Hκ' Hκκ'. assert (κ ∈ dom (gset _) I) by set_solver +HI Hκ. auto. } { move=> κ κ'. rewrite !elem_of_K=>-[?[Hd|?]] ??; split=>//. @@ -229,10 +237,10 @@ Proof. by eapply gmultiset_elem_of_subseteq. - right. by eapply gmultiset_elem_of_subseteq. } { rewrite big_sepS_impl. iApply "HinvD". - iIntros "!# !#" (κ [[? _]_]%elem_of_filter) "[[$ _]|[_ %]]". + iIntros "!> !>" (κ [[? _]_]%elem_of_filter) "[[$ _]|[_ %]]". by edestruct @lft_alive_and_dead_in. } { rewrite big_sepS_impl. iApply "HinvK". - iIntros "!# !#" (κ [_ Hκ]%elem_of_K) "[[H %]|H]"; [iLeft|iRight]; iFrame "H %". + iIntros "!> !>" (κ [_ Hκ]%elem_of_K) "[[H %]|H]"; [iLeft|iRight]; iFrame "H %". rewrite /lft_dead /=. iExists _, _, _. iIntros "{$HΛ} !%". split; [|done]. case Hκ=>// Hd. by destruct (lft_alive_and_dead_in A κ). } iModIntro. iMod ("Hclose" with "[-]") as "_"; last first. @@ -252,11 +260,11 @@ Proof. rewrite /own_alft_auth /to_alftUR fmap_insert. iFrame "HA HI". rewrite HI !big_sepS_union //. iSplitL "HinvD". - rewrite big_sepS_impl. iApply "HinvD". - iIntros "!#" (κ [[??]?]%elem_of_filter) "Halive". + iIntros "!>" (κ [[??]?]%elem_of_filter) "Halive". rewrite /lft_inv. iExists (Vs κ). iSplitR; [by auto|]. iLeft. iFrame "Halive". iPureIntro. by apply lft_alive_in_insert_false. - rewrite big_sepS_impl /lft_inv. iApply "HinvK". - iIntros "!#" (κ [? HκK]%elem_of_K) "Hdead". iExists (Vs κ). iSplitR; [by auto|]. + iIntros "!>" (κ [? HκK]%elem_of_K) "Hdead". iExists (Vs κ). iSplitR; [by auto|]. iRight. iFrame. iPureIntro. destruct HκK. by apply lft_dead_in_insert_false. by apply lft_dead_in_insert_false'. Qed. diff --git a/theories/lifetime/model/definitions.v b/theories/lifetime/model/definitions.v index 37b4fd6ce5183c57b9722bc556a385f512bfb643..a27b1a154cb4761864a8c371011ea97446ed382c 100644 --- a/theories/lifetime/model/definitions.v +++ b/theories/lifetime/model/definitions.v @@ -6,9 +6,9 @@ Set Default Proof Using "Type". Import uPred. Implicit Types (Lat : latticeT) (E : coPset). -Definition borN : namespace := lftN .@ "bor". Definition inhN : namespace := lftN .@ "inh". Definition mgmtN : namespace := lftN .@ "mgmt". +Definition borN : namespace := lftN .@ "bor". Definition atomic_lft := positive. (* HACK : We need to make sure this is not in the top-level context, @@ -21,13 +21,13 @@ End lft_notation. Definition static : lft := (∅ : gmultiset _). Instance lft_intersect : Meet lft := λ κ κ', κ ⊎ κ'. +Definition positive_to_lft (p : positive) : lft := {[+ p +]}. Inductive bor_state {Lat} := | Bor_in (V : Lat) | Bor_open (q : frac) (Vtok : Lat) (V : Lat) | Bor_rebor (κ : lft) (V : Lat). Arguments bor_state : clear implicits. - -Canonical Structure bor_stateO Lat := leibnizO (bor_state Lat). +Canonical bor_stateO Lat := leibnizO (bor_state Lat). Record lft_names := LftNames { bor_name : gname; @@ -42,7 +42,7 @@ Definition ilftUR := gmapUR lft (agreeR lft_namesO). Definition borUR Lat := gmapUR slice_name (exclR (bor_stateO Lat)). Definition inhUR := gset_disjUR slice_name. -Class lftG Lat E0 Σ := LftG { +Class lftG Σ Lat (userE : coPset) := LftG { lft_box :> boxG Lat Σ; alft_inG :> inG Σ (authR (alftUR Lat)); alft_name : gname; @@ -51,11 +51,11 @@ Class lftG Lat E0 Σ := LftG { lft_bor_inG :> inG Σ (authR (borUR Lat)); lft_cnt_inG :> inG Σ (authR natUR); lft_inh_inG :> inG Σ (authR inhUR); - E0_lftN_disj : ↑lftN ## E0; + userE_lftN_disj : ↑lftN ## userE; }. Definition lftG' := lftG. -Class lftPreG Lat Σ := LftPreG { +Class lftPreG Σ Lat := LftPreG { lft_preG_box :> boxG Lat Σ; alft_preG_inG :> inG Σ (authR (alftUR Lat)); ilft_preG_inG :> inG Σ (authR ilftUR); @@ -69,7 +69,7 @@ Definition lftΣ Lat : gFunctors := #[ boxΣ Lat; GFunctor (authR (alftUR Lat)); GFunctor (authR ilftUR); GFunctor (authR (borUR Lat)); GFunctor (authR natUR); GFunctor (authR inhUR) ]. Instance subG_lftPreG Lat Σ : - subG (lftΣ Lat) Σ → lftPreG Lat Σ. + subG (lftΣ Lat) Σ → lftPreG Σ Lat. Proof. solve_inG. Qed. Definition bor_filled {Lat} (s : bor_state Lat) : option Lat := @@ -84,7 +84,7 @@ Definition to_borUR {Lat} : gmap slice_name (bor_state Lat) → borUR Lat := fmap Excl. Section defs. - Context {Σ Lat} `{!invG Σ, !lftG Lat E0 Σ}. + Context `{!invG Σ, !lftG Σ Lat userE}. Notation iProp := (iProp Σ). Notation monPred := (monPred (lat_bi_index Lat) (uPredI (iResUR Σ))). @@ -141,8 +141,7 @@ Section defs. Definition lft_vs_inv_go (κ : lft) (lft_inv_alive : ∀ κ', κ' ⊂ κ → Lat → iProp) (Vs : lft → Lat) (I : gmap lft lft_names) : iProp := - (lft_bor_dead κ ∗ - own_ilft_auth I ∗ + (own_ilft_auth I ∗ [∗ set] κ' ∈ dom _ I, ∀ Hκ : κ' ⊂ κ, lft_inv_alive κ' Hκ (Vs κ'))%I. Definition lft_vs_go (κ : lft) (lft_inv_alive : ∀ κ', κ' ⊂ κ → monPred) @@ -153,7 +152,7 @@ Section defs. ⌜Vκ ⊑ Vs κ⌠-∗ ⌜∀ κ', κ' ∈ dom (gset _) I → κ' ⊆ κ → Vs κ ⊑ Vs κ'⌠-∗ lft_vs_inv_go κ lft_inv_alive Vs I -∗ â–· Pb (Vs κ) -∗ (∃ V, lft_dead κ V) - ={↑borN ∪ E0}=∗ + ={userE ∪ ↑borN}=∗ lft_vs_inv_go κ lft_inv_alive Vs I ∗ â–· Pi (Vs κ) ∗ own_cnt κ (â—¯ n)⎤)%I. Definition lft_inv_alive_go (κ : lft) @@ -237,7 +236,7 @@ Typeclasses Opaque lft_tok lft_dead bor_cnt lft_bor_alive lft_bor_dead idx_bor_own idx_bor raw_bor bor. Section basic_properties. -Context `{!invG Σ, !lftG Lat E0 Σ}. +Context `{!invG Σ, !lftG Σ Lat userE}. Implicit Types κ : lft. Notation iProp := (iProp Σ). @@ -248,7 +247,7 @@ Lemma lft_vs_inv_go_ne κ (f f' : ∀ κ', κ' ⊂ κ → monPred) Vs I n : (∀ κ' (Hκ : κ' ⊂ κ), f κ' Hκ ≡{n}≡ f' κ' Hκ) → lft_vs_inv_go κ f Vs I ≡{n}≡ lft_vs_inv_go κ f' Vs I. Proof. - intros Hf. apply sep_ne, sep_ne, big_opS_ne=> // κ' ?. + intros Hf. apply sep_ne, big_opS_ne=> // κ' ?. apply forall_ne=> Hκ. by rewrite Hf. Qed. @@ -279,7 +278,7 @@ Proof. unfold pointwise_relation. auto using lft_inv_alive_go_ne. Qed. Lemma lft_vs_inv_unfold κ (I : gmap lft lft_names) Vs : - lft_vs_inv κ Vs I ⊣⊢ lft_bor_dead κ ∗ + lft_vs_inv κ Vs I ⊣⊢ own_ilft_auth I ∗ [∗ set] κ' ∈ dom _ I, ⌜κ' ⊂ κ⌠→ lft_inv_alive κ' (Vs κ'). Proof. @@ -291,7 +290,7 @@ Lemma lft_vs_unfold κ Pb Pi : ∀ (I : gmap lft lft_names) (Vs : lft → Lat), ⌜Vκ ⊑ Vs κ⌠-∗ ⌜∀ κ', κ' ∈ dom (gset _) I → κ' ⊆ κ → Vs κ ⊑ Vs κ'⌠-∗ - lft_vs_inv κ Vs I -∗ â–· Pb (Vs κ) -∗ (∃ V, lft_dead κ V) ={↑borN ∪ E0}=∗ + lft_vs_inv κ Vs I -∗ â–· Pb (Vs κ) -∗ (∃ V, lft_dead κ V) ={userE ∪ ↑borN}=∗ lft_vs_inv κ Vs I ∗ â–· Pi (Vs κ) ∗ own_cnt κ (â—¯ n)⎤. Proof. done. Qed. @@ -357,4 +356,11 @@ Proof. rewrite /idx_bor_own. apply _. Qed. Global Instance lft_ctx_persistent : Persistent lft_ctx. Proof. rewrite /lft_ctx. apply _. Qed. + +Global Instance positive_to_lft_inj : Inj eq eq positive_to_lft. +Proof. + unfold positive_to_lft. intros x y Hxy. + apply gmultiset_elem_of_singleton. rewrite -Hxy. + set_solver. +Qed. End basic_properties. diff --git a/theories/lifetime/model/faking.v b/theories/lifetime/model/faking.v index 22409bd2aef568d704817a894677b26e50260bf7..cb4d43b105f9103f621dfec63a3d897b79c840d2 100644 --- a/theories/lifetime/model/faking.v +++ b/theories/lifetime/model/faking.v @@ -4,7 +4,7 @@ From iris.proofmode Require Import tactics. Set Default Proof Using "Type*". Section faking. -Context `{!invG Σ, !lftG Lat E0 Σ, LatBottom Lat bot}. +Context `{!invG Σ, !lftG Σ Lat userE, LatBottom Lat bot}. Implicit Types κ : lft. Lemma ilft_create A I κ : @@ -82,7 +82,7 @@ Proof. * iRight. by iDestruct "Hdeadandalive" as "[$ _]". } iPureIntro=>Λ?. rewrite lookup_union_with lookup_gset_to_gmap option_guard_True; [|by apply gmultiset_elem_of_dom]. destruct (A!!Λ); apply lat_bottom_sqsubseteq. - - iApply (@big_sepS_impl with "[$Hinv]"). iIntros "!# * _ H". unfold lft_inv. + - iApply (@big_sepS_impl with "[$Hinv]"). iIntros "!> * _ H". unfold lft_inv. iDestruct "H" as (Vκ) "[HV Hinv]". iExists Vκ. iSplit. + iDestruct "HV" as %HV. iPureIntro. intros Λ HΛ. specialize (HV Λ HΛ). rewrite lookup_union_with. by destruct (A !! Λ), (gset_to_gmap _ _ !! Λ). @@ -110,7 +110,7 @@ Proof. - iExists γ. iSplitL. + iDestruct (monPred_in_intro True%I with "[//]") as (V) "[? _]". iExists bot. iFrame. by rewrite [bot]lat_bottom_sqsubseteq. - + iExists P. iIntros "{$Hslice} !> !# **". by iApply (bi.iff_refl True%I). + + iExists P. iIntros "{$Hslice} !> !> **". by iApply (bi.iff_refl True%I). Qed. Lemma raw_bor_fake' E κ P : diff --git a/theories/lifetime/model/in_at_borrow.v b/theories/lifetime/model/in_at_borrow.v index cc7eb365bb7991197f98b2ec5ff87f42de20536e..a3459a3936d1456abfbf3e3108758e4553befafc 100644 --- a/theories/lifetime/model/in_at_borrow.v +++ b/theories/lifetime/model/in_at_borrow.v @@ -16,7 +16,7 @@ Set Default Proof Using "Type". Definition in_atN : namespace := lftN .@ "in_at". -Definition in_at_bor {Σ Lat} `{invG Σ, lftG Lat E0 Σ} κ P := +Definition in_at_bor `{invG Σ, lftG Σ Lat userE} κ P := (∃ P' i, (κ ⊑ (i.1)) ∗ â–· â–¡ (P' ↔ P) ∗ ⎡slice borN (i.2) P'⎤ ∗ ⎡inv in_atN (∃ V', idx_bor_own i V')⎤)%I. @@ -24,7 +24,7 @@ Notation "&in_at{ κ }" := (in_at_bor κ) (format "&in_at{ κ }") : bi_scope. Instance in_at_bor_params : Params (@in_at_bor) 5 := {}. Section in_at. -Context `{invG Σ, lftG Lat E0 Σ}. +Context `{invG Σ, lftG Σ Lat userE}. Global Instance in_at_bor_ne κ n : Proper (dist n ==> dist n) (in_at_bor κ). Proof. solve_proper. Qed. @@ -38,7 +38,7 @@ Lemma in_at_bor_iff κ P P' : â–· â–¡ (P ↔ P') -∗ &in_at{κ}P -∗ &in_at{κ}P'. Proof. iIntros "#HPP' H". iDestruct "H" as (P0 i) "#(? & HPP0 & ? & ?)". - iExists P0, i. iFrame "#%". iNext. iIntros "!# **". iSplit; iIntros. + iExists P0, i. iFrame "#%". iNext. iIntros "!> **". iSplit; iIntros. { iApply "HPP'". by iApply "HPP0". } { iApply "HPP0". by iApply "HPP'". } Qed. diff --git a/theories/lifetime/model/primitive.v b/theories/lifetime/model/primitive.v index 999912e1933f9521cdae3ecf137f1b6f8666e398..1e54efe6ed7dca34f5a1021ffaa53f3328e509e8 100644 --- a/theories/lifetime/model/primitive.v +++ b/theories/lifetime/model/primitive.v @@ -5,13 +5,13 @@ From lrust.lifetime.model Require Export definitions boxes. Set Default Proof Using "Type". Import uPred. -Lemma lft_init `{!invG Σ, !lftPreG Lat Σ} E E0 : - ↑lftN ⊆ E → ↑lftN ## E0 → ⊢ |={E}=> ∃ _ : lftG Lat E0 Σ, lft_ctx. +Lemma lft_init `{!invG Σ, !lftPreG Σ Lat} E userE : + ↑lftN ⊆ E → ↑lftN ## userE → ⊢ |={E}=> ∃ _ : lftG Σ Lat userE, lft_ctx. Proof. - iIntros (? HE0). rewrite /lft_ctx. + iIntros (? HuserE). rewrite /lft_ctx. iMod (own_alloc (◠∅ : authR (alftUR Lat))) as (γa) "Ha"; first by apply auth_auth_valid. iMod (own_alloc (◠∅ : authR ilftUR)) as (γi) "Hi"; first by apply auth_auth_valid. - set (HlftG := LftG _ E0 _ _ _ γa _ γi _ _ _ HE0). iExists HlftG. + set (HlftG := LftG _ _ _ _ _ γa _ γi _ _ _ HuserE). iExists HlftG. iMod (inv_alloc _ _ lfts_inv with "[Ha Hi]") as "$"; last done. iModIntro. rewrite /lfts_inv /own_alft_auth /own_ilft_auth. iExists ∅, ∅. rewrite /to_alftUR /to_ilftUR !fmap_empty. iFrame. @@ -19,7 +19,7 @@ Proof. Qed. Section primitive. -Context `{!invG Σ, !lftG Lat E0 Σ}. +Context `{!invG Σ, !lftG Σ Lat userE}. Implicit Types κ : lft. Lemma to_borUR_included (B : gmap slice_name (bor_state Lat)) i s : @@ -298,6 +298,13 @@ Proof. by rewrite /lft_tok big_sepMS_empty. Qed. Lemma lft_dead_static : [†static] -∗ False. Proof. rewrite /lft_dead. iDestruct 1 as (Λ V') "[% H']". set_solver. Qed. +Lemma lft_intersect_static_cancel_l κ κ' : κ ⊓ κ' = static → κ = static. +Proof. + rewrite !gmultiset_eq=>Heq Λ. move:(Heq Λ). + rewrite multiplicity_disj_union multiplicity_empty Nat.eq_add_0. + naive_solver. +Qed. + (* Fractional and AsFractional instances *) Global Instance lft_tok_fractional κ : Fractional (λ q, q.[κ])%I. Proof. @@ -345,7 +352,7 @@ Proof. reflexivity. Qed. Lemma lft_intersect_incl_l κ κ' : ⊢ κ ⊓ κ' ⊑ κ. Proof. - unfold lft_incl. iIntros "!#" (q). rewrite <-lft_tok_sep. iIntros "[H Hf]". + unfold lft_incl. iIntros "!>" (q). rewrite <-lft_tok_sep. iIntros "[H Hf]". iExists q. iIntros "{$H} !> Hf'". rewrite <-lft_tok_sep. by iFrame. Qed. @@ -353,11 +360,11 @@ Lemma lft_intersect_incl_r κ κ' : ⊢ κ ⊓ κ' ⊑ κ'. Proof. rewrite comm. apply lft_intersect_incl_l. Qed. Lemma lft_incl_refl κ : ⊢ κ ⊑ κ. -Proof. unfold lft_incl. iIntros "!#"; auto 10 with iFrame. Qed. +Proof. unfold lft_incl. iIntros "!>"; auto 10 with iFrame. Qed. Lemma lft_incl_trans κ κ' κ'' : (κ ⊑ κ') -∗ (κ' ⊑ κ'') -∗ (κ ⊑ κ''). Proof. - rewrite /lft_incl. iIntros "#H1 #H2 !#" (q) "Hκ". + rewrite /lft_incl. iIntros "#H1 #H2 !>" (q) "Hκ". iMod ("H1" with "Hκ") as (q') "[Hκ' Hclose]". iMod ("H2" with "Hκ'") as (q'') "[Hκ'' Hclose']". iExists q''. iIntros "{$Hκ''} !> Hκ''". @@ -374,7 +381,7 @@ Qed. Lemma lft_incl_glb κ κ' κ'' : κ ⊑ κ' -∗ κ ⊑ κ'' -∗ κ ⊑ κ' ⊓ κ''. Proof. - rewrite /lft_incl. iIntros "#H1 #H2 !#" (q) "[Hκ'1 Hκ'2]". + rewrite /lft_incl. iIntros "#H1 #H2 !>" (q) "[Hκ'1 Hκ'2]". iMod ("H1" with "Hκ'1") as (q') "[Hκ' Hclose']". iMod ("H2" with "Hκ'2") as (q'') "[Hκ'' Hclose'']". iDestruct (lft_intersect_acc with "Hκ' Hκ''") as (qq) "[Hqq Hclose]". @@ -489,7 +496,7 @@ Proof. Qed. Lemma lft_vs_cons κ Pb Pb' Pi : - (⎡lft_bor_dead κ⎤ -∗ â–· Pb' -∗ <subj> [†κ] ={↑borN ∪ E0}=∗ ⎡lft_bor_dead κ⎤ ∗ â–· Pb) -∗ + (â–· Pb' -∗ <subj> [†κ] ={userE ∪ ↑borN}=∗ â–· Pb) -∗ lft_vs κ Pb Pi -∗ lft_vs κ Pb' Pi. Proof. iIntros "Hcons Hvs". rewrite !lft_vs_unfold. @@ -497,13 +504,10 @@ Proof. iCombine "HVκ" "Hcons" as "HVκHcons". iDestruct (monPred_in_intro with "HVκHcons") as (Vκ') "(#HVκ' & % & Hcons)". iExists Vκ', n. iIntros "{$HVκ' $Hnâ—} !>" (I Vs HVVs HVs). - rewrite {1}lft_vs_inv_unfold. iIntros "(Hdead & Hinv & Hκs) HPb #Hκ†". - iAssert ⌜κ ∈ dom (gset _) IâŒ%I as "%". - { rewrite /lft_bor_dead. iDestruct "Hdead" as (??) "[Hdead _]". - iApply (own_bor_auth with "Hinv Hdead"). } - rewrite HVVs. iMod ("Hcons" with "[$Hdead] [HPb //] [Hκ†]") as "[Hdead HPb]". + rewrite {1}lft_vs_inv_unfold. iIntros "(Hinv & Hκs) HPb #Hκ†". + rewrite HVVs. iMod ("Hcons" with "[HPb //] [Hκ†]") as "HPb". - by rewrite monPred_subjectively_unfold. - - iApply ("Hvs" $! I Vs with "[%] [//] [Hdead Hinv Hκs] HPb Hκ†"). + - iApply ("Hvs" $! I Vs with "[%] [//] [Hinv Hκs] HPb Hκ†"). + by etrans. + rewrite lft_vs_inv_unfold. by iFrame. Qed. diff --git a/theories/lifetime/model/reborrow.v b/theories/lifetime/model/reborrow.v index 80592d5d2315c92732dc5cf059c42a71d57593f2..6a20e460ff37e25b9c28fc9f55b78b7365e92f86 100644 --- a/theories/lifetime/model/reborrow.v +++ b/theories/lifetime/model/reborrow.v @@ -17,7 +17,7 @@ Lemma and_extract_own `{inG Σ A} γ (x : A) (P : iProp Σ) : own γ x ∧ P -∗ own γ x ∗ (own γ x -∗ P). Proof. rewrite own.own_eq. apply and_extract_ownM. Qed. -Lemma and_extract_own_bor `{lftG Lat E0 Σ} κ σ (P : iProp Σ) : +Lemma and_extract_own_bor `{lftG Σ Lat userE} κ σ (P : iProp Σ) : own_bor κ σ ∧ P -∗ own_bor κ σ ∗ (own_bor κ σ -∗ P). Proof. rewrite /own_bor bi.and_exist_r bi.exist_elim=>// γ. @@ -29,7 +29,7 @@ Proof. Qed. Section reborrow. -Context `{!invG Σ, !lftG Lat E0 Σ, LatBottom Lat}. +Context `{!invG Σ, !lftG Σ Lat userE, LatBottom Lat}. Implicit Types κ : lft. (* Notice that taking lft_inv for both κ and κ' already implies @@ -102,11 +102,8 @@ Proof. clear dependent Iinv I. iNext. rewrite lft_vs_unfold. iExists Vκ', (S n). iFrame "Hnâ—". iSplit; [done|]. iIntros (I Vs) "% % Hinv [HP HPb] #Hκ†". - rewrite {1}lft_vs_inv_unfold; iDestruct "Hinv" as "(Hκ'dead & HI & Hinv)". + rewrite {1}lft_vs_inv_unfold; iDestruct "Hinv" as "(HI & Hinv)". iDestruct (own_bor_auth with "HI Hi") as %?. - iAssert ⌜κ' ∈ dom (gset _) IâŒ%I as "%" . - { rewrite /lft_bor_dead. iDestruct "Hκ'dead" as (??) "[H _]". - iApply (own_bor_auth with "HI H"). } iDestruct (big_sepS_delete _ _ κ with "Hinv") as "[Hκalive Hinv]"; [done|]. rewrite lft_inv_alive_unfold. iDestruct ("Hκalive" with "[//]") as (Pb' Pi') "(Hκalive & Hvs' & Hinh)". @@ -139,9 +136,9 @@ Proof. { by rewrite lookup_fmap HB. } iModIntro. rewrite big_sepM_delete //. simpl. iDestruct "HB" as "[(% & _ & H1â—¯) HB]". - iMod ("Hvs" $! I Vs with "[%] [//] [Hκ'dead HI Hinv Hvs' Hinh HBâ— Hbox HB] + iMod ("Hvs" $! I Vs with "[%] [//] [HI Hinv Hvs' Hinh HBâ— Hbox HB] [$HPb $Hi] Hκ†") as "($ & $ & Hcnt')"; [solve_lat| |]. - { rewrite lft_vs_inv_unfold. iFrame "Hκ'dead HI". + { rewrite lft_vs_inv_unfold. iFrame "HI". iApply (big_sepS_delete _ (dom (gset lft) I) with "[- $Hinv]"); [done|]. iIntros (_). rewrite lft_inv_alive_unfold. iExists Pb', Pi'. iFrame "Hvs' Hinh". rewrite /lft_bor_alive. @@ -199,7 +196,7 @@ Proof. iExists (delete i' B). rewrite -fmap_delete. iFrame. rewrite fmap_delete -insert_delete delete_insert ?lookup_delete //=. } { iModIntro. iModIntro. iApply (lft_vs_cons with "[] Hvs"). - iIntros (??) "HD %% [??] %% _ !>". iFrame "HD". iNext. iRewrite "HeqPb'". iFrame. } + iIntros "%% [??] %% _ !>". iNext. iRewrite "HeqPb'". iFrame. } iDestruct "HH" as "([HI Hinvκ] & Hb & Halive & Hvs)". iMod ("Hclose" with "[-Hb]"); last first. { iApply monPred_in_elim. done. by iFrame. } diff --git a/theories/lifetime/na_borrow.v b/theories/lifetime/na_borrow.v index 1ee27269f2fa91fd57d5b5511161f6496f2cb620..bf16113dc361f263067ef7250b16fe63350ffc14 100644 --- a/theories/lifetime/na_borrow.v +++ b/theories/lifetime/na_borrow.v @@ -3,7 +3,7 @@ From gpfsl.logic Require Export na_invariants. From iris.proofmode Require Import tactics. Set Default Proof Using "Type". -Definition na_bor `{!invG Σ, !lftG View E0 Σ, na_invG View bot Σ} +Definition na_bor `{!invG Σ, !lftG Σ View userE, na_invG View bot Σ} (κ : lft) (tid : na_inv_pool_name) (N : namespace) (P : monPred _ _) := (∃ i, &{κ,i}P ∗ na_inv tid N (idx_bor_own i))%I. @@ -11,7 +11,7 @@ Notation "&na{ κ , tid , N }" := (na_bor κ tid N) (format "&na{ κ , tid , N }") : bi_scope. Section na_bor. - Context `{!invG Σ, !lftG View E0 Σ, na_invG View bot Σ} + Context `{!invG Σ, !lftG Σ View userE, na_invG View bot Σ} (tid : na_inv_pool_name) (N : namespace). Global Instance na_bor_ne κ n : Proper (dist n ==> dist n) (na_bor κ tid N). diff --git a/theories/logic/gps.v b/theories/logic/gps.v index 1408ae0b4625b47044b103ff8b50f5f2b0bb0bcf..114d6d1ca98f8cbc7ccab23673e3692df22a2890 100644 --- a/theories/logic/gps.v +++ b/theories/logic/gps.v @@ -4,7 +4,7 @@ From lrust.lifetime Require Import at_borrow. From gpfsl.gps Require Import middleware. Section gps_at_bor_SP. -Context `{!noprolG Σ, !gpsG Σ Prtcl, !lftG view_Lat (↑histN) Σ} (N: namespace). +Context `{!noprolG Σ, !gpsG Σ Prtcl, !lftG Σ view_Lat (↑histN)} (N: namespace). Local Notation vProp := (vProp Σ). Implicit Types (IP : interpO Σ Prtcl) (IPC: interpCasO Σ Prtcl) (l : loc) @@ -231,7 +231,7 @@ Qed. End gps_at_bor_SP. Section gps_at_bor_PP. -Context `{!noprolG Σ, !gpsG Σ Prtcl, !gpsExAgG Σ, !lftG view_Lat (↑histN) Σ} +Context `{!noprolG Σ, !gpsG Σ Prtcl, !gpsExAgG Σ, !lftG Σ view_Lat (↑histN)} (N: namespace). Local Notation vProp := (vProp Σ). diff --git a/theories/typing/base.v b/theories/typing/base.v index cd2a0b9b9d909100a9bdfd765314fffd70837145..ab0431c20fb586b3b3b11fa4378fb3bfd40ad2f7 100644 --- a/theories/typing/base.v +++ b/theories/typing/base.v @@ -7,6 +7,10 @@ From lrust.lang Require Export new_delete. Open Scope Z_scope. +(* The "user mask" of the lifetime logic, which can be extended for extensions + of RustBelt. *) +Definition lft_userE : coPset := ↑histN. + Create HintDb lrust_typing. Create HintDb lrust_typing_merge. @@ -20,21 +24,21 @@ Ltac solve_typing := (typeclasses eauto with lrust_typing typeclass_instances core; fail) || (typeclasses eauto with lrust_typing lrust_typing_merge typeclass_instances core; fail). -Hint Constructors Forall Forall2 elem_of_list : lrust_typing. -Hint Resolve submseteq_cons submseteq_inserts_l submseteq_inserts_r +Global Hint Constructors Forall Forall2 elem_of_list : lrust_typing. +Global Hint Resolve submseteq_cons submseteq_inserts_l submseteq_inserts_r : lrust_typing. -Hint Extern 1 (@eq nat _ (Z.to_nat _)) => +Global Hint Extern 1 (@eq nat _ (Z.to_nat _)) => (etrans; [symmetry; exact: Nat2Z.id | reflexivity]) : lrust_typing. -Hint Extern 1 (@eq nat (Z.to_nat _) _) => +Global Hint Extern 1 (@eq nat (Z.to_nat _) _) => (etrans; [exact: Nat2Z.id | reflexivity]) : lrust_typing. -Hint Resolve <- elem_of_app : lrust_typing. +Global Hint Resolve <- elem_of_app : lrust_typing. (* done is there to handle equalities with constants *) -Hint Extern 100 (_ ≤ _) => simpl; first [done|lia] : lrust_typing. -Hint Extern 100 (@eq Z _ _) => simpl; first [done|lia] : lrust_typing. -Hint Extern 100 (@eq nat _ _) => simpl; first [done|lia] : lrust_typing. +Global Hint Extern 100 (_ ≤ _) => simpl; first [done|lia] : lrust_typing. +Global Hint Extern 100 (@eq Z _ _) => simpl; first [done|lia] : lrust_typing. +Global Hint Extern 100 (@eq nat _ _) => simpl; first [done|lia] : lrust_typing. (* Compatibility layer between lrust and gpfsl. *) diff --git a/theories/typing/bool.v b/theories/typing/bool.v index a4ef85d3e2cc3c9ac93cf0eaf02ae30c7fe2485c..436691de6bc6542b0ad0aa22dd8d088daec72653 100644 --- a/theories/typing/bool.v +++ b/theories/typing/bool.v @@ -25,7 +25,7 @@ Section typing. Lemma type_bool_instr (b : Datatypes.bool) : typed_val #b bool. Proof. - iIntros (E L tid) "_ _ $ $ _". iApply wp_value. + iIntros (E L tid ?) "_ _ $ $ _". iApply wp_value. rewrite tctx_interp_singleton tctx_hasty_val' //. by destruct b. Qed. @@ -40,7 +40,7 @@ Section typing. typed_body E L C T e1 -∗ typed_body E L C T e2 -∗ typed_body E L C T (if: p then e1 else e2). Proof. - iIntros (Hp) "He1 He2". iIntros (tid) "#LFT #HE Htl HL HC HT". + iIntros (Hp) "He1 He2". iIntros (tid qmax) "#LFT #HE Htl HL HC HT". iDestruct (big_sepL_elem_of _ _ _ Hp with "HT") as "#Hp". wp_bind p. iApply (wp_hasty with "Hp"); [done|]. iIntros ([[| |[|[]|]]|]) "_ H1"; try done; wp_case. diff --git a/theories/typing/borrow.v b/theories/typing/borrow.v index 46a453127be5601e5e1c89d8784b4d58024aee20..91971e0521e340bd766f0a4e89665d6ed6cbe5da 100644 --- a/theories/typing/borrow.v +++ b/theories/typing/borrow.v @@ -14,7 +14,7 @@ Section borrow. Lemma tctx_borrow E L p n ty κ : tctx_incl E L [p â— own_ptr n ty] [p â— &uniq{κ}ty; p â—{κ} own_ptr n ty]. Proof. - iIntros (tid ?) "#LFT _ $ [H _]". + iIntros (tid ??) "#LFT _ $ [H _]". iDestruct "H" as ([[]|]) "[% Hown]"; try done. iDestruct "Hown" as "[Hmt ?]". iMod (bor_create with "LFT Hmt") as "[Hbor Hext]". done. iModIntro. rewrite /tctx_interp /=. iSplitL "Hbor"; last iSplit; last done. @@ -25,7 +25,7 @@ Section borrow. Lemma tctx_share E L p κ ty : lctx_lft_alive E L κ → tctx_incl E L [p â— &uniq{κ}ty] [p â— &shr{κ}ty]. Proof. - iIntros (Hκ ??) "#LFT #HE HL Huniq". + iIntros (Hκ ???) "#LFT #HE HL Huniq". iMod (Hκ with "HE HL") as (q) "[Htok Hclose]"; [try done..|]. rewrite !tctx_interp_singleton /=. iDestruct "Huniq" as ([[]|]) "[% Huniq]"; try done. @@ -86,8 +86,9 @@ Section borrow. lctx_lft_alive E L κ → ⊢ typed_instruction_ty E L [p â— &uniq{κ}(own_ptr n ty)] (!p) (&uniq{κ} ty). Proof. - iIntros (Hκ tid) "#LFT HE $ HL Hp". + iIntros (Hκ tid qmax) "#LFT HE $ HL Hp". rewrite tctx_interp_singleton. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Hκ with "HE HL") as (q) "[Htok Hclose]"; first solve_ndisj. wp_apply (wp_hasty with "Hp"); [done|]. iIntros ([[]|]) "_ Hown"; try done. iMod (bor_acc_cons with "LFT Hown Htok") as "[H↦ Hclose']". done. @@ -97,7 +98,8 @@ Section borrow. with "[] [H↦ Hown H†]") as "[Hbor Htok]"; last 1 first. - iMod (bor_sep with "LFT Hbor") as "[_ Hbor]". done. iMod (bor_sep with "LFT Hbor") as "[_ Hbor]". done. - iMod ("Hclose" with "Htok") as "($ & $)". + iMod ("Hclose" with "Htok") as "HL". + iDestruct ("HLclose" with "HL") as "$". by rewrite tctx_interp_singleton tctx_hasty_val' //=. - iIntros "!>(?&?&?)!>". iNext. iExists _. rewrite -heap_mapsto_vec_singleton. iFrame. by iFrame. @@ -116,8 +118,9 @@ Section borrow. lctx_lft_alive E L κ → ⊢ typed_instruction_ty E L [p â— &shr{κ}(own_ptr n ty)] (!p) (&shr{κ} ty). Proof. - iIntros (Hκ tid) "#LFT HE $ HL Hp". + iIntros (Hκ tid qmax) "#LFT HE $ HL Hp". rewrite tctx_interp_singleton. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Hκ with "HE HL") as (q) "[[Htok1 Htok2] Hclose]"; first solve_ndisj. wp_apply (wp_hasty with "Hp"); [done|]. iIntros ([[]|]) "_ Hown"; try done. iDestruct "Hown" as (l') "#[H↦b #Hown]". @@ -126,7 +129,8 @@ Section borrow. - iApply ("Hown" with "[%] Htok2"); first solve_ndisj. - iApply wp_fupd. wp_read. iIntros "!>[#Hshr Htok2]". iMod ("Hclose'" with "[H↦]") as "Htok1"; first by auto. - iMod ("Hclose" with "[Htok1 Htok2]") as "($ & $)"; first by iFrame. + iMod ("Hclose" with "[Htok1 Htok2]") as "HL"; first by iFrame. + iDestruct ("HLclose" with "HL") as "$". rewrite tctx_interp_singleton tctx_hasty_val' //. auto. Qed. @@ -142,9 +146,10 @@ Section borrow. lctx_lft_alive E L κ → lctx_lft_incl E L κ κ' → ⊢ typed_instruction_ty E L [p â— &uniq{κ}(&uniq{κ'}ty)] (!p) (&uniq{κ} ty). Proof. - iIntros (Hκ Hincl tid) "#LFT #HE $ HL Hp". + iIntros (Hκ Hincl tid qmax) "#LFT #HE $ HL Hp". rewrite tctx_interp_singleton. - iDestruct (Hincl with "HL HE") as "#Hincl". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". + iDestruct (Hincl with "HL HE") as "%". iMod (Hκ with "HE HL") as (q) "[Htok Hclose]"; first solve_ndisj. wp_apply (wp_hasty with "Hp"); [done|]. iIntros ([[]|]) "_ Hown"; try done. iMod (bor_exists with "LFT Hown") as (vl) "Hbor". done. @@ -156,10 +161,11 @@ Section borrow. iMod (bor_unnest with "LFT Hbor") as "Hbor"; [done|]. iApply wp_fupd. wp_read. iMod ("Hclose'" with "[H↦]") as "[H↦ Htok]"; first by auto. - iMod ("Hclose" with "Htok") as "($ & $)". + iMod ("Hclose" with "Htok") as "HL". + iDestruct ("HLclose" with "HL") as "$". rewrite tctx_interp_singleton tctx_hasty_val' //. iApply (bor_shorten with "[] Hbor"). - iApply (lft_incl_glb with "Hincl"). iApply lft_incl_refl. + iApply lft_incl_glb. by iApply lft_incl_syn_sem. iApply lft_incl_refl. Qed. Lemma type_deref_uniq_uniq {E L} κ κ' x p e ty C T T' : @@ -174,23 +180,27 @@ Section borrow. lctx_lft_alive E L κ → lctx_lft_incl E L κ κ' → ⊢ typed_instruction_ty E L [p â— &shr{κ}(&uniq{κ'}ty)] (!p) (&shr{κ}ty). Proof. - iIntros (Hκ Hincl tid) "#LFT HE $ HL Hp". rewrite tctx_interp_singleton. - iDestruct (Hincl with "HL HE") as "#Hincl". + iIntros (Hκ Hincl tid qmax) "#LFT HE $ HL Hp". rewrite tctx_interp_singleton. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". + iDestruct (Hincl with "HL HE") as "%". iMod (Hκ with "HE HL") as (q) "[[Htok1 Htok2] Hclose]"; first solve_ndisj. wp_apply (wp_hasty with "Hp"); [done|]. iIntros ([[]|]) "_ Hshr"; try done. iDestruct "Hshr" as (l') "[H↦ Hown]". iMod (frac_bor_acc with "LFT H↦ Htok1") as (q'') "[>H↦ Hclose']". done. iAssert (κ ⊑ κ' ⊓ κ)%I as "#Hincl'". - { iApply (lft_incl_glb with "Hincl []"). iApply lft_incl_refl. } + { iApply lft_incl_glb. + + by iApply lft_incl_syn_sem. + + by iApply lft_incl_refl. } iMod (lft_incl_acc with "Hincl' Htok2") as (q2) "[Htok2 Hclose'']"; first solve_ndisj. iApply (wp_step_fupd _ _ (_∖_) with "[Hown Htok2]"); try done. - - iApply ("Hown" with "[%] Htok2"); first solve_ndisj. - - iApply wp_fupd. wp_read. iIntros "!>[#Hshr Htok2]". - iMod ("Hclose''" with "Htok2") as "Htok2". - iMod ("Hclose'" with "[H↦]") as "Htok1"; first by auto. - iMod ("Hclose" with "[Htok1 Htok2]") as "($ & $)"; first by iFrame. - rewrite tctx_interp_singleton tctx_hasty_val' //. - by iApply (ty_shr_mono with "Hincl' Hshr"). + { iApply ("Hown" with "[%] Htok2"); first solve_ndisj. } + iApply wp_fupd. wp_read. iIntros "!>[#Hshr Htok2]". + iMod ("Hclose''" with "Htok2") as "Htok2". + iMod ("Hclose'" with "[H↦]") as "Htok1"; first by auto. + iMod ("Hclose" with "[Htok1 Htok2]") as "HL"; first by iFrame. + iDestruct ("HLclose" with "HL") as "$". + rewrite tctx_interp_singleton tctx_hasty_val' //. + by iApply (ty_shr_mono with "Hincl' Hshr"). Qed. Lemma type_deref_shr_uniq {E L} κ κ' x p e ty C T T' : @@ -202,7 +212,7 @@ Section borrow. Proof. iIntros. iApply type_let; [by apply type_deref_shr_uniq_instr|solve_typing|done]. Qed. End borrow. -Hint Resolve tctx_extract_hasty_borrow tctx_extract_hasty_borrow_share +Global Hint Resolve tctx_extract_hasty_borrow tctx_extract_hasty_borrow_share | 10 : lrust_typing. -Hint Resolve tctx_extract_hasty_share | 10 : lrust_typing. -Hint Resolve tctx_extract_hasty_share_samelft | 9 : lrust_typing. +Global Hint Resolve tctx_extract_hasty_share | 10 : lrust_typing. +Global Hint Resolve tctx_extract_hasty_share_samelft | 9 : lrust_typing. diff --git a/theories/typing/cont.v b/theories/typing/cont.v index edb597f958cf71b6b0a398bc78125d4a981f6503..bbcf9c017015b4515f48473033998ac5b78bf9f6 100644 --- a/theories/typing/cont.v +++ b/theories/typing/cont.v @@ -15,9 +15,11 @@ Section typing. tctx_incl E L T (T' (list_to_vec argsv)) → ⊢ typed_body E L C T (k args). Proof. - iIntros (Hargs HC Hincl tid) "#LFT #HE Hna HL HC HT". + iIntros (Hargs HC Hincl tid qmax) "#LFT #HE Hna HL HC HT". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Hincl with "LFT HE HL HT") as "(HL & HT)". iSpecialize ("HC" with "[]"); first done. + iDestruct ("HLclose" with "HL") as "HL". assert (args = of_val <$> argsv) as ->. { clear -Hargs. induction Hargs as [|a av ?? [<-%of_to_val| ->] _ ->]=>//=. } rewrite -{3}(vec_to_list_to_vec argsv). iApply ("HC" with "Hna HL HT"). @@ -31,7 +33,7 @@ Section typing. (subst_v (kb::argsb) (k:::args) econt)) -∗ typed_body E L2 C T (letcont: kb argsb := econt in e2). Proof. - iIntros (Hc1 Hc2) "He2 #Hecont". iIntros (tid) "#LFT #HE Htl HL HC HT". + iIntros (Hc1 Hc2) "He2 #Hecont". iIntros (tid qmax) "#LFT #HE Htl HL HC HT". rewrite (_ : (rec: kb argsb := econt)%E = of_val (rec: kb argsb := econt)%V); last by unlock. wp_let. iApply ("He2" with "LFT HE Htl HL [HC] HT"). iLöb as "IH". iIntros (x) "H". @@ -47,7 +49,7 @@ Section typing. typed_body E L1 C (T' args) (subst_v (kb :: argsb) (k:::args) econt)) -∗ typed_body E L2 C T (letcont: kb argsb := econt in e2). Proof. - iIntros (Hc1 Hc2) "He2 Hecont". iIntros (tid) "#LFT #HE Htl HL HC HT". + iIntros (Hc1 Hc2) "He2 Hecont". iIntros (tid qmax) "#LFT #HE Htl HL HC HT". rewrite (_ : (rec: kb argsb := econt)%E = of_val (rec: kb argsb := econt)%V); last by unlock. wp_let. iApply ("He2" with "LFT HE Htl HL [HC Hecont] HT"). iIntros (x) "H". diff --git a/theories/typing/cont_context.v b/theories/typing/cont_context.v index ee8bb53dadc5c39401ae90e9ed62af79836e5a10..d6923a3111fdcba2cf1d405265fcdc964b687381 100644 --- a/theories/typing/cont_context.v +++ b/theories/typing/cont_context.v @@ -21,21 +21,21 @@ Notation "k â—cont( L , T )" := (CCtx_iscont k L _ T) Section cont_context. Context `{!typeG Σ}. - Definition cctx_elt_interp (tid : type.thread_id) (x : cctx_elt) : vProp Σ := + Definition cctx_elt_interp (tid : type.thread_id) (qmax: Qp) (x : cctx_elt) : vProp Σ := let '(k â—cont(L, T)) := x in (∀ args, na_own tid top -∗ - llctx_interp L 1 -∗ tctx_interp tid (T args) -∗ + llctx_interp qmax L -∗ tctx_interp tid (T args) -∗ WP k (base.of_val <$> (args : list _)) in tid {{ _, cont_postcondition }})%I. - Definition cctx_interp (tid : type.thread_id) (C : cctx) : vProp Σ := - (∀ (x : cctx_elt), ⌜x ∈ C⌠-∗ cctx_elt_interp tid x)%I. + Definition cctx_interp (tid : type.thread_id) (qmax: Qp) (C : cctx) : vProp Σ := + (∀ (x : cctx_elt), ⌜x ∈ C⌠-∗ cctx_elt_interp tid qmax x)%I. - Global Instance cctx_interp_permut tid: - Proper ((≡ₚ) ==> (⊣⊢)) (cctx_interp tid). + Global Instance cctx_interp_permut tid qmax : + Proper ((≡ₚ) ==> (⊣⊢)) (cctx_interp tid qmax). Proof. solve_proper. Qed. - Lemma cctx_interp_cons tid x T : - cctx_interp tid (x :: T) ≡ (cctx_elt_interp tid x ∧ cctx_interp tid T)%I. + Lemma cctx_interp_cons tid qmax x T : + cctx_interp tid qmax (x :: T) ≡ (cctx_elt_interp tid qmax x ∧ cctx_interp tid qmax T)%I. Proof. iSplit. - iIntros "H". iSplit; [|iIntros (??)]; iApply "H"; rewrite elem_of_cons; auto. @@ -44,22 +44,22 @@ Section cont_context. + iDestruct "H" as "[_ H]". by iApply "H". Qed. - Lemma cctx_interp_nil tid : cctx_interp tid [] ≡ True%I. + Lemma cctx_interp_nil tid qmax : cctx_interp tid qmax [] ≡ True%I. Proof. iSplit; first by auto. iIntros "_". iIntros (? Hin). inversion Hin. Qed. - Lemma cctx_interp_app tid T1 T2 : - cctx_interp tid (T1 ++ T2) ≡ (cctx_interp tid T1 ∧ cctx_interp tid T2)%I. + Lemma cctx_interp_app tid qmax T1 T2 : + cctx_interp tid qmax (T1 ++ T2) ≡ (cctx_interp tid qmax T1 ∧ cctx_interp tid qmax T2)%I. Proof. induction T1 as [|?? IH]; first by rewrite /= cctx_interp_nil left_id. by rewrite /= !cctx_interp_cons IH assoc. Qed. - Lemma cctx_interp_singleton tid x : - cctx_interp tid [x] ≡ cctx_elt_interp tid x. + Lemma cctx_interp_singleton tid qmax x : + cctx_interp tid qmax [x] ≡ cctx_elt_interp tid qmax x. Proof. rewrite cctx_interp_cons cctx_interp_nil right_id //. Qed. - Lemma fupd_cctx_interp tid C : - (|={⊤}=> cctx_interp tid C) -∗ cctx_interp tid C. + Lemma fupd_cctx_interp tid qmax C : + (|={⊤}=> cctx_interp tid qmax C) -∗ cctx_interp tid qmax C. Proof. rewrite /cctx_interp. iIntros "H". iIntros ([k L n T]) "%". iIntros (args) "Hna HL HT". iMod "H". @@ -67,30 +67,32 @@ Section cont_context. Qed. Definition cctx_incl (E : elctx) (C1 C2 : cctx): Prop := - ∀ tid, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ cctx_interp tid C1 -∗ cctx_interp tid C2. + ∀ tid qmax, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ cctx_interp tid qmax C1 -∗ cctx_interp tid qmax C2. Global Instance cctx_incl_preorder E : PreOrder (cctx_incl E). Proof. split. - - iIntros (??) "_ _ $". - - iIntros (??? H1 H2 ?) "#LFT #HE HC". + - iIntros (???) "_ _ $". + - iIntros (??? H1 H2 ??) "#LFT #HE HC". iApply (H2 with "LFT HE"). by iApply (H1 with "LFT HE"). Qed. Lemma incl_cctx_incl E C1 C2 : C1 ⊆ C2 → cctx_incl E C2 C1. Proof. - rewrite /cctx_incl /cctx_interp. iIntros (HC1C2 tid) "_ #HE H * %". + rewrite /cctx_incl /cctx_interp. iIntros (HC1C2 tid ?) "_ #HE H * %". iApply ("H" with "[%]"). by apply HC1C2. Qed. - Lemma cctx_incl_cons_match E k L n (T1 T2 : vec val n → tctx) C1 C2 : + Lemma cctx_incl_cons E k L n (T1 T2 : vec val n → tctx) C1 C2 : cctx_incl E C1 C2 → (∀ args, tctx_incl E L (T2 args) (T1 args)) → cctx_incl E (k â—cont(L, T1) :: C1) (k â—cont(L, T2) :: C2). Proof. - iIntros (HC1C2 HT1T2 ?) "#LFT #HE H". rewrite /cctx_interp. + iIntros (HC1C2 HT1T2 ??) "#LFT #HE H". rewrite /cctx_interp. iIntros (x) "Hin". iDestruct "Hin" as %[->|Hin]%elem_of_cons. - iIntros (args) "Htl HL HT". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (HT1T2 with "LFT HE HL HT") as "(HL & HT)". + iDestruct ("HLclose" with "HL") as "HL". iApply ("H" $! (_ â—cont(_, _)) with "[%] Htl HL HT"). constructor. - iApply (HC1C2 with "LFT HE [H] [%]"); last done. @@ -100,18 +102,19 @@ Section cont_context. Lemma cctx_incl_nil E C : cctx_incl E C []. Proof. apply incl_cctx_incl. by set_unfold. Qed. - Lemma cctx_incl_cons E k L n (T1 T2 : vec val n → tctx) C1 C2 : + (* Extra strong cctx inclusion rule that we do not have on paper. *) + Lemma cctx_incl_cons_dup E k L n (T1 T2 : vec val n → tctx) C1 C2 : k â—cont(L, T1) ∈ C1 → (∀ args, tctx_incl E L (T2 args) (T1 args)) → cctx_incl E C1 C2 → cctx_incl E C1 (k â—cont(L, T2) :: C2). Proof. - intros Hin ??. rewrite <-cctx_incl_cons_match; try done. - iIntros (?) "_ #HE HC". + intros Hin ??. rewrite <-cctx_incl_cons; try done. + clear -Hin. iIntros (??) "_ #HE HC". rewrite cctx_interp_cons. iSplit; last done. clear -Hin. iInduction Hin as [] "IH"; rewrite cctx_interp_cons; [iDestruct "HC" as "[$ _]" | iApply "IH"; iDestruct "HC" as "[_ $]"]. Qed. End cont_context. -Hint Resolve cctx_incl_nil cctx_incl_cons : lrust_typing. +Global Hint Resolve cctx_incl_nil cctx_incl_cons : lrust_typing. diff --git a/theories/typing/examples/get_x.v b/theories/typing/examples/get_x.v index 4e66d19cd4c24c73a941072441ddbdbed2d86d26..58c0a6710a7b7cce326f4a961f33ef5852a9c40d 100644 --- a/theories/typing/examples/get_x.v +++ b/theories/typing/examples/get_x.v @@ -13,7 +13,7 @@ Section get_x. Lemma get_x_type : typed_val get_x (fn(∀ α, ∅; &uniq{α}(Î [int; int])) → &shr{α}int). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret p). + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret p). inv_vec p=>p. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (p'); simpl_subst. iApply (type_letalloc_1 (&shr{α}int)); [solve_typing..|]. iIntros (r). simpl_subst. diff --git a/theories/typing/examples/init_prod.v b/theories/typing/examples/init_prod.v index 6672bf91cd173ad84797469339b89d14e0f8139d..0a4187ec1338e529c2b427918733162f10262ace 100644 --- a/theories/typing/examples/init_prod.v +++ b/theories/typing/examples/init_prod.v @@ -13,7 +13,7 @@ Section init_prod. Lemma init_prod_type : typed_val init_prod (fn(∅; int, int) → Î [int;int]). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([] Ï ret p). inv_vec p=>x y. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (y'). simpl_subst. diff --git a/theories/typing/examples/lazy_lft.v b/theories/typing/examples/lazy_lft.v index c7b0777ad2a02d5333e786c949b1a7ccf29879bd..6b5416ece24865e0c68f9cd9724b30050197a77a 100644 --- a/theories/typing/examples/lazy_lft.v +++ b/theories/typing/examples/lazy_lft.v @@ -18,7 +18,7 @@ Section lazy_lft. Lemma lazy_lft_type : typed_val lazy_lft (fn(∅) → unit). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([] Ï ret p). inv_vec p. simpl_subst. iApply (type_newlft []). iIntros (α). iApply (type_new_subtype (Î [uninit 1;uninit 1])); [solve_typing..|]. diff --git a/theories/typing/examples/nonlexical.v b/theories/typing/examples/nonlexical.v index 4f7623e2219b2e56fe61acb60903cfb2f7da8041..bef584f9ac4a6029bf223f97be178d4e4c4712d9 100644 --- a/theories/typing/examples/nonlexical.v +++ b/theories/typing/examples/nonlexical.v @@ -74,7 +74,7 @@ Section non_lexical. Lemma get_default_type : typed_val get_default (fn(∀ m, ∅; &uniq{m} hashmap, K) → &uniq{m} V). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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. @@ -115,11 +115,20 @@ Section non_lexical. iIntros (unwrap'). simpl_subst. iApply (type_letcall ()); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_jump; solve_typing. - - iApply type_equivalize_lft. + - (* 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 "!# *". inv_vec args=>r. simpl_subst. + iIntros "!> *". inv_vec args=>r. simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_delete; [solve_typing..|]. iApply type_delete; [solve_typing..|]. diff --git a/theories/typing/examples/rebor.v b/theories/typing/examples/rebor.v index ba08aef83d14f45f1cb695462c9cb47406005c40..42f1c255e6f2a5db5c894355c9eeb123d0ec824b 100644 --- a/theories/typing/examples/rebor.v +++ b/theories/typing/examples/rebor.v @@ -18,7 +18,7 @@ Section rebor. Lemma rebor_type : typed_val rebor (fn(∅; Î [int; int], Î [int; int]) → int). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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. diff --git a/theories/typing/examples/unbox.v b/theories/typing/examples/unbox.v index b954b686bdd9191ddab695b8e32637bcce090509..2815771a313ecc5586c9f51f941cc5069e746852 100644 --- a/theories/typing/examples/unbox.v +++ b/theories/typing/examples/unbox.v @@ -13,7 +13,7 @@ Section unbox. Lemma ubox_type : typed_val unbox (fn(∀ α, ∅; &uniq{α}(box(Î [int; int]))) → &uniq{α}int). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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. diff --git a/theories/typing/fixpoint.v b/theories/typing/fixpoint.v index c039ac08ccb60d202fdd4fea99d549b712a4fea3..dc64f4ae91a499cc682e52b0c69f5f235a010274 100644 --- a/theories/typing/fixpoint.v +++ b/theories/typing/fixpoint.v @@ -93,7 +93,7 @@ Section fixpoint. Proof. unfold eqtype, subtype, type_incl. setoid_rewrite <-fixpointK_unfold; [| by apply type_contractive_ne, _..]. - split; iIntros (qL) "_ !# _"; (iSplit; first done); iSplit; iIntros "!#*$". + split; iIntros (qmax qL) "_ !> _"; (iSplit; first done); iSplit; iIntros "!>*$". Qed. End fixpoint. @@ -141,11 +141,11 @@ Section subtyping. Proof. intros. by rewrite fixpoint_unfold_eqtype. Qed. End subtyping. -Hint Resolve fixpoint_mono fixpoint_proper : lrust_typing. +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 *) -Hint Resolve fixpoint_unfold_subtype_l|100 : lrust_typing. -Hint Resolve fixpoint_unfold_subtype_r|100 : lrust_typing. -Hint Resolve fixpoint_unfold_eqtype_l|100 : lrust_typing. -Hint Resolve fixpoint_unfold_eqtype_r|100 : lrust_typing. +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. diff --git a/theories/typing/function.v b/theories/typing/function.v index b0fe6d6dbf7674a069eb3e8e1dc7b8c47d6d7658..9f1325a986e5d30eb5edbe369ee422510ba7aaad 100644 --- a/theories/typing/function.v +++ b/theories/typing/function.v @@ -82,7 +82,8 @@ Section fn. (* 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. do 16 ((eapply fp_E_proper; try reflexivity) || exact: Hfp || f_equiv). + do 12 f_equiv. f_contractive. + 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. @@ -151,7 +152,7 @@ Section typing. 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=>//= qL. iIntros "HL0". + 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) Ï) ∗ @@ -159,17 +160,17 @@ Section typing. 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 with "HL0") as "#Htys"; first done. + iDestruct (subtype_Forall2_llctx_noend with "HL0") as "#Htys"; first done. iDestruct (Hty with "HL0") as "#Hty". - iClear "∗". iIntros "!# #HEE". + iClear "∗". iIntros "!> #HEE". iSplit; last iSplit. - by iApply "HE". - by iApply "Htys". - by iApply "Hty". } - iClear "∗". clear Hcons. iIntros "!# #HE0 * Hf". + iClear "∗". clear Hcons. iIntros "!> #HE0 * Hf". iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst. - iExists fb, kb, xb, e, _. iSplit. done. iSplit. done. iNext. - rewrite /typed_body. iIntros (x Ï k xl) "!# * #LFT #HE' Htl HL HC HT". + 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". @@ -185,7 +186,7 @@ Section typing. -{2}(fst_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?vec_to_list_length // -{2}(snd_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?vec_to_list_length // !zip_with_fmap_r !(zip_with_zip (λ _ _, (_ ∘ _) _ _)) !big_sepL_fmap. - iApply (big_sepL_impl with "HT"). iIntros "!#". + 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. @@ -228,32 +229,31 @@ Section typing. 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=>//= qL. - iIntros "_ !# _ * Hf". iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst. - iExists fb, kb, xb, e, _. iSplit. done. iSplit. done. + 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 qL tid + 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 L qL -∗ + ⎡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 L qL -∗ - qκs.[lft_intersect_list κs] -∗ + (∀ 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] in tid {{ _, cont_postcondition }}) -∗ WP (call: p ps → k) in tid {{ _, cont_postcondition }}. Proof. iIntros (HE [k' <-]) "#LFT #HE Htl HL Hκs Hf Hargs Hk". wp_apply (wp_hasty with "Hf"); [done|]. iIntros (v) "% Hf". - iApply (wp_app_vec _ _ _ (_::_) ((λ v, ⌜v = (λ: ["_r"], (#☠;; #☠) ;; k' ["_r"])%VâŒ)::: + 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]"). - done. @@ -271,29 +271,37 @@ Section typing. 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 (Ï) "[Htk #Hinh]"; first done. + 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'". - iMod (bor_create _ Ï with "LFT Hκs") as "[Hκs HκsI]"; first done. - iDestruct (frac_bor_lft_incl with "LFT [>Hκs]") as "#Hκs". - { iApply (bor_fracture with "LFT Hκs"); [|done]. - split; [by rewrite Qp_mul_1_r|]. apply _. } - iApply ("Hf" with "LFT [] Htl [Htk] [Hk HκsI HL]"). - + iApply "HE'". iIntros "{$# Hf Hinh HE' LFT HE %}". - iInduction κs as [|κ κs] "IH"=> //=. iSplitL. - { iApply lft_incl_trans; first done. iApply lft_intersect_incl_l. } - iApply "IH". iModIntro. iApply lft_incl_trans; first done. - iApply lft_intersect_incl_r. - + iSplitL; last done. iExists Ï. iSplit; first by rewrite /= left_id. - iFrame "#∗". - + iIntros (y) "IN". iDestruct "IN" as %->%stdpp.list.elem_of_list_singleton. + 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 κ'. simpl. wp_rec. wp_bind Endlft. - iSpecialize ("Hinh" with "Htk"). iClear "Hκs". - iApply (wp_mask_mono _ (↑lftN ∪ ↑histN)); first done. - iApply (wp_step_fupd with "Hinh"); [done|set_solver+|]. wp_seq. - iIntros "#Htok !>". wp_seq. iMod ("HκsI" with "Htok") as ">Hκs". - iApply ("Hk" with "Htl HL Hκs"). rewrite tctx_hasty_val. done. + 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']). @@ -315,7 +323,7 @@ Section typing. 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). by rewrite /llctx_interp. + - instantiate (1 := 1%Qp). instantiate (1 := 1%Qp). by rewrite /llctx_interp_noend. - iIntros "* Htl _". iApply "Hk". done. Qed. @@ -329,7 +337,7 @@ Section typing. T) (call: p ps → k). Proof. - iIntros (Hκs HE tid) "#LFT #HE Htl HL HC (Hf & Hargs & HT)". + 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". @@ -338,13 +346,13 @@ Section typing. rewrite tctx_interp_cons tctx_hasty_val. iFrame. Qed. - (* Specialized type_call': Adapted for use by solve_typing; fixed "list of - alive lifetimes" to be the local ones. *) - Lemma type_call {A} x E L C T T' T'' p (ps : list path) + (* 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) (L.*1) → - (∀ Ï, elctx_sat (((λ κ, Ï âŠ‘â‚‘ κ) <$> (L.*1)) ++ E) L ((fp x).(fp_E) Ï)) → + 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 → @@ -353,8 +361,9 @@ Section typing. Proof. intros Hfn HL HE HTT' HC HT'T''. rewrite -typed_body_mono /flip; last done; first by eapply type_call'. - - etrans. eapply (incl_cctx_incl _ [_]); first by intros ? ->%stdpp.list.elem_of_list_singleton. - apply cctx_incl_cons_match; first done. intros args. by inv_vec args. + - 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. @@ -381,12 +390,14 @@ Section typing. 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)). + 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. + 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. } @@ -407,11 +418,11 @@ Section typing. (subst_v (fb :: BNamed "return" :: argsb) (f ::: k ::: args) e)) -∗ typed_instruction_ty E L T ef (fn fp). Proof. - iIntros (<- ->) "#Hbody /=". iIntros (tid) "#LFT _ $ $ #HT". iApply wp_value. + iIntros (<- ->) "#Hbody /=". iIntros (tid qmax) "#LFT _ $ $ #HT". iApply wp_value. rewrite tctx_interp_singleton. iLöb as "IH". iExists _. iSplit. { unlock. rewrite /= decide_True_pi. done. } - iExists fb, _, argsb, e, _. iSplit. done. iSplit. done. iNext. - iIntros (x Ï k args) "!#". iIntros (tid') "_ HE Htl HL HC HT'". + 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. unlock. iFrame "IH HT'". by iApply sendc_change_tid. @@ -422,17 +433,18 @@ Section typing. IntoVal ef (funrec: <> argsb := e) → n = length argsb → â–¡ (∀ x Ï k (args : vec val (length argsb)), - typed_body ((fp x).(fp_E) Ï) [Ï âŠ‘â‚— []] + 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 "!# *". + 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. -Hint Resolve fn_subtype : lrust_typing. +Global Hint Resolve fn_subtype : lrust_typing. diff --git a/theories/typing/int.v b/theories/typing/int.v index 3290bd6354ada60d990fee46e01f1c811afb0dc7..595dcd76141295cfeaa4c5ccfd2677de4475b82e 100644 --- a/theories/typing/int.v +++ b/theories/typing/int.v @@ -26,7 +26,7 @@ Section typing. Lemma type_int_instr (z : Z) : typed_val #z int. Proof. - iIntros (E L tid) "_ _ $ $ _". iApply wp_value. + iIntros (E L tid ?) "_ _ $ $ _". iApply wp_value. by rewrite tctx_interp_singleton tctx_hasty_val' //. Qed. @@ -39,7 +39,7 @@ Section typing. Lemma type_plus_instr E L p1 p2 : ⊢ typed_instruction_ty E L [p1 â— int; p2 â— int] (p1 + p2) int. Proof. - iIntros (tid) "_ _ $ $ [Hp1 [Hp2 _]]". + iIntros (tid ?) "_ _ $ $ [Hp1 [Hp2 _]]". wp_apply (wp_hasty with "Hp1"); [done|]. iIntros ([[]|]) "_ H1"; try done. wp_apply (wp_hasty with "Hp2"); [done|]. iIntros ([[]|]) "_ H2"; try done. wp_op. by rewrite tctx_interp_singleton tctx_hasty_val' //. @@ -55,7 +55,7 @@ Section typing. Lemma type_minus_instr E L p1 p2 : ⊢ typed_instruction_ty E L [p1 â— int; p2 â— int] (p1 - p2) int. Proof. - iIntros (tid) "_ _ $ $ [Hp1 [Hp2 _]]". + iIntros (tid ?) "_ _ $ $ [Hp1 [Hp2 _]]". wp_apply (wp_hasty with "Hp1"); [done|]. iIntros ([[]|]) "_ H1"; try done. wp_apply (wp_hasty with "Hp2"); [done|]. iIntros ([[]|]) "_ H2"; try done. wp_op. by rewrite tctx_interp_singleton tctx_hasty_val' //. @@ -71,7 +71,7 @@ Section typing. Lemma type_le_instr E L p1 p2 : ⊢ typed_instruction_ty E L [p1 â— int; p2 â— int] (p1 ≤ p2) bool. Proof. - iIntros (tid) "_ _ $ $ [Hp1 [Hp2 _]]". + iIntros (tid ?) "_ _ $ $ [Hp1 [Hp2 _]]". wp_apply (wp_hasty with "Hp1"); [done|]. iIntros ([[]|]) "_ H1"; try done. wp_apply (wp_hasty with "Hp2"); [done|]. iIntros ([[]|]) "_ H2"; try done. wp_op; case_bool_decide; by rewrite tctx_interp_singleton tctx_hasty_val' //. diff --git a/theories/typing/lft_contexts.v b/theories/typing/lft_contexts.v index 73c7c8dffee0c9c93955d5fe778723cf7d0348df..03f38776ec8f8d5e3fc775791cd296edf8f90153 100644 --- a/theories/typing/lft_contexts.v +++ b/theories/typing/lft_contexts.v @@ -23,12 +23,12 @@ Notation llctx := (list llctx_elt). Notation "κ ⊑ₗ κl" := (@pair lft (list lft) κ κl) (at level 70). Section lft_contexts. - Context `{!invG Σ, !lftG view_Lat (↑histN) Σ}. + Context `{!invG Σ, !lftG Σ view_Lat lft_userE}. Implicit Type (κ : lft). (* External lifetime contexts. *) Definition elctx_elt_interp (x : elctx_elt) : vProp Σ := - (x.1 ⊑ x.2)%I. + ⌜(x.1 ⊑ˢʸ⿠x.2)âŒ%I. Definition elctx_interp (E : elctx) : vProp Σ := ([∗ list] x ∈ E, elctx_elt_interp x)%I. @@ -40,38 +40,80 @@ Section lft_contexts. Proof. apply _. Qed. (* Local lifetime contexts. *) - Definition llctx_elt_interp (x : llctx_elt) (q : Qp) : vProp Σ := + (* 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) : vProp Σ := let κ' := lft_intersect_list (x.2) in - (∃ κ0, ⌜x.1 = κ' ⊓ κ0⌠∗ q.[κ0] ∗ â–¡ (1.[κ0] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ [†κ0]))%I. - Global Instance llctx_elt_interp_fractional x : - Fractional (llctx_elt_interp x). + (* 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) : vProp Σ := + 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 Hq'] & #?)". + - 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 *. subst. - rewrite (inj ((lft_intersect_list κs) ⊓.) κ0' κ0) //. - iExists κ0. by iFrame "∗#". + 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 (L : llctx) (q : Qp) : vProp Σ := - ([∗ list] x ∈ L, llctx_elt_interp x q)%I. - Global Arguments llctx_interp _ _%Qp. - Global Instance llctx_interp_fractional L : - Fractional (llctx_interp L). + Definition llctx_interp (qmax : Qp) (L : llctx) : vProp Σ := + ([∗ list] x ∈ L, llctx_elt_interp qmax x)%I. + Definition llctx_interp_noend (qmax : Qp) (L : llctx) (q : Qp) : vProp Σ := + ([∗ 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 L q : - AsFractional (llctx_interp L q) (llctx_interp L) q. - Proof. split. done. apply _. Qed. - Global Instance llctx_interp_permut : - Proper ((≡ₚ) ==> eq ==> (⊣⊢)) llctx_interp. - Proof. intros ????? ->. by apply big_opL_permutation. Qed. - - Lemma lctx_equalize_lft qL κ1 κ2 `{!frac_borG Σ} : - ⎡lft_ctx⎤ -∗ llctx_elt_interp (κ1 ⊑ₗ [κ2]%list) qL ={⊤}=∗ - elctx_elt_interp (κ1 ⊑ₑ κ2) ∗ elctx_elt_interp (κ2 ⊑ₑ κ1). + 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 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". @@ -82,30 +124,53 @@ Section lft_contexts. + 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). - (* Lifetime inclusion *) Definition lctx_lft_incl κ κ' : Prop := - ∀ qL, llctx_interp L qL -∗ â–¡ (elctx_interp E -∗ κ ⊑ κ'). + ∀ 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 κ κ' : 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_relf κ : lctx_lft_incl κ κ. - Proof. iIntros (qL) "_ !# _". iApply lft_incl_refl. 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_relf. - iIntros (??? H1 H2 ?) "HL". + 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". iApply lft_incl_trans. - by iApply "H1". by iApply "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 : @@ -121,17 +186,18 @@ Section lft_contexts. Qed. Lemma lctx_lft_incl_static κ : lctx_lft_incl κ static. - Proof. iIntros (qL) "_ !# _". iApply lft_incl_static. Qed. + 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 qL) "H". + 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". - iApply lft_incl_trans; first iApply lft_intersect_incl_l. - by iApply lft_intersect_list_elem_of_incl. + 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 : @@ -140,7 +206,7 @@ Section lft_contexts. Lemma lctx_lft_incl_external κ κ' : κ ⊑ₑ κ' ∈ E → lctx_lft_incl κ κ'. Proof. - iIntros (??) "_ !# #HE". + iIntros (???) "_ !> #HE". rewrite /elctx_interp /elctx_elt_interp big_sepL_elem_of //. done. Qed. @@ -150,45 +216,51 @@ Section lft_contexts. Lemma lctx_lft_incl_intersect κ κ' κ'' : lctx_lft_incl κ κ' → lctx_lft_incl κ κ'' → - lctx_lft_incl κ (κ' ⊓ κ''). + lctx_lft_incl (κ ⊓ κ) (κ' ⊓ κ''). Proof. - iIntros (Hκ' Hκ'' ?) "HL". + iIntros (Hκ' Hκ'' ??) "HL". iDestruct (Hκ' with "HL") as "#Hκ'". iDestruct (Hκ'' with "HL") as "#Hκ''". - iIntros "!# #HE". iApply lft_incl_glb. by iApply "Hκ'". by iApply "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". + iIntros (Hκ' ??) "HL". iDestruct (Hκ' with "HL") as "#Hκ'". - iIntros "!# #HE". iApply lft_incl_trans. - by iApply lft_intersect_incl_l. by iApply "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". + iIntros (Hκ' ??) "HL". iDestruct (Hκ' with "HL") as "#Hκ'". - iIntros "!# #HE". iApply lft_incl_trans. - by iApply lft_intersect_incl_r. by iApply "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 qL, ↑lftN ⊆ F → elctx_interp E -∗ llctx_interp L qL ={F}=∗ - ∃ q', q'.[κ] ∗ (q'.[κ] ={F}=∗ llctx_interp L qL). + ∀ 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 κ F q : + Lemma lctx_lft_alive_tok_noend κ F qmax q : ↑lftN ⊆ F → - lctx_lft_alive κ → elctx_interp E -∗ llctx_interp L q ={F}=∗ - ∃ q', q'.[κ] ∗ llctx_interp L q' ∗ - (q'.[κ] -∗ llctx_interp L q' ={F}=∗ llctx_interp L q). + 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. @@ -198,17 +270,30 @@ Section lft_contexts. iApply "Hclose". iFrame. Qed. - Lemma lctx_lft_alive_tok_list κs F q : + 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 L q ={F}=∗ - ∃ q', q'.[lft_intersect_list κs] ∗ llctx_interp L q' ∗ - (q'.[lft_intersect_list κs] -∗ llctx_interp L q' ={F}=∗ llctx_interp L q). + 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 κ with "HE HL") as (q') "(Hκ & HL & Hclose1)"; [solve_typing..|]. + 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]". @@ -218,46 +303,65 @@ Section lft_contexts. 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 qL ?) "_ $". iExists 1%Qp. iSplitL. by iApply lft_tok_static. auto. + 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 qL ?) "#HE HL". - iDestruct "HL" as "[HL1 HL2]". rewrite {2}/llctx_interp /llctx_elt_interp. - iDestruct (big_sepL_lookup_acc with "HL2") as "[Hκ Hclose]". done. - iDestruct "Hκ" as (κ0) "(EQ & Htok & #Hend)". simpl. iDestruct "EQ" as %->. + 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 L (qL / 2)))%I + (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. iClear "Hend". + { move:(qL/2)%Qp=>qL'. clear HL. iInduction Hκs as [|κ κs Hκ ?] "IH" forall (qL'). - iExists 1%Qp. iFrame. iSplitR; last by auto. iApply lft_tok_static. - iDestruct "HL1" as "[HL1 HL2]". - iMod (Hκ with "HE HL1") as (q') "[Htok' Hclose]". done. + iMod (Hκ 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). - destruct (Qp_lower_bound q' (qL/2)) as (q0 & q'2 & q''2 & -> & ->). - iExists q0. rewrite -(lft_tok_sep q0). iDestruct "Htok" as "[$ Htok]". + 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. iFrame. auto. + 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 qL ?) "#HE HL". - iAssert (κ ⊑ κ')%I with "[#]" as "#Hincl". iApply (Hinc with "HL HE"). - iMod (Hal with "HE HL") as (q') "[Htok Hclose]". done. - iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose']". done. + 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. @@ -271,18 +375,18 @@ Section lft_contexts. (* External lifetime context satisfiability *) Definition elctx_sat E' : Prop := - ∀ qL, llctx_interp L qL -∗ â–¡ (elctx_interp E -∗ elctx_interp E'). + ∀ 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. + 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' qL) "HL". + iIntros (Hκκ' HE' qmax qL) "HL". iDestruct (Hκκ' with "HL") as "#Hincl". iDestruct (HE' with "HL") as "#HE'". - iClear "∗". iIntros "!# #HE". iSplit. + iClear "∗". iIntros "!> #HE". iSplit. - by iApply "Hincl". - by iApply "HE'". Qed. @@ -290,37 +394,41 @@ Section lft_contexts. Lemma elctx_sat_app E1 E2 : elctx_sat E1 → elctx_sat E2 → elctx_sat (E1 ++ E2). Proof. - iIntros (HE1 HE2 ?) "HL". + iIntros (HE1 HE2 ??) "HL". iDestruct (HE1 with "HL") as "#HE1". iDestruct (HE2 with "HL") as "#HE2". - iClear "∗". iIntros "!# #HE". + iClear "∗". iIntros "!> #HE". iDestruct ("HE1" with "HE") as "#$". iApply ("HE2" with "HE"). Qed. Lemma elctx_sat_refl : elctx_sat E. - Proof. iIntros (?) "_ !# ?". done. Qed. + Proof. iIntros (??) "_ !> ?". done. Qed. End lft_contexts. -Arguments lctx_lft_incl {_ _ _} _ _ _ _. -Arguments lctx_lft_eq {_ _ _} _ _ _ _. +Arguments lctx_lft_incl {_ _} _ _ _ _. +Arguments lctx_lft_eq {_ _} _ _ _ _. Arguments lctx_lft_alive {_ _ _} _ _ _. -Arguments elctx_sat {_ _ _} _ _ _. +Arguments elctx_sat {_ _} _ _ _. Arguments lctx_lft_incl_incl {_ _ _ _ _} _ _. +Arguments lctx_lft_incl_incl_noend {_ _ _ _} _ _. Arguments lctx_lft_alive_tok {_ _ _ _ _} _ _ _. +Arguments lctx_lft_alive_tok_noend {_ _ _ _ _} _ _ _. -Lemma elctx_sat_submseteq `{!invG Σ, !lftG view_Lat (↑histN) Σ} E E' L : +Lemma elctx_sat_submseteq `{!invG Σ, !lftG Σ view_Lat lft_userE} E E' L : E' ⊆+ E → elctx_sat E L E'. -Proof. iIntros (HE' ?) "_ !# H". by iApply big_sepL_submseteq. Qed. +Proof. iIntros (HE' ??) "_ !> H". by iApply big_sepL_submseteq. Qed. -Hint Resolve - lctx_lft_incl_relf lctx_lft_incl_static lctx_lft_incl_local' +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. -Hint Resolve elctx_sat_submseteq | 100 : lrust_typing. +Global Hint Extern 10 (lctx_lft_eq _ _ _ _) => split : lrust_typing. + +Global Hint Resolve elctx_sat_submseteq | 100 : lrust_typing. -Hint Opaque elctx_sat lctx_lft_alive lctx_lft_incl : lrust_typing. +Global Hint Opaque elctx_sat lctx_lft_alive lctx_lft_incl : lrust_typing. diff --git a/theories/typing/lib/arc.v b/theories/typing/lib/arc.v index 40a7f3451617ee764b563167ffbdc2ec9501d5a2..92af7db87e46eae02cc2be22d01a1006d5300ddf 100644 --- a/theories/typing/lib/arc.v +++ b/theories/typing/lib/arc.v @@ -32,7 +32,7 @@ Section arc. (* because [weak_new] cannot prove ty_shr, even for a dead *) (* lifetime. *) (ty.(ty_shr) ν tid (l +â‚— 2) ∨ [†ν]) ∗ - â–¡ (1.[ν] ={↑lftN ∪ ↑histN ∪ ↑arc_endN}[↑histN]â–·=∗ + â–¡ (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 ν γi γs γw γsw l n : @@ -56,7 +56,7 @@ Section arc. 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. + - iIntros "!> Hν". iMod ("Hvs" with "Hν") as "H". iModIntro. iNext. iMod "H" as "($ & H & $)". iDestruct "H" as (vl) "[??]". iExists _. iFrame. by iApply "Hinclo". Qed. @@ -65,7 +65,7 @@ Section arc. arc_persist tid ν γi γs γw γsw l ty -∗ arc_persist tid' ν γi γs γw γsw l ty. Proof. iIntros "#($ & ? & Hvs)". rewrite sync_change_tid. iFrame "#". - iIntros "!# Hν". iMod ("Hvs" with "Hν") as "H". iModIntro. iNext. + 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. @@ -122,7 +122,7 @@ Section arc. with "[H]" as "H". { by iRight. } iMod (inv_alloc arc_endN with "[H]") as "#INV". { rewrite -embed_or. iNext. by iApply "H". } - iIntros "!> !# Hν". + iIntros "!> !> Hν". iMod (inv_acc with "INV") as "[[H†|[>$ Hvs]] Hclose]"; [set_solver|..]. { iDestruct "H†" as (?) ">H†". iDestruct (lft_tok_dead_subj with "Hν [H†]") as "[]". @@ -173,7 +173,7 @@ Section arc. ∃ q' ts tw, shared_arc_local _ _ _ _ _ ∗ &at{_,_} _)%I). iMod (inv_alloc shrN _ (idx_bor_own i V ∨ C V)%I with "[Hpbown]") as "#Hinv"; [by iLeft; iFrame|]. - iIntros "!> !# * % Htok". + iIntros "!> !> * % Htok". iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. iDestruct "INV" as "[>Hbtok|#Hshr]". - iAssert (⎡(&{κ} _) V⎤)%I with "[Hbtok]" as "Hb". @@ -245,7 +245,7 @@ Section arc. iDestruct "Hvl" as (???? ν t q) "(#Hpersist & Htk & Hν)". iRight. iExists _,_,_,_,_,_,_. iFrame "#∗". by iApply arc_persist_type_incl. - iIntros "* #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'. - iIntros "{$Hfrac} !# * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H". + iIntros "{$Hfrac} !> * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H". iModIntro. iNext. iMod "H" as "[$ H]". iDestruct "H" as (???? ν) "(Hlft & Hpersist & Hna)". iExists _, _, _, _,_. iFrame. by iApply arc_persist_type_incl. @@ -254,8 +254,8 @@ Section arc. 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". + 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. @@ -312,7 +312,7 @@ Section arc. ∃ q' ts tw, shared_arc_local _ _ _ _ _ ∗ &at{_,_} _)%I). iMod (inv_alloc shrN _ (idx_bor_own i V ∨ C V)%I with "[Hpbown]") as "#Hinv"; first by iLeft; iFrame. - iIntros "!> !# * % Htok". + iIntros "!> !> * % Htok". iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. iDestruct "INV" as "[>Hbtok|#Hshr]". - iAssert (⎡(&{κ} _) V⎤)%I with "[Hbtok]" as "Hb". @@ -367,7 +367,7 @@ Section arc. iDestruct "Hvl" as (???? ν ? ?) "(#Hpersist & Htk)". iExists _, _, _, _, _, _, _. iFrame "#∗". by iApply arc_persist_type_incl. - iIntros "* #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'. - iIntros "{$Hfrac} !# * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H". + 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. @@ -376,8 +376,8 @@ Section arc. 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". + 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. @@ -413,11 +413,11 @@ Section 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]". + 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") @@ -455,11 +455,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrc [Hrcbox _]]". rewrite !tctx_hasty_val. + 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_mapsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]". @@ -482,7 +482,7 @@ Section arc. { rewrite freeable_sz_full_S. iFrame. iExists _. iFrame. by rewrite vec_to_list_length. } iExists γi, γs, γw, γsw, ν, t, q. iFrame. iSplitL; first by auto. - iIntros "!>!>!# Hν". iDestruct (lft_tok_dead with "Hν H†") as "[]". } + iIntros "!>!>!> Hν". iDestruct (lft_tok_dead with "Hν H†") as "[]". } iApply type_jump; solve_typing. Qed. @@ -497,11 +497,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrcx [#Hrc' [Hx _]]]". + 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_mapsto_vec_singleton. @@ -541,11 +541,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -561,7 +561,7 @@ Section arc. wp_apply (strong_count_spec (P1 ν) (P2 l' ty.(ty_size)) _ γi γs γw _ _ ts vs (q.[α])%I with "[] [] [$SR $Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!# Hα". + { iIntros "!> Hα". iMod (at_bor_acc with "LFT Hrctokb Hα") as (Vb) "[Htok Hclose1]"; [solve_ndisj..|]. iExists Vb, q'. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro. solve_ndisj. @@ -588,11 +588,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -607,7 +607,7 @@ Section arc. iDestruct "lc" as (vs vw) "[SR WR]". iModIntro. wp_let. wp_apply (weak_count_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ _ _ _ tw vw ts (q.[α])%I with "[] [] WR [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!# Hα". + { iIntros "!> Hα". iMod (at_bor_acc with "LFT Hrctokb Hα") as (Vb) "[Htok Hclose1]"; [solve_ndisj..|]. iExists Vb, q'. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro. solve_ndisj. iFrame. @@ -636,11 +636,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -656,7 +656,7 @@ Section arc. wp_apply (clone_arc_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ _ _ _ ts vs (q.[α])%I with "[] [] [] SR [$Hα1 $Hα2]"); [by iDestruct "Hpersist" as "[$ _]"| |by iExists _,_|]. - { iIntros "!# Hα". + { iIntros "!> Hα". iMod (at_bor_acc with "LFT Hrctokb Hα") as (Vb) "[Htok Hclose1]"; [solve_ndisj..|]. iExists Vb, q'. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro. solve_ndisj. iFrame. @@ -685,11 +685,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -705,7 +705,7 @@ Section arc. wp_apply (clone_weak_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ _ _ _ tw vw (q.[α])%I with "[] [] [SR] WR [$Hα1 $Hα2]"); [by iDestruct "Hpersist" as "[$ _]"| |by iExists _,_|]. - { iIntros "!# Hα". + { iIntros "!> Hα". iMod (at_bor_acc with "LFT Hrctokb Hα") as (Vb) "[Htok Hclose1]"; [solve_ndisj..|]. iExists Vb, q'. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro. solve_ndisj. iFrame. @@ -733,11 +733,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -752,7 +752,7 @@ Section arc. iDestruct "lc" as (vs vw) "[SR WR]". iModIntro. wp_let. wp_apply (downgrade_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ _ _ _ tw vw ts (q.[α])%I with "[] [] WR [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!# Hα". + { iIntros "!> Hα". iMod (at_bor_acc with "LFT Hrctokb Hα") as (Vb) "[Htok Hclose1]"; [solve_ndisj..|]. iExists Vb, q'. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro. solve_ndisj. iFrame. @@ -783,11 +783,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -803,7 +803,7 @@ Section arc. wp_apply (upgrade_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ _ _ _ ts vs tw (q.[α])%I with "[] [] [] SR [$Hα1 $Hα2]"); [by iDestruct "Hpersist" as "[$ _]"| |by iExists _,_|]. - { iIntros "!# Hα". + { iIntros "!> Hα". iMod (at_bor_acc with "LFT Hrctokb Hα") as (Vb) "[Htok Hclose1]"; [solve_ndisj..|]. iExists Vb, q0. iFrame. iMod fupd_mask_subseteq as "Hclose2"; last iModIntro. solve_ndisj. iFrame. @@ -847,12 +847,12 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hr [Hrcx [Hrc' _]]]". + 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. } @@ -911,11 +911,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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 (γi γs γw γsw ν t q) "[#Hpersist Htok]". wp_bind (drop_weak _). iApply (drop_weak_spec with "[] [Htok]"); @@ -955,12 +955,12 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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. } @@ -1031,11 +1031,11 @@ Section arc. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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..|]. @@ -1123,10 +1123,10 @@ Section arc. 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. + 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) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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..|]. @@ -1185,7 +1185,7 @@ Section arc. rewrite heap_mapsto_vec_singleton. wp_write. wp_let. wp_bind (of_val clone). iApply (wp_wand with "[Hna]"). - { iApply (Hclone _ [] with "LFT HE Hna"); rewrite /llctx_interp /tctx_interp //. } + { 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]". diff --git a/theories/typing/lib/brandedvec.v b/theories/typing/lib/brandedvec.v new file mode 100644 index 0000000000000000000000000000000000000000..76c38b11370452a992dc9407dafb86be7f7e6855 --- /dev/null +++ b/theories/typing/lib/brandedvec.v @@ -0,0 +1,420 @@ +From iris.algebra Require Import auth numbers. +From iris.proofmode Require Import tactics. +From lrust.lang Require Import notation new_delete. +From lrust.lifetime Require Import meta at_borrow. +From lrust.typing Require Import typing. +From lrust.typing.lib Require Import option. +Set Default Proof Using "Type". + +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Σ]. +Instance subG_brandidxΣ {Σ} : subG brandidxΣ Σ → brandidxG Σ. +Proof. solve_inG. Qed. + +Definition brandidxN := lrustN .@ "brandix". +Definition brandvecN := lrustN .@ "brandvec". + +Section brandedvec. + Context `{!typeG Σ, !brandidxG Σ}. + Implicit Types (q : Qp) (α : lft) (γ : gname) (n m : nat). + Local Notation vProp := (vProp Σ). + + Definition brandvec_inv α n : vProp := + (⎡∃ γ, 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 : vProp := + (⎡∃ γ, 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". + { iExists [_]. iNext. iFrame. iExists _. eauto with iFrame. } + iIntros "Htok2 Hmt1". iDestruct "Hmt1" as (vl') "[>Hmt1 #Hown']". + iDestruct ("Htok" with "Htok2") as "$". + iAssert (â–· ⌜1 = length vl'âŒ)%I as ">%". + { iNext. iDestruct (ty_size_eq with "Hown'") as %->. done. } + destruct vl' as [|v' vl']; first done. + destruct vl' as [|]; last first. { simpl in *. lia. } + rewrite !heap_mapsto_vec_singleton. + iDestruct (heap_mapsto_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 %<-. + iDestruct (own_valid_2 with "Hn Hm") as %[?%max_nat_included ?]%auth_both_valid_discrete. + iPureIntro. simpl in *. lia. + Qed. + +End brandedvec. + +Section typing. + Context `{!typeG Σ, !brandidxG Σ}. + Implicit Types (q : Qp) (α : lft) (γ : gname) (n m : nat). + Local Notation iProp := (iProp Σ). + + Definition brandvec_new (call_once : val) (R_F : type) : val := + funrec: <> ["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. rewrite bool_decide_false; [|lia]. + wp_case. + wp_alloc n as "Hn" "Hdead". + wp_let. + rewrite !heap_mapsto_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_mapsto_vec_singleton. + iFrame "Htok". + iExists n, (Vector.cons (LitV 0) Vector.nil). + iSplitR; first done. + simpl. + rewrite freeable_sz_full. + iFrame. + iExists 0%nat. + iSplitL; last done. + iExists γ; iSplitR; by eauto. } + 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 := + funrec: <> ["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 (qβ) "(Hβ & HL & Hclose)"; [solve_typing..|]. + iMod (frac_bor_acc with "LFT Hmem Hβ") as (qβ') "[>Hn'↦ Hcloseβ]"; first done. + rewrite !heap_mapsto_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 with "LFT Hghost Hβ") as (?) "[Hidx Hcloseg]"; [solve_ndisj..|]. + rewrite /brandvec_inv -embed_later monPred_in_embed'. + iDestruct "Hidx" as (γ) "> (#Hidx & Hown)". + iAssert (|==> ⎡own γ (â— (MaxNat m)) ∗ own γ (â—¯ (MaxNat m))⎤)%I with "[Hown]" as "> [Hown Hlb]". + { rewrite -own_op. iMod (own_update with "Hown"); [|done]. apply auth_update_alloc. + by apply max_nat_local_update. } + iMod ("Hcloseg" with "[Hown]") as "Hβ". + { iExists _. iNext. by 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". iModIntro. + 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 := + funrec: <> ["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 (qβ) "((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_mapsto_vec_singleton. + wp_read. + wp_op. + wp_read. + wp_op. + wp_case. + iApply fupd_wp. + iMod (at_bor_acc with "LFT Hnidx Hβ3") as (?) "[Hnidx Hcloseβ3]"; [solve_ndisj..|]. + rewrite /brandvec_inv -embed_later monPred_in_embed'. + iDestruct "Hnidx" as ">Hnidx". + 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 := + funrec: <> ["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 (qβ) "(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_mapsto_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_mapsto_vec_singleton. + iFrame "∗". + iIntros "!>". + iExists (n + 1)%nat. + iSplitL; last by (iPureIntro; do 3 f_equal; lia). + iExists γ. by iFrame. + } + 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 _. by iSplit. } + iApply type_assign; [solve_typing..|]. + iApply type_jump; solve_typing. + Qed. +End typing. diff --git a/theories/typing/lib/cell.v b/theories/typing/lib/cell.v index 7848c9620a54b82999352049f9f3d8862da38aa1..4beedf279001141c6e728fe8334a8c3e9ea5663c 100644 --- a/theories/typing/lib/cell.v +++ b/theories/typing/lib/cell.v @@ -34,10 +34,10 @@ Section cell. 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". + 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"]. + 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". @@ -88,10 +88,10 @@ Section typing. 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 "/= !#". + 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. + iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done. Qed. (* The other direction: getting ownership out of a cell. *) @@ -100,10 +100,10 @@ Section typing. 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 "/= !#". + 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. + iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done. Qed. Definition cell_get_mut : val := @@ -112,10 +112,10 @@ Section typing. 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 "/= !#". + 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. + iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done. Qed. Definition cell_from_mut : val := @@ -124,10 +124,10 @@ Section typing. 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 "/= !#". + 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. + iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done. Qed. Definition cell_into_box : val := @@ -136,10 +136,10 @@ Section typing. 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 "/= !#". + 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. + iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done. Qed. Definition cell_from_box : val := @@ -148,10 +148,10 @@ Section typing. 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 "/= !#". + 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. + iIntros (???) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done. Qed. (** Reading from a cell *) @@ -164,7 +164,7 @@ Section typing. 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 "/= !#". + 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]. @@ -186,13 +186,13 @@ Section typing. 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 "/= !#". + 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; [solve_typing..|]; iIntros (r); simpl_subst. (* Drop to Iris level. *) - iIntros (tid) "#LFT #HE Htl HL HC". + 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. @@ -237,9 +237,9 @@ Section typing. 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 "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk HT". + 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. @@ -251,4 +251,4 @@ Section typing. Qed. End typing. -Hint Resolve cell_mono' cell_proper' : lrust_typing. +Global Hint Resolve cell_mono' cell_proper' : lrust_typing. diff --git a/theories/typing/lib/diverging_static.v b/theories/typing/lib/diverging_static.v index 62fac42d7df3fe3035d9756f58b8c28870287fe6..8d0768aa9a19d5013c81ff4b3ea0790385b9aa7e 100644 --- a/theories/typing/lib/diverging_static.v +++ b/theories/typing/lib/diverging_static.v @@ -24,15 +24,15 @@ Section diverging_static. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk (Hcall & Hx & Hf & _)". + 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 (qÏ) "(HÏ & HL & _)"; [solve_typing..|]. - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & _ & _)"; + 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 _]]". @@ -51,6 +51,48 @@ Section diverging_static. { 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. + + 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 "!> *". inv_vec args=>r. simpl_subst. + iApply (type_cont [] [] (λ r, [])). + { iIntros (kloop). simpl_subst. iApply type_jump; solve_typing. } + iIntros "!> *". inv_vec args. simpl_subst. + iApply type_jump; solve_typing. Qed. End diverging_static. diff --git a/theories/typing/lib/fake_shared.v b/theories/typing/lib/fake_shared.v index 283ea8baaf801ef5b7a9bda0f7f81c1e2892ef27..1cd113064358ef3f7add80c9be1ae4be592e43c5 100644 --- a/theories/typing/lib/fake_shared.v +++ b/theories/typing/lib/fake_shared.v @@ -12,20 +12,21 @@ Section fake_shared. typed_val fake_shared_box (fn(∀ '(α, β), ∅; &shr{α}(&shr{β} ty)) → &shr{α}(box ty)). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk HT". + 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 "#Hαβ"; [solve_typing..|]. + 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 "!>!# *". + { iApply frac_bor_iff; last done. iIntros "!>!> *". rewrite heap_mapsto_vec_singleton. iSplit; auto. } - iDestruct "H" as "#H". iIntros "!# * % $". iApply step_fupd_intro. set_solver. - simpl. by iApply ty_shr_mono. } + iDestruct "H" as "#H". iIntros "!> * % $". iApply step_fupd_intro. 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. @@ -41,21 +42,23 @@ Section fake_shared. 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) "#LFT #HE Hna HL Hk HT". + 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 "#Hαβ"; [solve_typing..|]. + 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 "!>!# *". + { iApply frac_bor_iff; last done. iIntros "!>!> *". rewrite heap_mapsto_vec_singleton. iSplit; auto. } - iDestruct "H" as "#H". iIntros "!# * % $". iApply step_fupd_intro. set_solver. - simpl. iApply ty_shr_mono; last done. iApply lft_intersect_incl_l. } + iDestruct "H" as "#H". iIntros "!> * % $". iApply step_fupd_intro. 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. diff --git a/theories/typing/lib/join.v b/theories/typing/lib/join.v index c6a3fd721cc6b77bf4aa5b796496fc11d8316a8c..129de0b61a2b66e75f68a59b78d1959972d60f25 100644 --- a/theories/typing/lib/join.v +++ b/theories/typing/lib/join.v @@ -38,12 +38,12 @@ Section join. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk (HfB & HfA & HenvA & HenvB & _)". + 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. *) @@ -61,7 +61,7 @@ Section join. wp_rec. iApply (finish_spec with "[$Hfin Hret HÏ1]"); [solve_ndisj..| |auto]. rewrite right_id. iFrame. by iApply @send_change_tid. } iNext. iIntros (γc γe c) "Hjoin". wp_let. wp_let. - iMod (lctx_lft_alive_tok Ï with "HE HL") as (qÏ2) "(HÏ2 & HL & Hclose2)"; + 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]"). @@ -87,5 +87,4 @@ Section join. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. Qed. - End join. diff --git a/theories/typing/lib/mutex/mutex.v b/theories/typing/lib/mutex/mutex.v index b82ba926e2a6bb4382e4ff7364eb666977471840..adb30abf60e5c2ce32d57d19ef77ef2553313d8d 100644 --- a/theories/typing/lib/mutex/mutex.v +++ b/theories/typing/lib/mutex/mutex.v @@ -83,14 +83,14 @@ Section mutex. 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. + 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 "!# * Hvl". destruct vl as [|[[| |n]|]vl]; try done. + - iIntros "!> * Hvl". destruct vl as [|[[| |n]|]vl]; try done. simpl. iDestruct "Hvl" as "[$ [$ Hvl]]". by iApply "Howni". - - iIntros "!# * Hshr". iDestruct "Hshr" as (κ') "[Hincl Hshr]". + - iIntros "!> * Hshr". iDestruct "Hshr" as (κ') "[Hincl Hshr]". iExists _. iFrame "Hincl". iDestruct "Hshr" as (γ) "Hproto". iExists γ. iApply lock_proto_iff_proper; [|iFrame]. @@ -133,7 +133,7 @@ Section code. 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 "/= !#". + 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. *) @@ -141,7 +141,7 @@ Section code. (* FIXME: The following should work. We could then go into Iris later. iApply (type_memcpy ty); [solve_typing..|]. *) (* Switch to Iris. *) - iIntros (tid) "#LFT #HE Hna HL Hk [Hm [Hx _]]". + 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_mapsto_vec_cons with "Hm") as "[Hm0 Hm]". @@ -173,11 +173,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hm _]]". + 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. @@ -212,11 +212,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hm [Hm' _]]". + 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 (qα) "(Hα & HL & Hclose1)"; diff --git a/theories/typing/lib/mutex/mutexguard.v b/theories/typing/lib/mutex/mutexguard.v index de85d02cebdb992db14f69cfebcf3b4a4fe5463e..1c455e50f7648f967a6cea3c4e0e224e287c3ac9 100644 --- a/theories/typing/lib/mutex/mutexguard.v +++ b/theories/typing/lib/mutex/mutexguard.v @@ -63,7 +63,7 @@ Section mguard. iIntros (??????) "#? H". iDestruct "H" as (l') "[#Hf #H]". iExists _. iSplit. - by iApply frac_bor_shorten. - - iIntros "!# * % Htok". + - iIntros "!> * % Htok". iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. { iApply lft_intersect_mono. iApply lft_incl_refl. done. } iMod ("H" with "[] Htok") as "Hshr". done. iModIntro. iNext. @@ -86,16 +86,17 @@ Section mguard. Global Instance mutexguard_mono E L : Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) mutexguard. Proof. - intros α1 α2 Hα ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold. iIntros (Hty' q) "HL". + intros α1 α2 Hα ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold. + iIntros (Hty' qmax qL) "HL". iDestruct (Hty' with "HL") as "#Hty". clear Hty'. iDestruct (Hα with "HL") as "#Hα". - iIntros "!# #HE". iDestruct ("Hα" with "HE") as "{Hα} 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. - + by iApply lft_incl_trans. + + iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done. + iDestruct "Hinv" as (γ) "Hproto". - iExists γ. iApply (lock_proto_lft_mono with "Hα"). + iExists γ. iApply lock_proto_lft_mono; first by iApply lft_incl_syn_sem. iApply lock_proto_iff_proper; [|iFrame]. iApply bor_iff_proper. iNext. iApply heap_mapsto_pred_iff_proper. iModIntro; iIntros; iSplit; iIntros; by iApply "Ho". @@ -103,13 +104,15 @@ Section mguard. iApply heap_mapsto_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". + iDestruct "H" as "[$ #H]". iIntros "!> * % Htok". iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. done. iApply lft_incl_refl. } + { iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. } iMod ("H" with "[] Htok") as "Hshr". done. iModIntro. iNext. iMod "Hshr" as "[Hshr Htok]". iMod ("Hclose" with "Htok") as "$". - iApply ty_shr_mono; try done. iApply lft_intersect_mono. done. iApply lft_incl_refl. - by iApply "Hs". + 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 : @@ -124,8 +127,18 @@ Section mguard. Qed. (* POSIX requires the unlock to occur from the thread that acquired - the lock, so Rust does not implement Send for RwLockWriteGuard. We could - prove this. *) + 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". unfold lock_proto. + 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. @@ -142,12 +155,12 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hg [Hx [Hm _]]]". + 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†)". @@ -180,11 +193,11 @@ Section code. typed_val mutexguard_derefmut (fn(∀ '(α, β), ∅; &uniq{α}(mutexguard β ty)) → &uniq{α}ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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) "#LFT #HE Hna HL Hk [Hg [Hg' _]]". + 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. @@ -201,7 +214,7 @@ Section code. 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 α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + 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. *) @@ -210,7 +223,7 @@ Section code. { 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; done. } + 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. @@ -222,34 +235,34 @@ Section code. typed_val mutexguard_derefmut (fn(∀ '(α, β), ∅; &shr{α}(mutexguard β ty)) → &shr{α}ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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) "#LFT #HE Hna HL Hk [Hg [Hg' _]]". + 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 (qα) "([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 β with "HE HL") as (qβ) "(Hβ & HL & Hclose3)"; + iMod (lctx_lft_alive_tok_noend β with "HE HL") as (qβ) "(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β]"); - [done| |by iApply ("Hshr" with "[] Hα2β")|]; first done. + [| |by iApply ("Hshr" with "[] Hα2β")|]; [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 "#Hαβ"; [solve_typing..|]. + 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. 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. @@ -265,11 +278,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hg [Hm _]]". + 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. *) diff --git a/theories/typing/lib/option.v b/theories/typing/lib/option.v index b7d0b2fd5ed038782bf65f517db52866c20dafd3..15028885d135ac2f641d26659c6cbd4dea9f05c9 100644 --- a/theories/typing/lib/option.v +++ b/theories/typing/lib/option.v @@ -7,6 +7,12 @@ Section option. 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. @@ -26,8 +32,8 @@ Section option. 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. + 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..| |]. @@ -38,7 +44,7 @@ Section option. 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. + - iIntros "/= !>". iIntros (k args). inv_vec args. simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. Qed. @@ -56,7 +62,7 @@ Section option. Lemma option_unwrap_or_type Ï„ `{!TyWf Ï„} : typed_val (option_unwrap_or Ï„) (fn(∅; option Ï„, Ï„) → Ï„). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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..|]. @@ -81,7 +87,7 @@ Section option. Lemma option_unwrap_type Ï„ `{!TyWf Ï„} : typed_val (option_unwrap Ï„) (fn(∅; option Ï„) → Ï„). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + 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|]. diff --git a/theories/typing/lib/panic.v b/theories/typing/lib/panic.v index f6a6b094f87a94a6605e424d6cdb27f7ab83d584..58a6b02eed73ca1f685aa781f16f811f405b3224 100644 --- a/theories/typing/lib/panic.v +++ b/theories/typing/lib/panic.v @@ -18,8 +18,8 @@ Section panic. Lemma panic_type : typed_val panic (fn(∅) → ∅). Proof. - intros E L. iApply type_fn; [done|]. iIntros "!# *". - inv_vec args. iIntros (tid) "LFT HE Hna HL Hk HT /=". simpl_subst. + intros E L. iApply type_fn; [done|]. iIntros "!> *". + inv_vec args. iIntros (tid qmax) "LFT HE Hna HL Hk HT /=". simpl_subst. by iApply wp_value. Qed. End panic. diff --git a/theories/typing/lib/rc/rc.v b/theories/typing/lib/rc/rc.v index 6e46c823737a2f649a31df83dd4bf22a2e0f2aea..71b859886f56a481cb9c83a1da1f69e8f450a434 100644 --- a/theories/typing/lib/rc/rc.v +++ b/theories/typing/lib/rc/rc.v @@ -1,5 +1,4 @@ -From Coq.QArith Require Import Qcanon. -From iris.algebra Require Import auth excl csum frac agree numbers. +From iris.algebra Require Import auth csum frac agree excl numbers. From lrust.lang Require Import memcpy. From lrust.lifetime Require Import na_borrow. From lrust.typing Require Export type. @@ -86,7 +85,7 @@ Section rc. because [weak_new] cannot prove ty_shr, even for a dead lifetime. *) (ty.(ty_shr) ν tid (l +â‚— 2) ∨ [†ν]) ∗ - â–¡ (1.[ν] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ [†ν]))%I. + â–¡ (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. @@ -143,7 +142,7 @@ Section rc. set (C := (∃ γ ν q', _ ⊑ _ ∗ rc_persist _ _ _ _ _ ∗ &na{_,_,_} _)%I). iMod (inv_alloc shrN _ (idx_bor_own i V ∨ C V)%I with "[Hpbown]") as "#Hinv"; [by iLeft; iFrame|]. - iIntros "!> !#" (F q' ?) "Htok". + iIntros "!> !>" (F q' ?) "Htok". iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. iDestruct "INV" as "[>Hbtok|#Hshr]". - iAssert (⎡(&{κ} _) V⎤)%I with "[Hbtok]" as "Hb". @@ -226,7 +225,7 @@ Section rc. iDestruct "Hvl" as (γ ν q) "(#Hpersist & Htk & Hν)". iRight. iExists _, _, _. iFrame "#∗". by iApply rc_persist_type_incl. - iIntros "* #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'. - iIntros "{$Hfrac} !# * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H". + 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. @@ -235,8 +234,8 @@ Section rc. 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". + 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. @@ -254,16 +253,16 @@ Section code. (((⌜strong = 1%positive⌠∗ (∃ weak : Z, (l +â‚— 1) ↦ #weak ∗ ((⌜weak = 1⌠∗ - |={⊤}[↑histN]â–·=> ⎡†l…(2 + ty.(ty_size))⎤ ∗ + |={⊤}[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) - ={⊤}[↑histN]â–·=∗ â–· (l +â‚— 2) ↦∗: ty.(ty_own) tid ∗ + ={⊤}[lft_userE]â–·=∗ â–· (l +â‚— 2) ↦∗: ty.(ty_own) tid ∗ ((l +â‚— 2) ↦∗: (λ vl, ⌜length vl = ty.(ty_size)âŒ) ={⊤}=∗ na_own tid F)))))) ∧ - (l ↦ #0 ={⊤}[↑histN]â–·=∗ + (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 ∗ @@ -302,12 +301,13 @@ Section code. iMod (own_update_2 with "Hst Htok") as "Hst". { apply auth_update_dealloc. rewrite -{1}(right_id (None, 0%nat) _ (_, _)). apply cancel_local_update_unit, _. } - rewrite -[in (1).[ν]%I]Hqq' -[(|={↑histN,⊤}=>_)%I]fupd_trans. + 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†". - iApply fupd_mask_mono; [|iMod ("Hν†" with "H†") as "Hty"]; [set_solver|]. - iModIntro. + 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". @@ -316,12 +316,14 @@ Section code. 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' -[(|={↑histN,⊤}=>_)%I]fupd_trans. + 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†". - iApply fupd_mask_mono; [|iMod ("Hν†" with "H†") as "Hty"]; [set_solver|]. - iModIntro. iMod (own_update_2 with "Hst Htok") as "Hst". + 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". @@ -335,12 +337,14 @@ Section code. { 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 -[(|={↑histN,⊤}=>_)%I]fupd_trans -[in (1).[ν]%I]Hqq'. + 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†". - iApply fupd_mask_mono; [|iMod ("Hν†" with "H†") as "Hty"]; [set_solver|]. - iModIntro. iMod ("Hclose" with "[Hst $Hna Hl1 Hl2]") as "$"; + 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". } @@ -397,11 +401,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -456,11 +460,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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. *) @@ -516,11 +520,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]". + 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") @@ -559,11 +563,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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†)". @@ -621,11 +625,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hrcx [#Hrc' [Hx _]]]". + 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_mapsto_vec_singleton. @@ -684,7 +688,7 @@ Section code. 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 "/= !#". + 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 [] [Ï âŠ‘â‚— []] @@ -695,7 +699,7 @@ Section code. iApply type_jump; solve_typing. } iIntros (k). simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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' ]). @@ -779,7 +783,7 @@ Section code. 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 "/= !#". + 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 [] [Ï âŠ‘â‚— []] @@ -790,7 +794,7 @@ Section code. iApply type_jump; solve_typing. } iIntros (k). simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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' ]). @@ -869,7 +873,7 @@ Section code. 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 "/= !#". + 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 [] [Ï âŠ‘â‚— []] @@ -880,7 +884,7 @@ Section code. iApply type_jump; solve_typing. } iIntros (k). simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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..|]. @@ -1002,7 +1006,7 @@ Section code. 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. + 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)])); @@ -1012,7 +1016,7 @@ Section code. iApply type_jump; solve_typing. } iIntros (k). simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". + 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)"; @@ -1088,7 +1092,7 @@ Section code. 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 _ [] with "LFT HE 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. diff --git a/theories/typing/lib/rc/weak.v b/theories/typing/lib/rc/weak.v index 0195aa92e9c791c607ec8b747a30ff2bc1f17476..ed5b7d8837ebda83cd580993bfe02a7799139fc3 100644 --- a/theories/typing/lib/rc/weak.v +++ b/theories/typing/lib/rc/weak.v @@ -39,7 +39,7 @@ Section weak. set (C := (∃ _ _, rc_persist _ _ _ _ _ ∗ &na{_,_,_} _)%I). iMod (inv_alloc shrN _ (idx_bor_own i V ∨ C V)%I with "[Hpbown]") as "#Hinv"; [by iLeft; iFrame|]. - iIntros "!> !#" (F q' ?) "Htok". + iIntros "!> !>" (F q' ?) "Htok". iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. iDestruct "INV" as "[>Hbtok|#Hshr]". - iAssert (⎡(&{κ} _) V⎤)%I with "[Hbtok]" as "Hb". @@ -92,7 +92,7 @@ Section weak. iDestruct "Hvl" as (γ ν) "(#Hpersist & Htok)". iExists _, _. iFrame "Htok". by iApply rc_persist_type_incl. - iIntros "* #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'. - iIntros "{$Hfrac} !# * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H". + 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. @@ -100,8 +100,8 @@ Section weak. 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. + 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. @@ -131,7 +131,7 @@ Section code. 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 "/= !#". + 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 [] [Ï âŠ‘â‚— []] @@ -142,7 +142,7 @@ Section code. iApply type_jump; solve_typing. } iIntros (k). simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (w'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hw [Hw' [Hr _]]]". + 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†)". @@ -240,11 +240,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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†)". @@ -303,11 +303,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". + 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†)". @@ -379,7 +379,7 @@ Section code. 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 "/= !#". + 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. @@ -389,7 +389,7 @@ Section code. iApply type_jump; solve_typing. } iIntros (k). simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (w'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hw [Hw' _]]". + 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 ∗ @@ -454,11 +454,11 @@ Section code. 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 "/= !#". + 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) "#LFT #HE Hna HL Hk [Hw [Hrcbox _]]". rewrite !tctx_hasty_val. + 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_mapsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]". @@ -469,7 +469,7 @@ Section code. 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+..|]. + 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. diff --git a/theories/typing/lib/refcell/ref.v b/theories/typing/lib/refcell/ref.v index 4028a108c4cf002342581e5e00d8078af7396762..8c3e2f0cab611084fa127d787269fc6d0632d358 100644 --- a/theories/typing/lib/refcell/ref.v +++ b/theories/typing/lib/refcell/ref.v @@ -79,22 +79,22 @@ Section ref. Global Instance ref_mono E L : Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) ref. Proof. - iIntros (α1 α2 Hα ty1 ty2 Hty q) "HL". + iIntros (α1 α2 Hα ty1 ty2 Hty qmax qL) "HL". iDestruct (Hty with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". - iIntros "!# #HE". iDestruct ("Hα" with "HE") as "Hα1α2". + iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2. iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro]. - done. - iIntros (tid [|[[]|][|[[]|][]]]) "H"=>//=. iDestruct "H" as (ν q' γ β ty') "(#Hshr & #H⊑ & #Hinv & Htok & Hown)". iExists ν, q', γ, β, ty'. iFrame "∗#". iSplit. + iApply ty_shr_mono; last by iApply "Hs". - iApply lft_intersect_mono. done. iApply lft_incl_refl. - + by iApply lft_incl_trans. + iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. + + iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done. - iIntros (κ tid l) "H /=". iDestruct "H" as (ν q' γ β ty' lv lrc) "H". iExists ν, q', γ, β, ty', lv, lrc. iDestruct "H" as "#($&$&?&?&$&$)". iSplit. + iApply ty_shr_mono; last by iApply "Hs". - iApply lft_intersect_mono. done. iApply lft_incl_refl. - + by iApply lft_incl_trans. + iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. + + iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done. Qed. Global Instance ref_mono_flip E L : Proper (lctx_lft_incl E L ==> flip (subtype E L) ==> flip (subtype E L)) ref. @@ -112,4 +112,4 @@ Section ref. Proof. intros. by eapply ref_proper. Qed. End ref. -Hint Resolve refcell_mono' refcell_proper' : lrust_typing. +Global Hint Resolve refcell_mono' refcell_proper' : lrust_typing. diff --git a/theories/typing/lib/refcell/ref_code.v b/theories/typing/lib/refcell/ref_code.v index 583d03407a613cab176170549649421b45e8dcea..82ced030145b076eb7684c0aa9bcc645835e92c2 100644 --- a/theories/typing/lib/refcell/ref_code.v +++ b/theories/typing/lib/refcell/ref_code.v @@ -13,7 +13,7 @@ Section ref_functions. ⎡own γ (â—¯ reading_stR q ν)⎤ -∗ ∃ (q' : Qp) n, l ↦ #(Zpos n) ∗ ⌜(q ≤ q')%Qp⌠∗ ⎡own γ (â— (refcell_st_to_R $ Some (ν, false, q', n)) â‹… â—¯ reading_stR q ν)⎤ ∗ - ((1).[ν] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ &{α}((l +â‚— 1) ↦∗: ty_own ty tid)) ∗ + ((1).[ν] ={↑lftN ∪ lft_userE}[lft_userE]â–·=∗ &{α}((l +â‚— 1) ↦∗: ty_own ty tid)) ∗ (∃ q'', ⌜(q' + q'' = 1)%Qp⌠∗ q''.[ν]) ∗ ty.(ty_shr) (α ⊓ ν) tid (l +â‚— 1). Proof. @@ -47,17 +47,17 @@ Section ref_functions. typed_val ref_clone (fn(∀ '(α, β), ∅; &shr{α}(ref β ty)) → ref β ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=. iDestruct "Hx'" as (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hβδ & Hinv & Hâ—¯inv)". wp_op. iMod (lctx_lft_alive_tok β with "HE HL") as (qβ) "(Hβ & HL & Hclose)"; [solve_typing..|]. iMod (lft_incl_acc with "Hβδ Hβ") as (qδ) "[Hδ Hcloseβ]". done. - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "([Hα1 Hα2] & HL & Hclose')"; + iMod (lctx_lft_alive_tok_noend α with "HE HL") as (qα) "([Hα1 Hα2] & HL & Hclose')"; [solve_typing..|]. iMod (frac_bor_acc with "LFT Hfrac Hα1") as (qlx') "[H↦ Hcloseα1]". done. iMod (na_bor_acc with "LFT Hinv Hδ Hna") as "(INV & Hna & Hcloseδ)"; [done..|]. @@ -107,10 +107,10 @@ Section ref_functions. typed_val ref_deref (fn(∀ '(α, β), ∅; &shr{α}(ref β ty)) → &shr{α}ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=. iDestruct "Hx'" as (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hx')". @@ -119,12 +119,13 @@ Section ref_functions. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. iMod "H↦" as "[H↦1 H↦2]". wp_read. wp_let. iMod ("Hcloseα" with "[$H↦1 $H↦2]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|]. iApply (type_type _ _ _ [ x â— box (&shr{α}(ref β ty)); #lv â— &shr{α}ty] with "[] LFT HE Hna HL Hk"); first last. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - iFrame. iApply (ty_shr_mono with "[] Hshr"). by iApply lft_incl_glb. } + iFrame. iApply (ty_shr_mono with "[] Hshr"). + iApply lft_incl_glb; last done. by iApply lft_incl_syn_sem. } iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. @@ -143,9 +144,9 @@ Section ref_functions. Lemma ref_drop_type ty `{!TyWf ty} : typed_val ref_drop (fn(∀ α, ∅; ref α ty) → unit). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val. + iIntros (tid qmax) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val. destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". iDestruct "Hx" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hx"; simpl; try iDestruct "Hx" as "[_ >[]]". @@ -159,7 +160,7 @@ Section ref_functions. as (q' n) "(H↦lrc & >% & Hâ—â—¯ & H†& Hq' & Hshr)". iDestruct "Hq'" as (q'') ">[% Hν']". wp_read. wp_let. wp_op. wp_write. - iAssert (|={↑lftN ∪ ↑histN}[↑histN]â–·=> refcell_inv tid lrc γ β ty')%I + iAssert (|={↑lftN ∪ lft_userE}[lft_userE]â–·=> refcell_inv tid lrc γ β ty')%I with "[H↦lrc Hâ—â—¯ Hν Hν' Hshr H†]" as "INV". { iDestruct (own_valid with "Hâ—â—¯") as %[[[[_ ?]?]|[[_ [q0 Hq0]]%prod_included [n' Hn']]%prod_included] @@ -177,10 +178,10 @@ Section ref_functions. { apply auth_update_dealloc. rewrite -(agree_idemp (to_agree _)) !pair_op Some_op. apply (cancel_local_update_unit (reading_stR q ν)), _. } - iApply step_fupd_intro; first set_solver. iExists (q+q'')%Qp. iFrame. + iApply step_fupd_intro; first set_solver-. iExists (q+q'')%Qp. iFrame. by rewrite assoc (comm _ q0 q). } - wp_bind Endlft. iApply (wp_mask_mono _ (↑lftN ∪ ↑histN)); first done. - iApply (wp_step_fupd with "INV"); [set_solver..|]. wp_seq. + wp_bind Endlft. iApply (wp_mask_mono _ (↑lftN ∪ lft_userE)); first done. + iApply (wp_step_fupd with "INV"); [set_solver-..|]. wp_seq. iIntros "INV !>". wp_seq. iMod ("Hcloseβ" with "[$INV] Hna") as "[Hβ Hna]". iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". iApply (type_type _ _ _ [ #lx â— box (uninit 2)] with "[] LFT HE Hna HL Hk"); @@ -208,10 +209,11 @@ Section ref_functions. typed_val call_once (fn(∀ α, ∅; fty, &shr{α}ty1) → &shr{α}ty2) → typed_val (ref_map call_once) (fn(∀ α, ∅; ref α ty1, fty) → ref α ty2). Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret arg). + intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (α Ï ret arg). inv_vec arg=>ref env. simpl_subst. iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". + iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done. iDestruct "Href" as "[Href Href†]". iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl; @@ -223,10 +225,10 @@ Section ref_functions. iIntros (lx) "(H†& Hlx)". rewrite heap_mapsto_vec_singleton. wp_let. wp_write. wp_let. rewrite tctx_hasty_val. iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|]. - iMod (lctx_lft_alive_tok Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. + iMod (lctx_lft_alive_tok_noend Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". iDestruct (lft_intersect_acc with "Hαν HÏ") as (?) "[HÎ±Î½Ï Hclose4]". - rewrite -[Ï in (α ⊓ ν) ⊓ Ï](right_id_L _ (⊓)). + rewrite -[Ï in (α ⊓ ν) ⊓ Ï]right_id_L. iApply (type_call_iris _ [α ⊓ ν; Ï] (α ⊓ ν) [_; _] with "LFT HE Hna [HανÏ] Hf' [$Henv Hlx H†]"); [solve_typing|done| |]. { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. @@ -286,10 +288,11 @@ Section ref_functions. typed_val (ref_map_split call_once) (fn(∀ α, ∅; ref α ty, fty) → Î [ref α ty1; ref α ty2]). Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret arg). + intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (α Ï ret arg). inv_vec arg=>ref env. simpl_subst. iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". + iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done. iDestruct "Href" as "[Href Href†]". iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl; @@ -301,10 +304,10 @@ Section ref_functions. iIntros (lx) "(H†& Hlx)". rewrite heap_mapsto_vec_singleton. wp_let. wp_write. wp_let. rewrite tctx_hasty_val. iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|]. - iMod (lctx_lft_alive_tok Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. + iMod (lctx_lft_alive_tok_noend Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". iDestruct (lft_intersect_acc with "Hαν HÏ") as (?) "[HÎ±Î½Ï Hclose4]". - rewrite -[Ï in (α ⊓ ν) ⊓ Ï](right_id_L static (⊓)). + rewrite -[Ï in (α ⊓ ν) ⊓ Ï]right_id_L. iApply (type_call_iris _ [α ⊓ ν; Ï] (α ⊓ ν) [_; _] with "LFT HE Hna [HανÏ] Hf' [$Henv Hlx H†]"); [solve_typing|done| |]. { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. diff --git a/theories/typing/lib/refcell/refcell.v b/theories/typing/lib/refcell/refcell.v index 76b1c4087766bad9bfc4a595ef6c1a02919d37c9..1ddebc8d8933636ea00d737b1bd0807a6b9d5f0d 100644 --- a/theories/typing/lib/refcell/refcell.v +++ b/theories/typing/lib/refcell/refcell.v @@ -48,7 +48,7 @@ Section refcell_inv. (* Not borrowed. *) &{α}((l +â‚— 1) ↦∗: ty.(ty_own) tid) | Some (ν, rw, q, _) => - (1.[ν] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ &{α}((l +â‚— 1) ↦∗: ty.(ty_own) tid)) ∗ + (1.[ν] ={↑lftN ∪ lft_userE}[lft_userE]â–·=∗ &{α}((l +â‚— 1) ↦∗: ty.(ty_own) tid)) ∗ (∃ q', ⌜(q + q')%Qp = 1%Qp⌠∗ q'.[ν]) ∗ if rw then (* Mutably borrowed. *) True else (* Immutably borrowed. *) ty.(ty_shr) (α ⊓ ν) tid (l +â‚— 1) @@ -67,19 +67,19 @@ Section refcell_inv. eapply refcell_inv_type_ne, type_dist_dist2. done. Qed. - Lemma refcell_inv_proper E L ty1 ty2 q : + Lemma refcell_inv_proper E L ty1 ty2 qmax qL : eqtype E L ty1 ty2 → - llctx_interp L q -∗ ∀ tid l γ α, â–¡ (elctx_interp E -∗ + llctx_interp_noend qmax L qL -∗ ∀ tid l γ α, â–¡ (elctx_interp E -∗ refcell_inv tid l γ α ty1 -∗ refcell_inv tid l γ α ty2). Proof. (* TODO : this proof is essentially [solve_proper], but within the logic. *) (* It would easily gain from some automation. *) rewrite eqtype_unfold. iIntros (Hty) "HL". - iDestruct (Hty with "HL") as "#Hty". iIntros "* !# #HE H". + iDestruct (Hty with "HL") as "#Hty". iIntros "* !> #HE H". iDestruct ("Hty" with "HE") as "(% & #Hown & #Hshr)". iAssert (â–¡ (&{α}((l +â‚— 1) ↦∗: ty_own ty1 tid) -∗ &{α}((l +â‚— 1) ↦∗: ty_own ty2 tid)))%I as "#Hb". - { iIntros "!# H". iApply bor_iff; last done. + { iIntros "!> H". iApply bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl; iFrame; by iApply "Hown". } iDestruct "H" as (st) "H"; iExists st; @@ -150,11 +150,11 @@ Section refcell. iExists γ, _. iFrame "Hst Hn Hshr". iSplitR "Htok2"; last by iExists _; iFrame; rewrite Qp_div_2. iIntros "!> !> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν". - iApply fupd_mask_mono; last iApply "Hh". set_solver. rewrite -lft_dead_or. auto. + iApply fupd_mask_mono; last iApply "Hh". set_solver-. rewrite -lft_dead_or. auto. - iMod (own_alloc (â— (refcell_st_to_R $ Some (static, true, (1/2)%Qp, n)))) as (γ) "Hst". { by apply auth_auth_valid. } iFrame "Htok'". iExists γ, _. iFrame. iSplitR. - { rewrite -step_fupd_intro. auto. set_solver+. } + { rewrite -step_fupd_intro. auto. set_solver-. } iSplitR; [|done]. iExists (1/2)%Qp. rewrite Qp_div_2. iSplitR; [done|]. iApply lft_tok_static. Qed. @@ -182,12 +182,12 @@ Section refcell. Proof. (* TODO : this proof is essentially [solve_proper], but within the logic. It would easily gain from some automation. *) - iIntros (ty1 ty2 EQ qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'. + iIntros (ty1 ty2 EQ qmax qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'. iDestruct (EQ' with "HL") as "#EQ'". iDestruct (refcell_inv_proper with "HL") as "#Hty1ty2"; first done. iDestruct (refcell_inv_proper with "HL") as "#Hty2ty1"; first by symmetry. - iIntros "!# #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)". - iSplit; [|iSplit; iIntros "!# * H"]. + iIntros "!> #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)". + iSplit; [|iSplit; iIntros "!> * H"]. - iPureIntro. simpl. congruence. - destruct vl as [|[[]|]]=>//=. by iApply "Hown". - simpl. iDestruct "H" as (a γ) "[Ha H]". iExists a, γ. iFrame. @@ -208,4 +208,4 @@ Section refcell. Proof. move=>???[|[[]|]]//=. Qed. End refcell. -Hint Resolve refcell_mono' refcell_proper' : lrust_typing. +Global Hint Resolve refcell_mono' refcell_proper' : lrust_typing. diff --git a/theories/typing/lib/refcell/refcell_code.v b/theories/typing/lib/refcell/refcell_code.v index e6aa46e9db7ac380ac045f303f8ed37eb21de089..8e46f894ffe8be47411d1161529ccf80000db75a 100644 --- a/theories/typing/lib/refcell/refcell_code.v +++ b/theories/typing/lib/refcell/refcell_code.v @@ -20,10 +20,10 @@ Section refcell_functions. Lemma refcell_new_type ty `{!TyWf ty} : typed_val (refcell_new ty) (fn(∅; ty) → refcell ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (_ Ï ret arg). inv_vec arg=>x. simpl_subst. iApply (type_new (S ty.(ty_size))); [solve_typing..|]. - iIntros (r tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (r tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hr Hx]". iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x. @@ -54,10 +54,10 @@ Section refcell_functions. Lemma refcell_into_inner_type ty `{!TyWf ty} : typed_val (refcell_into_inner ty) (fn(∅; refcell ty) → ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (_ Ï ret arg). inv_vec arg=>x. simpl_subst. iApply (type_new ty.(ty_size)); [solve_typing..|]. - iIntros (r tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (r tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hr Hx]". iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & >Hx↦ & Hx & Hx†)". subst x. @@ -87,10 +87,10 @@ Section refcell_functions. typed_val refcell_get_mut (fn(∀ α, ∅; &uniq{α} (refcell ty)) → &uniq{α} ty)%T. Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL HC HT". + iIntros (tid qmax) "#LFT #HE Hna HL HC HT". rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try iDestruct "Hx'" as "[]". iAssert (&{α} (∃ (z : Z), lx' ↦ #z) ∗ @@ -139,16 +139,16 @@ Section refcell_functions. typed_val refcell_try_borrow (fn(∀ α, ∅; &shr{α}(refcell ty)) → option (ref α ty)). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. iApply (type_cont [] [Ï âŠ‘â‚— []] (λ _, [x â— box (&shr{α}(refcell ty)); r â— box (option (ref α ty))])); - [iIntros (k)|iIntros "/= !#"; iIntros (k arg); inv_vec arg]; simpl_subst; last first. + [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; simpl_subst; last first. { iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. } iApply type_deref; [solve_typing..|]. - iIntros (x' tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (x' tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]=>//=. iDestruct "Hx'" as (β γ) "#[Hαβ Hinv]". @@ -192,7 +192,7 @@ Section refcell_functions. - iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]". done. iMod (own_update with "Hownst") as "[Hownst Hreading]"; first by apply auth_update_alloc, (op_local_update_discrete _ _ (reading_stR (1/2)%Qp ν)). - rewrite right_id. iExists _, _. iFrame "Htok1 Hreading". + rewrite (right_id None). iExists _, _. iFrame "Htok1 Hreading". iDestruct (lft_intersect_acc with "Hβtok2 Htok2") as (q) "[Htok Hclose]". iApply (fupd_mask_mono (↑lftN)); first done. iMod (rebor _ _ (β ⊓ ν) with "LFT [] Hst") as "[Hst Hh]". done. @@ -201,8 +201,11 @@ Section refcell_functions. iDestruct ("Hclose" with "Htok") as "[$ Htok2]". iExists _. iFrame "∗#". iSplitR "Htok2". + iIntros "!> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. - iNext. iMod "Hν". iApply fupd_mask_mono; last iApply "Hh". - set_solver. rewrite -lft_dead_or. auto. + iNext. iMod "Hν". + iMod fupd_mask_subseteq as "Hclose"; last iMod ("Hh" with "[Hν]") as "$". + { set_solver-. } + * rewrite -lft_dead_or. auto. + * iMod "Hclose". done. + iExists _. iFrame. by rewrite Qp_div_2. } iMod ("Hclose''" with "[$INV] Hna") as "[Hβtok1 Hna]". iMod ("Hclose'" with "[$Hβtok1 $Hβtok2]") as "Hα". @@ -245,17 +248,17 @@ Section refcell_functions. typed_val refcell_try_borrow_mut (fn(∀ α, ∅; &shr{α}(refcell ty)) → option (refmut α ty))%T. Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. iApply (type_cont [] [Ï âŠ‘â‚— []] (λ _, [x â— box (&shr{α}(refcell ty)); r â— box (option (refmut α ty))])); - [iIntros (k)|iIntros "/= !#"; iIntros (k arg); inv_vec arg]; + [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; simpl_subst; last first. { iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. } iApply type_deref; [solve_typing..|]. - iIntros (x' tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (x' tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]=>//=. iDestruct "Hx'" as (β γ) "#[Hαβ Hinv]". @@ -272,14 +275,17 @@ Section refcell_functions. iMod (own_update with "Hownst") as "[Hownst ?]". { by eapply auth_update_alloc, (op_local_update_discrete _ _ (refcell_st_to_R $ Some (ν, true, (1/2)%Qp, xH))). } - rewrite right_id. + rewrite (right_id None). iApply fupd_wp. iApply (fupd_mask_mono (↑lftN)); first done. iMod (rebor _ _ (β ⊓ ν) with "LFT [] Hb") as "[Hb Hbh]". done. { iApply lft_intersect_incl_l. } iModIntro. iMod ("Hclose''" with "[Hlx Hownst Hbh Htok1] Hna") as "[Hβtok Hna]". { iExists _. iFrame. iNext. iSplitL "Hbh". - iIntros "Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν". - iApply fupd_mask_mono; last iApply "Hbh". set_solver. rewrite -lft_dead_or. auto. + iMod fupd_mask_subseteq as "Hclose"; last iMod ("Hbh" with "[Hν]") as "$". + { set_solver-. } + * rewrite -lft_dead_or. auto. + * iMod "Hclose". done. - iSplitL; [|done]. iExists _. iFrame. by rewrite Qp_div_2. } iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". iApply (type_type _ _ _ diff --git a/theories/typing/lib/refcell/refmut.v b/theories/typing/lib/refcell/refmut.v index 6a2cecc462b0f0fe25e9faf33140443e23c2a577..5136a8e21a28c72dcf395fc9a57c75e4bba67df3 100644 --- a/theories/typing/lib/refcell/refmut.v +++ b/theories/typing/lib/refcell/refmut.v @@ -65,7 +65,7 @@ Section refmut. iIntros (??????) "#? H". iDestruct "H" as (lv lrc) "[#Hf #H]". iExists _, _. iSplit. - by iApply frac_bor_shorten. - - iIntros "!# * % Htok". + - iIntros "!> * % Htok". iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. { iApply lft_intersect_mono. iApply lft_incl_refl. done. } iMod ("H" with "[] Htok") as "Hshr". done. iModIntro. iNext. @@ -84,27 +84,29 @@ Section refmut. Global Instance refmut_mono E L : Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) refmut. Proof. - intros α1 α2 Hα ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (q) "HL". + intros α1 α2 Hα ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (qmax qL) "HL". iDestruct (Hty with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". - iIntros "!# #HE". iDestruct ("Hα" with "HE") as "Hα1α2". + iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2. iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro]. - done. - iIntros (tid [|[[]|][|[[]|][]]]) "H"=>//=. iDestruct "H" as (ν γ q' β ty') "(Hb & #H⊑ & #Hinv & Hν & Hown)". iExists ν, γ, q', β, ty'. iFrame "∗#". iSplit. + iApply bor_shorten; last iApply bor_iff; last done. - * iApply lft_intersect_mono; first done. iApply lft_incl_refl. + * iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. + iApply lft_incl_refl. * iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]"; iExists vl; iFrame; by iApply "Ho". - + by iApply lft_incl_trans. + + iApply lft_incl_trans; last done. by iApply lft_incl_syn_sem. - iIntros (κ tid l) "H /=". iDestruct "H" as (lv lrc) "H". iExists lv, lrc. - iDestruct "H" as "[$ #H]". iIntros "!# * % Htok". + iDestruct "H" as "[$ #H]". iIntros "!> * % Htok". iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. done. iApply lft_incl_refl. } - iMod ("H" with "[] Htok") as "Hshr". done. iModIntro. iNext. + { 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. done. iApply lft_incl_refl. - by iApply "Hs". + iApply ty_shr_mono; try done. + + iApply lft_intersect_mono. by iApply lft_incl_syn_sem. iApply lft_incl_refl. + + by iApply "Hs". Qed. Global Instance refmut_mono_flip E L : Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) refmut. @@ -122,4 +124,4 @@ Section refmut. Proof. intros. by eapply refmut_proper. Qed. End refmut. -Hint Resolve refmut_mono' refmut_proper' : lrust_typing. +Global Hint Resolve refmut_mono' refmut_proper' : lrust_typing. diff --git a/theories/typing/lib/refcell/refmut_code.v b/theories/typing/lib/refcell/refmut_code.v index 550751122920464cf292fe1c9dcf0574f33eec34..47db4a1ade14c687653adafeeb192af9ad665c6d 100644 --- a/theories/typing/lib/refcell/refmut_code.v +++ b/theories/typing/lib/refcell/refmut_code.v @@ -19,10 +19,10 @@ Section refmut_functions. Lemma refmut_deref_type ty `{!TyWf ty} : typed_val refmut_deref (fn(∀ '(α, β), ∅; &shr{α}(refmut β ty)) → &shr{α}ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=. iDestruct "Hx'" as (lv lrc) "#(Hfrac & #Hshr)". @@ -30,21 +30,21 @@ Section refmut_functions. [solve_typing..|]. iMod (frac_bor_acc with "LFT Hfrac Hα1") as (qlx') "[H↦ Hcloseα1]". done. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iMod (lctx_lft_alive_tok β with "HE HL") as (qβ) "(Hβ & HL & Hclose'')"; + iMod (lctx_lft_alive_tok_noend β with "HE HL") as (qβ) "(Hβ & HL & Hclose'')"; [solve_typing..|]. iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hcloseβα2]". wp_bind (!#lx')%E. iApply (wp_step_fupd with "[Hα2β]"); - [done| |by iApply ("Hshr" with "[] Hα2β")|]; first done. + [| |by iApply ("Hshr" with "[] Hα2β")|]; [done..|]. iMod "H↦" as "[H↦1 H↦2]". wp_read. iIntros "[#Hshr' Hα2β] !>". wp_let. iDestruct ("Hcloseβα2" with "Hα2β") as "[Hβ Hα2]". iMod ("Hcloseα1" with "[$H↦1 $H↦2]") as "Hα1". iMod ("Hclose''" with "Hβ HL") as "HL". iMod ("Hclose'" with "[$] HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|]. iApply (type_type _ _ _ [ x â— box (&shr{α}(refmut β ty)); #lv â— &shr{α}ty] with "[] LFT HE Hna HL Hk"); last first. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. iFrame. iApply (ty_shr_mono with "[] Hshr'"). - by iApply lft_incl_glb; last iApply lft_incl_refl. } + iApply lft_incl_glb; last iApply lft_incl_refl. by iApply lft_incl_syn_sem. } iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. @@ -57,10 +57,10 @@ Section refmut_functions. typed_val refmut_derefmut (fn(∀ '(α, β), ∅; &uniq{α}(refmut β ty)) → &uniq{α}ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. iMod (bor_exists with "LFT Hx'") as (vl) "H". done. @@ -87,12 +87,13 @@ Section refmut_functions. wp_bind (!#lx')%E. iMod (bor_unnest with "LFT Hb") as "Hb"; first done. wp_read. wp_let. iMod "Hb". iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]". iMod ("Hclose'" with "Hα HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|]. iApply (type_type _ _ _ [ x â— box (uninit 1); #lv â— &uniq{α}ty] with "[] LFT HE Hna HL Hk"); last first. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. iFrame. iApply (bor_shorten with "[] Hb"). - by iApply lft_incl_glb; [iApply lft_incl_glb|iApply lft_incl_refl]. } + iApply lft_incl_glb; [iApply lft_incl_glb|iApply lft_incl_refl]; + try by iApply lft_incl_syn_sem. done. } iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. @@ -111,9 +112,9 @@ Section refmut_functions. Lemma refmut_drop_type ty `{!TyWf ty} : typed_val refmut_drop (fn(∀ α, ∅; refmut α ty) → unit). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val. + iIntros (tid qmax) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val. destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". iDestruct "Hx" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hx"; simpl; try iDestruct "Hx" as "[_ >[]]". @@ -125,7 +126,7 @@ Section refmut_functions. iMod (lft_incl_acc with "Hαβ Hα") as (qβ) "[Hβ Hcloseα]". done. iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hcloseβ)"; [done..|]. iDestruct "INV" as (st) "(H↦lrc & Hâ— & INV)". wp_read. wp_let. wp_op. wp_write. - iAssert (|={↑lftN ∪ ↑histN}[↑histN]â–·=> refcell_inv tid lrc γ β ty')%I + iAssert (|={↑lftN ∪ lft_userE}[lft_userE]â–·=> refcell_inv tid lrc γ β ty')%I with "[H↦lrc Hâ— Hâ—¯ Hν INV]" as "INV". { iDestruct (own_valid_2 with "Hâ— Hâ—¯") as %[[[=]|(? & [[? q'] ?] & [= <-] & Hst & INCL)]%option_included _] @@ -137,7 +138,8 @@ Section refmut_functions. iMod (own_update_2 with "Hâ— Hâ—¯") as "$". { apply auth_update_dealloc. rewrite -(right_id None op (Some _)). apply cancel_local_update_unit, _. } - iDestruct "INV" as "(H†& Hq & _)". iApply "H†". + iDestruct "INV" as "(H†& Hq & _)". + iApply "H†". iDestruct "Hq" as (q) "(<- & ?)". iFrame. - simpl in *. setoid_subst. iExists (Some (_, _, _, _)). iMod (own_update_2 with "Hâ— Hâ—¯") as "$". @@ -146,11 +148,11 @@ Section refmut_functions. iDestruct "INV" as "(H†& Hq & _)". rewrite /= (_:Z.neg (1%positive â‹… n') + 1 = Z.neg n'); last (rewrite pos_op_plus; lia). iFrame. - iApply step_fupd_intro; [set_solver+|]. iSplitL; [|done]. + iApply step_fupd_intro; [set_solver-|]. iSplitL; [|done]. iDestruct "Hq" as (q' ?) "?". iExists (q+q')%Qp. iFrame. rewrite assoc (comm _ q'' q) //. } - wp_bind Endlft. iApply (wp_mask_mono _ (↑lftN ∪ ↑histN)); first done. - iApply (wp_step_fupd with "INV"); [done|set_solver|]. wp_seq. iIntros "{Hb} Hb !>". + wp_bind Endlft. iApply (wp_mask_mono _ (↑lftN ∪ lft_userE)); first done. + iApply (wp_step_fupd with "INV"); [done|set_solver-|]. wp_seq. iIntros "{Hb} Hb !>". iMod ("Hcloseβ" with "Hb Hna") as "[Hβ Hna]". iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". wp_seq. iApply (type_type _ _ _ [ #lx â— box (uninit 2)] @@ -178,10 +180,11 @@ Section refmut_functions. typed_val call_once (fn(∀ α, ∅; fty, &uniq{α}ty1) → &uniq{α}ty2) → typed_val (refmut_map call_once) (fn(∀ α, ∅; refmut α ty1, fty) → refmut α ty2). Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret arg). + intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (α Ï ret arg). inv_vec arg=>ref env. simpl_subst. iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". + iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done. iDestruct "Href" as "[Href Href†]". iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl; @@ -193,7 +196,7 @@ Section refmut_functions. iIntros (lx) "(H†& Hlx)". rewrite heap_mapsto_vec_singleton. wp_let. wp_write. wp_let. rewrite tctx_hasty_val. iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|]. - iMod (lctx_lft_alive_tok Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. + iMod (lctx_lft_alive_tok_noend Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". iDestruct (lft_intersect_acc with "Hαν HÏ") as (?) "[HÎ±Î½Ï Hclose4]". rewrite -[Ï in (α ⊓ ν) ⊓ Ï](right_id_L _ (⊓)). @@ -256,10 +259,11 @@ Section refmut_functions. typed_val (refmut_map_split call_once) (fn(∀ α, ∅; refmut α ty, fty) → Î [refmut α ty1; refmut α ty2]). Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret arg). + intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (α Ï ret arg). inv_vec arg=>refmut env. simpl_subst. iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Hrefmut & Henv & _)". + iIntros (tid qmax) "#LFT #HE Hna HL Hk (#Hf' & Hrefmut & Henv & _)". rewrite (tctx_hasty_val _ refmut). destruct refmut as [[|lrefmut|]|]; try done. iDestruct "Hrefmut" as "[Hrefmut Hrefmut†]". iDestruct "Hrefmut" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hrefmut"; simpl; @@ -271,10 +275,10 @@ Section refmut_functions. iIntros (lx) "(H†& Hlx)". rewrite heap_mapsto_vec_singleton. wp_let. wp_write. wp_let. rewrite tctx_hasty_val. iMod (lctx_lft_alive_tok α with "HE HL") as (?) "(Hα & HL & Hclose1)";[solve_typing..|]. - iMod (lctx_lft_alive_tok Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. + iMod (lctx_lft_alive_tok_noend Ï with "HE HL") as (?) "(HÏ & HL & Hclose2)";[solve_typing..|]. iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". iDestruct (lft_intersect_acc with "Hαν HÏ") as (?) "[HÎ±Î½Ï Hclose4]". - rewrite -[Ï in (α ⊓ ν) ⊓ Ï](right_id_L static (⊓)). + rewrite -[Ï in (α ⊓ ν) ⊓ Ï]right_id_L. iApply (type_call_iris _ [α ⊓ ν; Ï] (α ⊓ ν) [_; _] with "LFT HE Hna [HανÏ] Hf' [$Henv Hlx H†Hbor]"); [solve_typing|done| |]. { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. diff --git a/theories/typing/lib/rwlock/rwlock.v b/theories/typing/lib/rwlock/rwlock.v index a36127024919eb8515a2e4b1632db9d32e5b7ba5..298264f6ee26c34ab4d126a331da1df75af39574 100644 --- a/theories/typing/lib/rwlock/rwlock.v +++ b/theories/typing/lib/rwlock/rwlock.v @@ -1,6 +1,6 @@ From iris.bi Require Import fractional. From iris.proofmode Require Import tactics. -From iris.algebra Require Import auth excl csum frac agree numbers. +From iris.algebra Require Import auth csum frac agree excl numbers. From gpfsl.gps Require Import middleware protocols. From lrust.lifetime Require Import at_borrow. From lrust.typing Require Import typing. @@ -176,7 +176,7 @@ Section rwlock_inv. GPS_SWSharedReader l t () #0 1%Qp γ | Some (inr (ν, q, n)) => (* Locked for read. *) GPS_SWSharedWriter l t () #(Z.pos n) γ ∗ - ∃ q', â–¡ (1.[ν] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ [†ν]) ∗ + ∃ q', â–¡ (1.[ν] ={↑lftN ∪ lft_userE}[lft_userE]â–·=∗ [†ν]) ∗ ([†ν] ={↑lftN ∪ ↑histN}=∗ &{α}tyO) ∗ (â–¡ tyS (α ⊓ ν)) ∗ ⌜(q + q')%Qp = 1%Qp⌠∗ q'.[ν] ∗ @@ -191,9 +191,12 @@ Section rwlock_inv. Definition rwlock_proto' l γ γs κ (tyO: vProp) (tyS: lft → vProp): vProp := (∃ t v, GPS_aSP_Reader rwlockN (rwlock_interp γs κ tyO tyS) rwlock_noCAS l κ t () v γ)%I. - Definition rwlock_proto l γ γs κ tyO tyS tid ty: vProp := - ((â–· â–¡ (tyO ↔ (l +â‚— 1) ↦∗: ty.(ty_own) tid)) ∗ - (â–· â–¡ (∀ α, tyS α ↔ ty.(ty_shr) α tid (l +â‚— 1))) ∗ + Definition rwlock_proto l γ γs κ tyO tyS tid_own tid_shr ty: vProp := + (* We close tyO and tyS over iff here because GPS_INV is not closed + under iff. Not doing this would make impossible to prove + [rwlock_proto_proper]. *) + ((â–· â–¡ (tyO ↔ (l +â‚— 1) ↦∗: ty.(ty_own) tid_own)) ∗ + (â–· â–¡ (∀ α, tyS α ↔ ty.(ty_shr) α tid_shr (l +â‚— 1))) ∗ rwlock_proto' l γ γs κ tyO tyS)%I. Lemma rwlock_interp_comparable γs α tyO tyS l γ t s v (n: Z): @@ -206,11 +209,11 @@ Section rwlock_inv. iExists _; iPureIntro ;(split; [done|by constructor]). Qed. - Global Instance rwlock_proto_persistent l γ γs κ tyO tyS tid ty : - Persistent (rwlock_proto l γ γs κ tyO tyS tid ty) := _. + Global Instance rwlock_proto_persistent l γ γs κ tyO tyS tid_own tid_shr ty : + Persistent (rwlock_proto l γ γs κ tyO tyS tid_own tid_shr ty) := _. - Global Instance rwlock_proto_type_ne n l γ γs κ tyO tyS tid : - Proper (type_dist2 n ==> dist n) (rwlock_proto l γ γs κ tyO tyS tid). + Global Instance rwlock_proto_type_ne n l γ γs κ tyO tyS tid_own tid_shr : + Proper (type_dist2 n ==> dist n) (rwlock_proto l γ γs κ tyO tyS tid_own tid_shr). Proof. move => ???. apply bi.sep_ne; [|apply bi.sep_ne]; [..|done]; @@ -222,23 +225,24 @@ Section rwlock_inv. apply bi.iff_ne; [done|apply ty_shr_type_dist]; [by apply type_dist2_S|done..]. Qed. - Global Instance rwlock_proto_ne l γ γs κ tyO tyS tid : - NonExpansive (rwlock_proto l γ γs κ tyO tyS tid). + Global Instance rwlock_proto_ne l γ γs κ tyO tyS tid_own tid_shr : + NonExpansive (rwlock_proto l γ γs κ tyO tyS tid_own tid_shr). Proof. intros n ???. by eapply rwlock_proto_type_ne, type_dist_dist2. Qed. - Lemma rwlock_proto_proper E L ty1 ty2 q : + Lemma rwlock_proto_proper E L ty1 ty2 qmax qL : eqtype E L ty1 ty2 → - llctx_interp L q -∗ ∀ l γ γs κ tyO tyS tid, â–¡ (elctx_interp E -∗ - rwlock_proto l γ γs κ tyO tyS tid ty1 -∗ rwlock_proto l γ γs κ tyO tyS tid ty2). + llctx_interp_noend qmax L qL -∗ ∀ l γ γs κ tyO tyS tid_own tid_shr, â–¡ (elctx_interp E -∗ + rwlock_proto l γ γs κ tyO tyS tid_own tid_shr ty1 -∗ + rwlock_proto l γ γs κ tyO tyS tid_own tid_shr ty2). Proof. (* TODO : this proof is essentially [solve_proper], but within the logic. It would easily gain from some automation. *) rewrite eqtype_unfold. iIntros (Hty) "HL". - iDestruct (Hty with "HL") as "#Hty". iIntros "* !# #HE H". + iDestruct (Hty with "HL") as "#Hty". iIntros "* !> #HE H". iDestruct ("Hty" with "HE") as "(% & #Hown & #Hshr)". - iDestruct "H" as "(#EqO & #EqS & $)". iSplit; iIntros "!> !#". + iDestruct "H" as "(#EqO & #EqS & $)". iSplit; iIntros "!> !>". - iSplit; iIntros "H1". + iDestruct ("EqO" with "H1") as (?) "[? ?]". iExists _. iFrame. by iApply "Hown". @@ -254,18 +258,18 @@ Section rwlock_inv. (q).[α] -∗ â–· ⎡ hist_inv ⎤ -∗ &{α} ((l +â‚— 1) ↦∗{1}: ty_own ty tid)%I -∗ &{α} (∃ n : Z, l ↦{1} #n ∗ ⌜-1 ≤ nâŒ) ={E}=∗ - (q).[α] ∗ (∃ γ γs tyO tyS, rwlock_proto l γ γs α tyO tyS tid ty). + (q).[α] ∗ (∃ γ γs tyO tyS, rwlock_proto l γ γs α tyO tyS tid tid ty). Proof. iIntros "#LFT [Htok Htok'] #HINV Hvl H". set tyO : vProp := ((l +â‚— 1) ↦∗{1}: ty.(ty_own) tid)%I. set tyS : lft → vProp := (λ α', ty.(ty_shr) α' tid (l +â‚— 1))%I. iAssert (â–¡ (tyO ↔ (l +â‚— 1) ↦∗{1}: ty.(ty_own) tid))%I as "#EqO". - { iIntros "!#". by iApply (bi.iff_refl True%I). } + { iIntros "!>". by iApply (bi.iff_refl True%I). } iAssert (â–¡ (∀ α', tyS α' ↔ ty.(ty_shr) α' tid (l +â‚— 1)))%I as "#EqS". - { iIntros "!#" (?). by iApply (bi.iff_refl True%I). } + { iIntros "!>" (?). by iApply (bi.iff_refl True%I). } iAssert (â–¡ (∀ κ q' E', ⌜lftE ⊆ E'⌠-∗ &{κ} tyO -∗ (q').[κ] ={E'}=∗ (â–¡ tyS κ) ∗ (q').[κ]))%I as "#Share". - { iIntros "!#" (????) "tyO tok". + { iIntros "!>" (????) "tyO tok". by iMod (ty_share with "LFT tyO tok") as "[#? $]". } iMod (bor_acc_cons with "LFT H Htok") as "[H Hclose]"; [done|]. set P : gname → val → vProp := @@ -331,7 +335,7 @@ Section rwlock_inv. iExists q'. iFrame "H†Hinh HtyS Eqq Htok'". iDestruct "Eqq" as %Eqq. rewrite -Eqq. iDestruct "R" as "[? $]". - iFrame "H". by iDestruct (GPS_SWWriter_shared with "W") as "[$ $]". } - { iIntros "!#" (t γ s v) "Int !> !>". + { iIntros "!>" (t γ s v) "Int !> !>". iDestruct "Int" as (st) "(#Eq & Hshr & Hst & Int)". iExists st. iFrame "Eq Hst". destruct st as [[|[[] n]]|]; [done|..]. - iDestruct "Int" as "[W Int]". @@ -343,7 +347,6 @@ Section rwlock_inv. iFrame "EqO EqS". iExists _, _. iFrame "R". Qed. - Lemma rwlock_interp_read_acq l t s v γ γs α tyO tyS tid: (â–½{tid} rwlock_interp γs α tyO tyS false l γ t s v : vProp) -∗ ∃ st, ⌜v = #(Z_of_rwlock_st st)âŒ. @@ -353,6 +356,21 @@ Section rwlock_inv. iDestruct (acq_pure_elim with "Int") as "?". iExists _. iFrame. Qed. + Lemma rwlock_proto_change_tid_own l γ γs α tyS tyO tid_own1 tid_own2 tid_shr ty : + Send ty → + rwlock_proto l γ γs α tyO tyS tid_own1 tid_shr ty ≡ + rwlock_proto l γ γs α tyO tyS tid_own2 tid_shr ty. + Proof. + intros ?. unfold rwlock_proto. do 7 f_equiv. iApply send_change_tid'. + Qed. + + Lemma rwlock_proto_change_tid_shr l γ γs α tyS tyO tid_own tid_shr1 tid_shr2 ty : + Sync ty → + rwlock_proto l γ γs α tyO tyS tid_own tid_shr1 ty ≡ + rwlock_proto l γ γs α tyO tyS tid_own tid_shr2 ty. + Proof. + intros ?. unfold rwlock_proto. do 7 f_equiv. iApply sync_change_tid'. + Qed. End rwlock_inv. Section rwlock. @@ -374,7 +392,7 @@ Section rwlock. | _ => False end)%I; ty_shr κ tid l := - (∃ α, κ ⊑ α ∗ ∃ γ γs tyO tyS, rwlock_proto l γ γs α tyO tyS tid ty)%I |}. + (∃ α, κ ⊑ α ∗ ∃ γ γs tyO tyS, rwlock_proto l γ γs α tyO tyS tid tid ty)%I |}. Next Obligation. iIntros (??[|[[]|]]); try iIntros "[? []]". rewrite ty_size_eq. iIntros "[_ [_ %]] !% /=". congruence. @@ -422,11 +440,11 @@ Section rwlock. Proof. (* TODO : this proof is essentially [solve_proper], but within the logic. It would easily gain from some automation. *) - iIntros (ty1 ty2 EQ qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'. + iIntros (ty1 ty2 EQ qmax qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'. iDestruct (EQ' with "HL") as "#EQ'". iDestruct (rwlock_proto_proper with "HL") as "#Hty1ty2"; first done. - iIntros "!# #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)". - iSplit; [|iSplit; iIntros "!# * H"]. + iIntros "!> #HE". iDestruct ("EQ'" with "HE") as "(% & #Hown & #Hshr)". + iSplit; [|iSplit; iIntros "!> * H"]. - iPureIntro. simpl. congruence. - destruct vl as [|[[]|]]; try done. iDestruct "H" as "[$ [$ H]]". by iApply "Hown". @@ -454,18 +472,8 @@ Section rwlock. Proof. move=>??????/=. apply bi.exist_mono=>?. apply bi.sep_mono_r. do 4 apply bi.exist_mono=>?. - apply bi.sep_mono; last apply bi.sep_mono_l; - apply bi.later_mono; iIntros "#H !#". - - iSplit; iIntros "H1". - + iDestruct ("H" with "H1") as (?) "[??]". - iExists _. iFrame. by iApply send_change_tid. - + iApply "H". iDestruct "H1" as (?) "[??]". - iExists _. iFrame. by iApply send_change_tid. - - iIntros (α). iSpecialize ("H" $! α). iSplit; iIntros "H1". - + iApply sync_change_tid. by iApply "H". - + iApply "H". by iApply sync_change_tid. + by rewrite rwlock_proto_change_tid_own rwlock_proto_change_tid_shr. Qed. - End rwlock. -Hint Resolve rwlock_mono' rwlock_proper' : lrust_typing. +Global Hint Resolve rwlock_mono' rwlock_proper' : lrust_typing. diff --git a/theories/typing/lib/rwlock/rwlock_code.v b/theories/typing/lib/rwlock/rwlock_code.v index 0e9e4c64091ec164992c47b4d4b880df824aad8d..619284d0a41db33c546eea6b7556cae7d478f99f 100644 --- a/theories/typing/lib/rwlock/rwlock_code.v +++ b/theories/typing/lib/rwlock/rwlock_code.v @@ -35,10 +35,10 @@ Section rwlock_functions. Lemma rwlock_new_type ty `{!TyWf ty} : typed_val (rwlock_new ty) (fn(∅; ty) → rwlock ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". - iIntros (_ ret Ï arg). inv_vec arg=>x. simpl_subst. + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (_ Ï ret arg). inv_vec arg=>x. simpl_subst. iApply (type_new (S ty.(ty_size))); [solve_typing..|]. - iIntros (r tid) "#LFT HE Hna HL Hk HT". simpl_subst. + iIntros (r tid qmax) "#LFT HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". iDestruct "Hx" as (vl) "[Hx↦ Hx]". @@ -71,10 +71,10 @@ Section rwlock_functions. Lemma rwlock_into_inner_type ty `{!TyWf ty} : typed_val (rwlock_into_inner ty) (fn(∅; rwlock ty) → ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". - iIntros (_ ret Ï arg). inv_vec arg=>x. simpl_subst. + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (_ Ï ret arg). inv_vec arg=>x. simpl_subst. iApply (type_new ty.(ty_size)); [solve_typing..|]. - iIntros (r tid) "#LFT HE Hna HL Hk HT". simpl_subst. + iIntros (r tid qmax) "#LFT HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". @@ -104,10 +104,10 @@ Section rwlock_functions. Lemma rwlock_get_mut_type ty `{!TyWf ty} : typed_val rwlock_get_mut (fn(∀ α, ∅; &uniq{α} (rwlock ty)) → &uniq{α} ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iIntros (tid) "#LFT HE Hna HL HC HT". + iIntros (tid qmax) "#LFT HE Hna HL HC HT". rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try iDestruct "Hx'" as "[]". iAssert (&{α} (⎡ hist_inv ⎤ ∗ ∃ (z : Z), lx' ↦ #z ∗ ⌜-1 ≤ zâŒ) ∗ @@ -156,21 +156,21 @@ Section rwlock_functions. typed_val rwlock_try_read (fn(∀ α, ∅; &shr{α}(rwlock ty)) → option (rwlockreadguard α ty)). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. iApply (type_cont [] [Ï âŠ‘â‚— []] (λ _, [x â— box (&shr{α}(rwlock ty)); r â— box (option (rwlockreadguard α ty))])); - [iIntros (k)|iIntros "/= !#"; iIntros (k arg); inv_vec arg]; + [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; simpl_subst; last first. { iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. } iApply (type_cont [] [Ï âŠ‘â‚— []] (λ _, [x â— _; x' â— _; r â— _])); - [iIntros (loop)|iIntros "/= !#"; iIntros (loop arg); inv_vec arg]; + [iIntros (loop)|iIntros "/= !>"; iIntros (loop arg); inv_vec arg]; simpl_subst. { iApply type_jump; solve_typing. } - iIntros (tid) "#LFT #HE Hna HL Hk HT". + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]; try done. iDestruct "Hx'" as (β) "#[Hαβ Hinv]". @@ -210,7 +210,7 @@ Section rwlock_functions. (∃ qν ν, (qν).[ν] ∗ (â–¡ tyS (β ⊓ ν)) ∗ rwown γs (reading_st qν ν) ∗ - (â–¡ ((1).[ν] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ [†ν])) ∗ + (â–¡ ((1).[ν] ={↑lftN ∪ lft_userE}[lft_userE]â–·=∗ [†ν])) ∗ GPS_SWSharedReader lx t (() : unitProtocol) #(Z_of_rwlock_st st' + 1) qν γ))%I. set Q1: time → () → vProp Σ := (λ t _, rwlock_interp γs β tyO tyS false lx γ t () #(Z_of_rwlock_st st'))%I. @@ -292,7 +292,7 @@ Section rwlock_functions. with "[] LFT HE Hna HL Hk"); first last. { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. iFrame "Hx Hr". - iExists _, _, _, _. iFrame "Hν Hreading H†Hαβ". + iExists _, _, _, _, _. iFrame "Hν Hreading H†Hαβ". iDestruct ("EqS" with "Hshr") as "Hshr". iSplitL "Hshr". - iApply ty_shr_mono; last done. @@ -332,16 +332,16 @@ Section rwlock_functions. typed_val rwlock_try_write (fn(∀ α, ∅; &shr{α}(rwlock ty)) → option (rwlockwriteguard α ty)). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply (type_cont [_] [Ï âŠ‘â‚— []] (λ r, [x â— box (&shr{α}(rwlock ty)); (r!!!0%fin:val) â— box (option (rwlockwriteguard α ty))])); - [iIntros (k)|iIntros "/= !#"; iIntros (k arg); inv_vec arg=>r]; + [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg=>r]; simpl_subst; last first. { iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. } iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. iApply type_deref; [solve_typing..|]. - iIntros (x' tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (x' tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "(Hx & Hx' & Hr)". destruct x' as [[|lx|]|]; try done. iDestruct "Hx'" as (β) "#[Hαβ Hinv]". @@ -372,7 +372,7 @@ Section rwlock_functions. iPureIntro. by constructor. } { simpl. iSplitL ""; last done. rewrite -(bi.True_sep' (∀ _ _, _)%I). - iApply (rel_sep_objectively with "[$rTrue]"). + iApply (rel_sep_objectively with "[$rTrue]"). iIntros "!>" (t1 [] [_ Ext1]). iSplit; last first; last iSplitL ""; last iSplitL "". { iIntros "!>" (v ?) "!>". iSplit; last iSplit; by iIntros "$". } { rewrite /rwlock_noCAS. iIntros "!> _" (Eq). by inversion Eq. } @@ -403,7 +403,7 @@ Section rwlock_functions. with "[] LFT HE Hna HL Hk"); first last. { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. iFrame "Hx Hr". - iExists γs, β. iSplitL "Hβ"; [by iApply (bor_iff with "[$EqO] Hβ")|]. + iExists γs, β, _. iSplitL "Hβ"; [by iApply (bor_iff with "[$EqO] Hβ")|]. iFrame "Hαβ wst". iExists _,_,_. iSplitL "W". - iExists _. iDestruct "W" as "[$ _]". - iFrame "EqO EqS". iExists _,_. iFrame "R AT". } diff --git a/theories/typing/lib/rwlock/rwlockreadguard.v b/theories/typing/lib/rwlock/rwlockreadguard.v index 432dc93804a0aac60ad3d3f9db5d45de81b84e12..8e2910431db1fa6877fdb805e120a793442ea8ca 100644 --- a/theories/typing/lib/rwlock/rwlockreadguard.v +++ b/theories/typing/lib/rwlock/rwlockreadguard.v @@ -21,11 +21,11 @@ Section rwlockreadguard. ty_own tid vl := match vl return _ with | [ #(LitLoc l) ] => - ∃ ν q γs β, ty.(ty_shr) (α ⊓ ν) tid (l +â‚— 1) ∗ + ∃ ν q γs β tid_own, ty.(ty_shr) (α ⊓ ν) tid (l +â‚— 1) ∗ α ⊑ β ∗ q.[ν] ∗ rwown γs (reading_st q ν) ∗ - â–¡ (1.[ν] ={↑lftN ∪ ↑histN}[↑histN]â–·=∗ [†ν]) ∗ + â–¡ (1.[ν] ={↑lftN ∪ lft_userE}[lft_userE]â–·=∗ [†ν]) ∗ (∃ γ tyO tyS, (∃ t n, GPS_SWSharedReader l t () #n q γ) ∗ - rwlock_proto l γ γs β tyO tyS tid ty) + rwlock_proto l γ γs β tyO tyS tid_own tid ty) | _ => False end; ty_shr κ tid l := @@ -44,6 +44,7 @@ Section rwlockreadguard. iMod (bor_exists with "LFT Hb") as (q') "Hb". done. iMod (bor_exists with "LFT Hb") as (γs) "Hb". done. iMod (bor_exists with "LFT Hb") as (β) "Hb". done. + iMod (bor_exists with "LFT Hb") as (tid_own) "Hb". done. iMod (bor_sep with "LFT Hb") as "[Hshr Hb]". done. iMod (bor_persistent with "LFT Hshr Htok") as "[#Hshr Htok]". done. iMod (bor_sep with "LFT Hb") as "[Hαβ Hb]". done. @@ -78,23 +79,23 @@ Section rwlockreadguard. Global Instance rwlockreadguard_mono E L : Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockreadguard. Proof. - iIntros (α1 α2 Hα ty1 ty2 Hty q) "HL". + iIntros (α1 α2 Hα ty1 ty2 Hty qmax qL) "HL". iDestruct (proj1 Hty with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". iDestruct (rwlock_proto_proper with "HL") as "#Hty1ty2"; first done. - iIntros "!# #HE". iDestruct ("Hα" with "HE") as "Hα1α2". + iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2. iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [|iSplit; iModIntro]. - done. - iIntros (tid [|[[]|][]]) "H"; try done. - iDestruct "H" as (ν q' γs β) "(#Hshr & #H⊑ & Htok & Hown & Hinh & Hinv)". - iExists ν, q', γs, β. iFrame "∗". iSplit; last iSplit. + iDestruct "H" as (ν q' γs β tid_own) "(#Hshr & #H⊑ & Htok & Hown & Hinh & Hinv)". + iExists ν, q', γs, β, tid_own. iFrame "∗". iSplit; last iSplit. + iApply ty_shr_mono; last by iApply "Hs". - iApply lft_intersect_mono. done. iApply lft_incl_refl. - + by iApply lft_incl_trans. + iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. + + iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done. + iDestruct "Hinv" as (γ tyO tyS) "(R & inv)". iExists γ, tyO, tyS. iFrame "R". by iApply ("Hty1ty2" with "HE"). - iIntros (κ tid l) "H". iDestruct "H" as (l') "[Hf Hshr]". iExists l'. iFrame. iApply ty_shr_mono; last by iApply "Hs". - iApply lft_intersect_mono. done. iApply lft_incl_refl. + iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. Qed. Global Instance rwlockreadguard_mono_flip E L : Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) @@ -120,8 +121,19 @@ Section rwlockreadguard. Qed. (* POSIX requires the unlock to occur from the thread that acquired - the lock, so Rust does not implement Send for RwLockWriteGuard. We could - prove this. *) + the lock, so Rust does not implement Send for RwLockWriteGuard. We can + prove this for our spinlock implementation, however. *) + Global Instance rwlockreadguard_send α ty : + Sync ty → Send (rwlockreadguard α ty). + Proof. + iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H". + iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_entails. + repeat lazymatch goal with + | |- (ty_shr _ _ _ _) ≡ (ty_shr _ _ _ _) => by apply sync_change_tid' + | |- (rwlock_proto _ _ _ _ _ _ _ _ _) ≡ _ => by apply rwlock_proto_change_tid_shr + | |- _ => f_equiv + end. + Qed. End rwlockreadguard. -Hint Resolve rwlockreadguard_mono' rwlockreadguard_proper' : lrust_typing. +Global Hint Resolve rwlockreadguard_mono' rwlockreadguard_proper' : lrust_typing. diff --git a/theories/typing/lib/rwlock/rwlockreadguard_code.v b/theories/typing/lib/rwlock/rwlockreadguard_code.v index acb8d15d890fefa911485a6c135d406e3ca8520b..3ac037613404cb0df1884bc6d673353ffd56b955 100644 --- a/theories/typing/lib/rwlock/rwlockreadguard_code.v +++ b/theories/typing/lib/rwlock/rwlockreadguard_code.v @@ -23,10 +23,10 @@ Section rwlockreadguard_functions. typed_val rwlockreadguard_deref (fn(∀ '(α, β), ∅; &shr{α}(rwlockreadguard β ty)) → &shr{α} ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. iDestruct "Hx'" as (l') "#[Hfrac Hshr]". @@ -34,12 +34,13 @@ Section rwlockreadguard_functions. iMod (frac_bor_acc with "LFT Hfrac Hα") as (qlx') "[H↦ Hcloseα]". done. rewrite heap_mapsto_vec_singleton. wp_read. wp_op. wp_let. iMod ("Hcloseα" with "[$H↦]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|]. iApply (type_type _ _ _ [ x â— box (&shr{α}(rwlockreadguard β ty)); #(l' +â‚— 1) â— &shr{α}ty] with "[] LFT HE Hna HL Hk"); first last. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - iFrame. iApply (ty_shr_mono with "[] Hshr"). iApply lft_incl_glb; first done. + iFrame. iApply (ty_shr_mono with "[] Hshr"). + iApply lft_incl_glb; first by iApply lft_incl_syn_sem. by iApply lft_incl_refl. } iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. @@ -62,18 +63,18 @@ Section rwlockreadguard_functions. Lemma rwlockreadguard_drop_type ty `{!TyWf ty} : typed_val rwlockreadguard_drop (fn(∀ α, ∅; rwlockreadguard α ty) → unit). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. iApply (type_cont [] [Ï âŠ‘â‚— []] (λ _, [x â— _; x' â— _])); - [iIntros (loop)|iIntros "/= !#"; iIntros (loop arg); inv_vec arg]; + [iIntros (loop)|iIntros "/= !>"; iIntros (loop arg); inv_vec arg]; simpl_subst. { iApply type_jump; solve_typing. } - iIntros (tid) "#LFT #HE Hna HL Hk HT". + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. simpl. - iDestruct "Hx'" as (ν q γs β) "(#Hx' & #Hαβ & Hν & Hâ—¯ & #H†& Hinv)". + iDestruct "Hx'" as (ν q γs β tid_own) "(#Hx' & #Hαβ & Hν & Hâ—¯ & #H†& Hinv)". iDestruct "Hinv" as (γ tyO tyS) "(R & #EqO & #EqS & R')". iDestruct "R'" as (tR vR) "R'". iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. @@ -159,7 +160,7 @@ Section rwlockreadguard_functions. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //=; simpl. iFrame "Hx". iExists _,_,_,_. iFrame "Hx' Hαβ Hν Hâ—¯ H†". - iExists _,_,_. iSplitL "R''". + iExists _,_,_,_. iSplitL "R''". - iExists _,_; by iDestruct "R''" as "[$ ?]". - iFrame "EqO EqS". iExists _,_. by iFrame "R'". } iApply type_jump; solve_typing. diff --git a/theories/typing/lib/rwlock/rwlockwriteguard.v b/theories/typing/lib/rwlock/rwlockwriteguard.v index 6cfd2e84803412f84fcbb7cb3eb66960afd5b035..dac7663d6e8bf8f720fd97ef9c9f2267a93b475f 100644 --- a/theories/typing/lib/rwlock/rwlockwriteguard.v +++ b/theories/typing/lib/rwlock/rwlockwriteguard.v @@ -22,10 +22,10 @@ Section rwlockwriteguard. ty_own tid vl := match vl return _ with | [ #(LitLoc l) ] => - ∃ γs β, &{β}((l +â‚— 1) ↦∗: ty.(ty_own) tid) ∗ + ∃ γs β tid_shr, &{β}((l +â‚— 1) ↦∗: ty.(ty_own) tid) ∗ α ⊑ β ∗ rwown γs writing_st ∗ (∃ γ tyO tyS, (∃ t, GPS_SWWriter l t () #(-1) γ) ∗ - rwlock_proto l γ γs β tyO tyS tid ty) + rwlock_proto l γ γs β tyO tyS tid tid_shr ty) | _ => False end; ty_shr κ tid l := @@ -43,6 +43,7 @@ Section rwlockwriteguard. try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". iMod (bor_exists with "LFT Hb") as (γs) "Hb". done. iMod (bor_exists with "LFT Hb") as (β) "Hb". done. + iMod (bor_exists with "LFT Hb") as (tid_shr) "Hb". done. iMod (bor_sep with "LFT Hb") as "[Hb H]". done. iMod (bor_sep with "LFT H") as "[Hαβ _]". done. iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ $]". done. @@ -55,7 +56,7 @@ Section rwlockwriteguard. iIntros (??????) "#? H". iDestruct "H" as (l') "[#Hf #H]". iExists _. iSplit. - by iApply frac_bor_shorten. - - iIntros "!# * % Htok". + - iIntros "!> * % Htok". iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. { iApply lft_intersect_mono. iApply lft_incl_refl. done. } iMod ("H" with "[] Htok") as "Hshr". done. iModIntro. iNext. @@ -78,28 +79,30 @@ Section rwlockwriteguard. Global Instance rwlockwriteguard_mono E L : Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockwriteguard. Proof. - intros α1 α2 Hα ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold. iIntros (Hty' q) "HL". + intros α1 α2 Hα ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold. + iIntros (Hty' qmax qL) "HL". iDestruct (Hty' with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". iDestruct (rwlock_proto_proper with "HL") as "#Hty1ty2"; first done. - iIntros "!# #HE". iDestruct ("Hα" with "HE") as "Hα1α2". + iIntros "!> #HE". iDestruct ("Hα" with "HE") as %Hα1α2. iDestruct ("Hty" with "HE") as "(%&#Ho&#Hs)". iSplit; [done|iSplit; iModIntro]. - iIntros (tid [|[[]|][]]) "H"; try done. - iDestruct "H" as (γs β) "(Hb & #H⊑ & Hown & Hinv)". - iExists γs, β. iFrame "Hown". iSplitL "Hb"; last iSplitR "Hinv". + iDestruct "H" as (γs β tid_shr) "(Hb & #H⊑ & Hown & Hinv)". + iExists γs, β, tid_shr. iFrame "Hown". iSplitL "Hb"; last iSplitR "Hinv". + iApply bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]"; iExists vl; iFrame; by iApply "Ho". - + by iApply lft_incl_trans. + + iApply lft_incl_trans; first by iApply lft_incl_syn_sem. done. + iDestruct "Hinv" as (γ tyO tyS) "(W & inv)". iExists γ, tyO, tyS. iFrame "W". by iApply ("Hty1ty2" with "HE"). - iIntros (κ tid l) "H". iDestruct "H" as (l') "H". iExists l'. - iDestruct "H" as "[$ #H]". iIntros "!# * % Htok". + iDestruct "H" as "[$ #H]". iIntros "!> * % Htok". iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. done. iApply lft_incl_refl. } - iMod ("H" with "[] Htok") as "Hshr". done. iModIntro. iNext. + { 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. done. iApply lft_incl_refl. - by iApply "Hs". + iApply ty_shr_mono; try done. + + iApply lft_intersect_mono; first by iApply lft_incl_syn_sem. iApply lft_incl_refl. + + by iApply "Hs". Qed. Global Instance rwlockwriteguard_mono_flip E L : Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) rwlockwriteguard. @@ -125,8 +128,19 @@ Section rwlockwriteguard. Qed. (* POSIX requires the unlock to occur from the thread that acquired - the lock, so Rust does not implement Send for RwLockWriteGuard. We could - prove this. *) + the lock, so Rust does not implement Send for RwLockWriteGuard. We can + prove this for our spinlock implementation, however. *) + Global Instance rwlockwriteguard_send α ty : + Send ty → Send (rwlockwriteguard α ty). + Proof. + iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H". + iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_entails. + repeat lazymatch goal with + | |- (ty_own _ _ _) ≡ (ty_own _ _ _) => by apply send_change_tid' + | |- (rwlock_proto _ _ _ _ _ _ _ _ _) ≡ _ => by apply rwlock_proto_change_tid_own + | |- _ => f_equiv + end. + Qed. End rwlockwriteguard. -Hint Resolve rwlockwriteguard_mono' rwlockwriteguard_proper' : lrust_typing. +Global Hint Resolve rwlockwriteguard_mono' rwlockwriteguard_proper' : lrust_typing. diff --git a/theories/typing/lib/rwlock/rwlockwriteguard_code.v b/theories/typing/lib/rwlock/rwlockwriteguard_code.v index d9ad06e62595656d63d25315d5970f2084d3802e..bd17c7904812187258fbf4a0aa7d31aad4beac0e 100644 --- a/theories/typing/lib/rwlock/rwlockwriteguard_code.v +++ b/theories/typing/lib/rwlock/rwlockwriteguard_code.v @@ -23,10 +23,10 @@ Section rwlockwriteguard_functions. typed_val rwlockwriteguard_deref (fn(∀ '(α, β), ∅; &shr{α}(rwlockwriteguard β ty)) → &shr{α} ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. iDestruct "Hx'" as (l') "#[Hfrac Hshr]". @@ -34,22 +34,23 @@ Section rwlockwriteguard_functions. [solve_typing..|]. iMod (frac_bor_acc with "LFT Hfrac Hα1") as (qlx') "[H↦ Hcloseα1]". done. rewrite heap_mapsto_vec_singleton. - iMod (lctx_lft_alive_tok β with "HE HL") as (qβ) "(Hβ & HL & Hclose')"; + iMod (lctx_lft_alive_tok_noend β with "HE HL") as (qβ) "(Hβ & HL & Hclose')"; [solve_typing..|]. iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hcloseβα2]". wp_bind (!(LitV lx'))%E. iApply (wp_step_fupd with "[Hα2β]"); - [done| |by iApply ("Hshr" with "[] Hα2β")|]; first done. + [| |by iApply ("Hshr" with "[] Hα2β")|]; [done..|]. iMod "H↦" as "[H↦1 H↦2]". wp_read. iIntros "[#Hshr' Hα2β]!>". wp_op. wp_let. iDestruct ("Hcloseβα2" with "Hα2β") as "[Hβ Hα2]". iMod ("Hcloseα1" with "[$H↦1 $H↦2]") as "Hα1". iMod ("Hclose'" with "Hβ HL") as "HL". iMod ("Hclose" with "[$] HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|]. iApply (type_type _ _ _ [ x â— box (&shr{α}(rwlockwriteguard β ty)); #(l' +â‚— 1) â— &shr{α}ty] with "[] LFT HE Hna HL Hk"); last first. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - iFrame. iApply (ty_shr_mono with "[] Hshr'"). iApply lft_incl_glb; first done. - by iApply lft_incl_refl. } + iFrame. iApply (ty_shr_mono with "[] Hshr'"). + iApply lft_incl_glb; first by iApply lft_incl_syn_sem. + by iApply lft_incl_refl. } iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. @@ -67,10 +68,10 @@ Section rwlockwriteguard_functions. typed_val rwlockwriteguard_derefmut (fn(∀ '(α, β), ∅; &uniq{α}(rwlockwriteguard β ty)) → &uniq{α}ty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros ([α β] Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". simpl_subst. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. iMod (bor_exists with "LFT Hx'") as (vl) "H". done. @@ -81,6 +82,7 @@ Section rwlockwriteguard_functions. rewrite heap_mapsto_vec_singleton. iMod (bor_exists with "LFT H") as (γ) "H". done. iMod (bor_exists with "LFT H") as (δ) "H". done. + iMod (bor_exists with "LFT H") as (tid_shr) "H". done. iMod (bor_sep with "LFT H") as "[Hb H]". done. iMod (bor_sep with "LFT H") as "[Hβδ _]". done. iMod (bor_persistent with "LFT Hβδ Hα") as "[#Hβδ Hα]". done. @@ -88,12 +90,13 @@ Section rwlockwriteguard_functions. wp_bind (!(LitV lx'))%E. iMod (bor_unnest with "LFT Hb") as "Hb"; first done. wp_read. wp_op. wp_let. iMod "Hb". iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]". iMod ("Hclose" with "Hα HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. + iDestruct (lctx_lft_incl_incl α β with "HL HE") as %?; [solve_typing..|]. iApply (type_type _ _ _ [ x â— box (uninit 1); #(l +â‚— 1) â— &uniq{α}ty] with "[] LFT HE Hna HL Hk"); last first. { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. iFrame. iApply (bor_shorten with "[] Hb"). iApply lft_incl_glb. - by iApply lft_incl_trans. by iApply lft_incl_refl. } + iApply lft_incl_trans; last done. by iApply lft_incl_syn_sem. + by iApply lft_incl_refl. } iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|]. iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. @@ -111,40 +114,34 @@ Section rwlockwriteguard_functions. typed_val rwlockwriteguard_drop (fn(∀ α, ∅; rwlockwriteguard α ty) → unit). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret arg). inv_vec arg=>x. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk HT". + iIntros (tid qmax) "#LFT #HE Hna HL Hk HT". rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. simpl. - iDestruct "Hx'" as (γs β) "(Hx' & #Hαβ & wst & Hinv)". + iDestruct "Hx'" as (γs β tid_own) "(Hx' & #Hαβ & wst & Hinv)". iDestruct "Hinv" as (γ tyO tyS) "(W & #EqO & #EqS & #R)". iDestruct "W" as (t) "W". iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. iMod (lft_incl_acc with "Hαβ Hα") as (qβ) "[Hβ Hcloseα]". done. wp_bind (#lx' <-ʳᵉˡ #0)%E. - iAssert (â–· â–¡ (∀ κ q' E', ⌜lftE ⊆ E'⌠-∗ &{κ} tyO -∗ - (q').[κ] ={E'}=∗ (â–¡ tyS κ) ∗ (q').[κ]))%I as "#Share". - { iIntros "!> !#" (κ???) "tyO tok". - iMod (ty_share with "LFT [tyO] tok") as "[#tyS $]" ;[done|..]. - - iApply (bor_iff with "EqO tyO"). - - iIntros "!> !#". by iApply "EqS". } iApply (GPS_aSP_SWWrite_rel rwlockN (rwlock_interp γs β tyO tyS) rwlock_noCAS _ β qβ (λ _, True)%I True%I (rwlock_interp γs β tyO tyS true lx' γ t () #(-1)) t () () #(-1) #0 with "[$LFT $Hβ $W wst Hx']"); [solve_ndisj|done..|iSplitL ""; last iSplitR ""|]. { iDestruct "R" as (??) "[R $]". } - { iIntros "!>" (t' Ext') "W". iDestruct 1 as (st' Eqv) "[_ [g INT]]". + { iIntros "!>" (t' Ext') "W". iDestruct 1 as (st' Eqv) "[#Share [g INT]]". iDestruct (rwown_global_writing_st with "g wst") as %?. subst st'. iMod (rwown_update_write_dealloc with "g wst") as "g". iModIntro. iSplitL ""; [by iIntros "!> _"|]. iDestruct (GPS_SWWriter_shared with "W") as "[W Rs]". iSplitR ""; [|done]. iExists None. iSplitL ""; [done|]. iFrame "Share g W Rs". iApply (bor_iff with "[] Hx'"). - iIntros "!> !#". iApply bi.iff_sym. by iApply "EqO". } + iIntros "!> !>". iApply bi.iff_sym. by iApply "EqO". } { by iIntros "!> !> $". } iIntros "!>" (?) "(_ & R' & Hβ & _)". wp_seq. iMod ("Hcloseα" with "Hβ") as "Hα". diff --git a/theories/typing/lib/spawn.v b/theories/typing/lib/spawn.v index 2da54c522e52cc2241f187d00da96e1419e850c4..bf928a57a36fcf0cc4ad0b5db4b09ac9b8ad1e28 100644 --- a/theories/typing/lib/spawn.v +++ b/theories/typing/lib/spawn.v @@ -33,7 +33,7 @@ Section join_handle. iIntros "#Hincl". iSplit; first done. iSplit; iModIntro. - iIntros "* Hvl". destruct vl as [|[[|vl|]|] [|]]; try done. simpl. iDestruct "Hvl" as (γc γe) "Hvl". iExists γc, γe. - iApply (join_handle_impl with "[] Hvl"). iIntros "!# * Hown" (tid'). + iApply (join_handle_impl with "[] Hvl"). iIntros "!> * Hown" (tid'). iDestruct (box_type_incl with "Hincl") as "{Hincl} (_ & Hincl & _)". iApply "Hincl". done. - iIntros "* _". auto. @@ -42,8 +42,8 @@ Section join_handle. Global Instance join_handle_mono E L : Proper (subtype E L ==> subtype E L) join_handle. Proof. - iIntros (ty1 ty2 Hsub ?) "HL". iDestruct (Hsub with "HL") as "#Hsub". - iIntros "!# #HE". iApply join_handle_subtype. iApply "Hsub"; done. + iIntros (ty1 ty2 Hsub ??) "HL". iDestruct (Hsub with "HL") as "#Hsub". + iIntros "!> #HE". iApply join_handle_subtype. iApply "Hsub"; done. Qed. Global Instance join_handle_proper E L : Proper (eqtype E L ==> eqtype E L) join_handle. @@ -84,13 +84,13 @@ Section spawn. let E Ï := ty_outlives_E fty static ++ ty_outlives_E retty static in typed_val (spawn call_once) (fn(E; fty) → join_handle retty). Proof. - intros Hf ? E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros Hf ? E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (_ Ï ret arg). inv_vec arg=>env. simpl_subst. iApply type_let; [apply Hf|solve_typing|]. iIntros (f'). simpl_subst. iApply (type_let _ _ _ _ ([f' â— _; env â— _]) (λ j, [j â— join_handle retty])); try solve_typing; [|]. { (* The core of the proof: showing that spawn is safe. *) - iIntros (tid) "#LFT #HE $ $ [Hf' [Henv _]]". rewrite !tctx_hasty_val [fn _]lock. + iIntros (tid qmax) "#LFT #HE $ $ [Hf' [Henv _]]". rewrite !tctx_hasty_val [fn _]lock. iApply (spawn_spec _ (join_inv retty) with "[-]"); last first; last simpl_subst. { iIntros "!> *". rewrite tctx_interp_singleton tctx_hasty_val. @@ -119,12 +119,12 @@ Section spawn. Lemma join_type retty `{!TyWf retty} : typed_val join (fn(∅; join_handle retty) → retty). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (_ Ï ret arg). inv_vec arg=>c. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (c'); simpl_subst. iApply (type_let _ _ _ _ ([c' â— _]) (λ r, [r â— box retty])); try solve_typing; [|]. - { iIntros (tid) "#LFT _ $ $". + { iIntros (tid qmax) "#LFT _ $ $". rewrite tctx_interp_singleton tctx_hasty_val. iIntros "Hc". destruct c' as [[|c'|]|]; try done. iDestruct "Hc" as (??) "Hc". iApply (join_spec with "Hc"); [solve_ndisj..|]. iNext. iIntros "* Hret". diff --git a/theories/typing/lib/swap.v b/theories/typing/lib/swap.v index d5fc5b1557cb9906e95cb5ad7b16ef016ccab850..cb46baaaa629fc1af4cbbdd6887146bcd0125ed8 100644 --- a/theories/typing/lib/swap.v +++ b/theories/typing/lib/swap.v @@ -16,11 +16,11 @@ Section swap. Lemma swap_type Ï„ `{!TyWf Ï„} : typed_val (swap Ï„) (fn(∀ α, ∅; &uniq{α} Ï„, &uniq{α} Ï„) → unit). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret p). + intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α Ï ret p). inv_vec p=>p1 p2. simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (p1'). simpl_subst. iApply type_deref; [solve_typing..|]. iIntros (p2'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (H2 & H2' & H1 & H1' & _)". + iIntros (tid qmax) "#LFT #HE Hna HL Hk (H2 & H2' & H1 & H1' & _)". rewrite !tctx_hasty_val. iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "([Hα1 Hα2] & HL & Hclose)"; [solve_typing..|]. diff --git a/theories/typing/lib/take_mut.v b/theories/typing/lib/take_mut.v index 2c43e944140d08b6f67c3f88631f73f685006995..b8c06c9b21200b2811c3e012e0d2cacc3ef9f463 100644 --- a/theories/typing/lib/take_mut.v +++ b/theories/typing/lib/take_mut.v @@ -22,18 +22,18 @@ Section code. typed_val call_once (fn(∅; fty, ty) → ty) → typed_val (take ty call_once) (fn(∀ α, ∅; &uniq{α} ty, fty) → unit). Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α Ï ret arg). - inv_vec arg=>x env. simpl_subst. + intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". + iIntros (α Ï ret arg). inv_vec arg=>x env. simpl_subst. iApply type_deref; [solve_typing..|]; iIntros (x'); simpl_subst. iApply type_let; [apply Hf|solve_typing|]; iIntros (f'); simpl_subst. iApply (type_new ty.(ty_size)); [solve_typing..|]; iIntros (t); simpl_subst. (* Switch to Iris. *) - iIntros (tid) "#LFT #HE Hna HL Hk (Ht & Hf' & Hx & Hx' & Henv & _)". + iIntros (tid qmax) "#LFT #HE Hna HL Hk (Ht & Hf' & Hx & Hx' & Henv & _)". rewrite !tctx_hasty_val [[x]]lock [[env]]lock [fn _]lock. iDestruct (ownptr_uninit_own with "Ht") as (tl tvl) "(% & >Htl & Htl†)". subst t. simpl. destruct x' as [[|lx'|]|]; try done. simpl. iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose1)"; [solve_typing..|]. - iMod (lctx_lft_alive_tok Ï with "HE HL") as (qÏ) "(HÏ & HL & Hclose2)"; [solve_typing..|]. + iMod (lctx_lft_alive_tok_noend Ï with "HE HL") as (qÏ) "(HÏ & HL & Hclose2)"; [solve_typing..|]. iMod (bor_acc with "LFT Hx' Hα") as "[Hx' Hclose3]"; first done. iDestruct (heap_mapsto_ty_own with "Hx'") as (x'vl) "[>Hx'↦ Hx'vl]". wp_apply (wp_memcpy with "[$Htl $Hx'↦]"); [by auto using vec_to_list_length..|]. diff --git a/theories/typing/own.v b/theories/typing/own.v index 7d53c63d82509dc2b384cb46dc4d7a833402a664..edc3e8536ca6091e1d86f60238145078498432d9 100644 --- a/theories/typing/own.v +++ b/theories/typing/own.v @@ -79,7 +79,7 @@ Section own. Next Obligation. intros _ ty κ κ' tid l. iIntros "#Hκ #H". iDestruct "H" as (l') "[Hfb #Hvs]". - iExists l'. iSplit. by iApply (frac_bor_shorten with "[]"). iIntros "!# *% Htok". + iExists l'. iSplit. by iApply (frac_bor_shorten with "[]"). iIntros "!> *% Htok". iApply (step_fupd_mask_mono F _ (F∖↑shrN)); [solve_ndisj..|]. iMod (lft_incl_acc with "Hκ Htok") as (q') "[Htok Hclose]"; first solve_ndisj. iMod ("Hvs" with "[%] Htok") as "Hvs'"; first solve_ndisj. iModIntro. iNext. @@ -99,7 +99,7 @@ Section own. iDestruct "Heq" as %->. iFrame. iApply (heap_mapsto_pred_wand with "Hmt"). iApply "Ho". - iIntros (???) "H". iDestruct "H" as (l') "[Hfb #Hvs]". - iExists l'. iFrame. iIntros "!#". iIntros (F' q) "% Htok". + iExists l'. iFrame. iIntros "!>". iIntros (F' q) "% Htok". iMod ("Hvs" with "[%] Htok") as "Hvs'". done. iModIntro. iNext. iMod "Hvs'" as "[Hshr $]". iApply ("Hs" with "Hshr"). Qed. @@ -107,9 +107,9 @@ Section own. Global Instance own_mono E L n : Proper (subtype E L ==> subtype E L) (own_ptr n). Proof. - intros ty1 ty2 Hincl. iIntros (qL) "HL". + intros ty1 ty2 Hincl. iIntros (qmax qL) "HL". iDestruct (Hincl with "HL") as "#Hincl". - iClear "∗". iIntros "!# #HE". + iClear "∗". iIntros "!> #HE". iApply own_type_incl; first by auto. iApply "Hincl"; auto. Qed. Lemma own_mono' E L n1 n2 ty1 ty2 : @@ -161,9 +161,9 @@ Section box. Global Instance box_mono E L : Proper (subtype E L ==> subtype E L) box. Proof. - intros ty1 ty2 Hincl. iIntros (qL) "HL". + intros ty1 ty2 Hincl. iIntros (qmax qL) "HL". iDestruct (Hincl with "HL") as "#Hincl". - iClear "∗". iIntros "!# #HE". + iClear "∗". iIntros "!> #HE". iApply box_type_incl. iApply "Hincl"; auto. Qed. Lemma box_mono' E L ty1 ty2 : @@ -221,8 +221,8 @@ Section typing. Lemma write_own {E L} ty ty' n : ty.(ty_size) = ty'.(ty_size) → ⊢ typed_write E L (own_ptr n ty') ty (own_ptr n ty). Proof. - rewrite typed_write_eq. iIntros (Hsz) "!#". - iIntros ([[]|] tid F qL ?) "_ _ $ Hown"; try done. + rewrite typed_write_eq. iIntros (Hsz) "!>". + iIntros ([[]|] tid F qmax qL ?) "_ _ $ Hown"; try done. rewrite /= Hsz. iDestruct "Hown" as "[H↦ $]". iDestruct "H↦" as (vl) "[>H↦ Hown]". iDestruct (ty_size_eq with "Hown") as "#>%". iExists _, _. iFrame "H↦". auto. Qed. @@ -230,8 +230,8 @@ Section typing. Lemma read_own_copy E L ty n : Copy ty → ⊢ typed_read E L (own_ptr n ty) ty (own_ptr n ty). Proof. - rewrite typed_read_eq. iIntros (Hsz) "!#". - iIntros ([[]|] tid F qL ?) "_ _ $ $ Hown"; try done. + rewrite typed_read_eq. iIntros (Hsz) "!>". + iIntros ([[]|] tid F qmax qL ?) "_ _ $ $ Hown"; try done. iDestruct "Hown" as "[H↦ H†]". iDestruct "H↦" as (vl) "[>H↦ #Hown]". iExists l, _, _. iFrame "∗#". iSplitR; first done. iIntros "!> Hl !>". iExists _. auto. @@ -241,7 +241,7 @@ Section typing. ⊢ typed_read E L (own_ptr n ty) ty (own_ptr n $ uninit ty.(ty_size)). Proof. rewrite typed_read_eq. iModIntro. - iIntros ([[]|] tid F qL ?) "_ _ $ $ Hown"; try done. + iIntros ([[]|] tid F qmax qL ?) "_ _ $ $ Hown"; try done. iDestruct "Hown" as "[H↦ H†]". iDestruct "H↦" as (vl) "[>H↦ Hown]". iDestruct (ty_size_eq with "Hown") as "#>%". iExists l, vl, _. iFrame "∗#". iSplitR; first done. iIntros "!> Hl !> !>". @@ -253,7 +253,7 @@ Section typing. ⊢ let n' := Z.to_nat n in typed_instruction_ty E L [] (new [ #n ]%E) (own_ptr n' (uninit n')). Proof. - iIntros (? tid) "#LFT #HE $ $ _". + iIntros (? tid qmax) "#LFT #HE $ $ _". iApply wp_new; try done. iModIntro. iIntros (l) "(H†& Hlft)". rewrite tctx_interp_singleton tctx_hasty_val. iNext. rewrite freeable_sz_full Z2Nat.id //. iFrame. @@ -286,7 +286,7 @@ Section typing. Z.of_nat (ty.(ty_size)) = n → ⊢ typed_instruction E L [p â— own_ptr ty.(ty_size) ty] (delete [ #n; p])%E (λ _, []). Proof. - iIntros (<- tid) "#LFT #HE $ $ Hp". rewrite tctx_interp_singleton. + iIntros (<- tid qmax) "#LFT #HE $ $ Hp". rewrite tctx_interp_singleton. wp_bind p. iApply (wp_hasty with "Hp"); [done|]. iIntros ([[]|]) "_ Hown"; try done. iDestruct "Hown" as "[H↦: >H†]". iDestruct "H↦:" as (vl) "[>H↦ Hown]". iDestruct (ty_size_eq with "Hown") as "#>EQ". @@ -363,8 +363,8 @@ Section typing. Qed. End typing. -Hint Resolve own_mono' own_proper' box_mono' box_proper' +Global Hint Resolve own_mono' own_proper' box_mono' box_proper' write_own read_own_copy : lrust_typing. (* By setting the priority high, we make sure copying is tried before moving. *) -Hint Resolve read_own_move | 100 : lrust_typing. +Global Hint Resolve read_own_move | 100 : lrust_typing. diff --git a/theories/typing/product.v b/theories/typing/product.v index 3bc52cac881b9db479202bde8509e2b45e4d3f02..cde5ec3ba8adee1a01525cab3e699910e19ddbe6 100644 --- a/theories/typing/product.v +++ b/theories/typing/product.v @@ -89,9 +89,9 @@ Section product. Global Instance product2_mono E L: Proper (subtype E L ==> subtype E L ==> subtype E L) product2. Proof. - iIntros (ty11 ty12 H1 ty21 ty22 H2). iIntros (qL) "HL". + iIntros (ty11 ty12 H1 ty21 ty22 H2). iIntros (qmax qL) "HL". iDestruct (H1 with "HL") as "#H1". iDestruct (H2 with "HL") as "#H2". - iClear "∗". iIntros "!# #HE". + iClear "∗". iIntros "!> #HE". iDestruct ("H1" with "HE") as "#(% & #Ho1 & #Hs1)". clear H1. iDestruct ("H2" with "HE") as "#(% & #Ho2 & #Hs2)". clear H2. iSplit; first by (iPureIntro; simpl; f_equal). iSplit; iModIntro. @@ -198,8 +198,8 @@ Section typing. Global Instance prod2_assoc E L : Assoc (eqtype E L) product2. Proof. - intros ???. apply eqtype_unfold. iIntros (?) "_ !# _". - iSplit; first by rewrite /= assoc. iSplit; iIntros "!# *"; iSplit; iIntros "H". + intros ???. apply eqtype_unfold. iIntros (??) "_ !> _". + iSplit; first by rewrite /= assoc. iSplit; iIntros "!> *"; iSplit; iIntros "H". - iDestruct "H" as (vl1 vl') "(% & Ho1 & H)". iDestruct "H" as (vl2 vl3) "(% & Ho2 & Ho3)". subst. iExists _, _. iSplit. by rewrite assoc. iFrame. iExists _, _. by iFrame. @@ -214,8 +214,8 @@ Section typing. Global Instance prod2_unit_leftid E L : LeftId (eqtype E L) unit product2. Proof. - intros ty. apply eqtype_unfold. iIntros (?) "_ !# _". iSplit; first done. - iSplit; iIntros "!# *"; iSplit; iIntros "H". + intros ty. apply eqtype_unfold. iIntros (??) "_ !> _". iSplit; first done. + iSplit; iIntros "!> *"; iSplit; iIntros "H". - iDestruct "H" as (? ?) "(% & % & ?)". by subst. - iExists [], _. eauto. - iDestruct "H" as "(? & ?)". rewrite shift_loc_0. @@ -226,8 +226,8 @@ Section typing. Global Instance prod2_unit_rightid E L : RightId (eqtype E L) unit product2. Proof. - intros ty. apply eqtype_unfold. iIntros (?) "_ !# _". - iSplit; first by rewrite /= -plus_n_O. iSplit; iIntros "!# *"; iSplit; iIntros "H". + intros ty. apply eqtype_unfold. iIntros (??) "_ !> _". + iSplit; first by rewrite /= -plus_n_O. iSplit; iIntros "!> *"; iSplit; iIntros "H". - iDestruct "H" as (? ?) "(% & ? & %)". subst. by rewrite app_nil_r. - iExists _, []. rewrite app_nil_r. eauto. - iDestruct "H" as "(? & _)". done. @@ -261,6 +261,6 @@ Section typing. End typing. Arguments product : simpl never. -Hint Opaque product : lrust_typing lrust_typing_merge. -Hint Resolve product_mono' product_proper' ty_outlives_E_elctx_sat_product +Global Hint Opaque product : lrust_typing lrust_typing_merge. +Global Hint Resolve product_mono' product_proper' ty_outlives_E_elctx_sat_product : lrust_typing. diff --git a/theories/typing/product_split.v b/theories/typing/product_split.v index b73d1152eff9297a111b3f0bd6ec1b1a9e634f25..eb5075dfbd6643d31c237fd2f853d4a25ab12f50 100644 --- a/theories/typing/product_split.v +++ b/theories/typing/product_split.v @@ -36,12 +36,12 @@ Section product_split. (∀ tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]âŒ) → ∀ p, tctx_incl E L [p â— ptr $ product tyl] (hasty_ptr_offsets p ptr tyl 0). Proof. - iIntros (Hsplit Hloc p tid q) "#LFT #HE HL H". + iIntros (Hsplit Hloc p tid qmax qL) "#LFT #HE HL H". iInduction tyl as [|ty tyl IH] "IH" forall (p). { rewrite tctx_interp_nil. auto. } rewrite product_cons. iMod (Hsplit with "LFT HE HL H") as "(HL & H)". cbn -[tctx_elt_interp]. - iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". iDestruct "Hp" as %Hp. + iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". iDestruct "Hp" as %Hp. iDestruct (Hloc with "Hty") as %[l [=->]]. iAssert (tctx_elt_interp tid (p +â‚— #0 â— ptr ty)) with "[Hty]" as "$". { iExists #l. iSplit; last done. simpl; by rewrite Hp shift_loc_0. } @@ -58,7 +58,7 @@ Section product_split. (∀ tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]âŒ) → ∀ p, tctx_incl E L (hasty_ptr_offsets p ptr tyl 0) [p â— ptr $ product tyl]. Proof. - iIntros (Hptr Htyl Hmerge Hloc p tid q) "#LFT #HE HL H". + iIntros (Hptr Htyl Hmerge Hloc p tid qmax qL) "#LFT #HE HL H". iInduction tyl as [|ty tyl IH] "IH" forall (p Htyl); first done. rewrite product_cons. rewrite /= tctx_interp_singleton tctx_interp_cons. iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". @@ -86,7 +86,7 @@ Section product_split. tctx_incl E L [p â— own_ptr n $ product2 ty1 ty2] [p â— own_ptr n ty1; p +â‚— #ty1.(ty_size) â— own_ptr n ty2]. Proof. - iIntros (tid q) "#LFT _ $ H". + iIntros (tid qmax qL) "#LFT _ $ H". rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as ([[]|]) "[#Hp H]"; try done. iDestruct "H" as "[H >H†]". iDestruct "H" as (vl) "[>H↦ H]". @@ -104,7 +104,7 @@ Section product_split. tctx_incl E L [p â— own_ptr n ty1; p +â‚— #ty1.(ty_size) â— own_ptr n ty2] [p â— own_ptr n $ product2 ty1 ty2]. Proof. - iIntros (tid q) "#LFT _ $ H". + iIntros (tid qmax qL) "#LFT _ $ H". rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[]|]) "(Hp1 & H1)"; try done. iDestruct "H1" as "(H↦1 & H†1)". @@ -141,7 +141,7 @@ Section product_split. tctx_incl E L [p â— &uniq{κ}(product2 ty1 ty2)] [p â— &uniq{κ} ty1; p +â‚— #ty1.(ty_size) â— &uniq{κ} ty2]. Proof. - iIntros (tid q) "#LFT _ $ H". + iIntros (tid qmax qL) "#LFT _ $ H". rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as ([[]|]) "[Hp H]"; try done. iDestruct "Hp" as %Hp. rewrite /= split_prod_mt. iMod (bor_sep with "LFT H") as "[H1 H2]"; first solve_ndisj. @@ -153,7 +153,7 @@ Section product_split. tctx_incl E L [p â— &uniq{κ} ty1; p +â‚— #ty1.(ty_size) â— &uniq{κ} ty2] [p â— &uniq{κ}(product2 ty1 ty2)]. Proof. - iIntros (tid q) "#LFT _ $ H". + iIntros (tid qmax qL) "#LFT _ $ H". rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[]|]) "[Hp1 H1]"; try done. iDestruct "Hp1" as %Hp1. iDestruct "H2" as (v2) "(Hp2 & H2)". rewrite /= Hp1. @@ -191,7 +191,7 @@ Section product_split. tctx_incl E L [p â— &shr{κ}(product2 ty1 ty2)] [p â— &shr{κ} ty1; p +â‚— #ty1.(ty_size) â— &shr{κ} ty2]. Proof. - iIntros (tid q) "#LFT _ $ H". + iIntros (tid qmax qL) "#LFT _ $ H". rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as ([[]|]) "[Hp H]"; try iDestruct "H" as "[]". iDestruct "H" as "[H1 H2]". iDestruct "Hp" as %Hp. @@ -202,7 +202,7 @@ Section product_split. tctx_incl E L [p â— &shr{κ} ty1; p +â‚— #ty1.(ty_size) â— &shr{κ} ty2] [p â— &shr{κ}(product2 ty1 ty2)]. Proof. - iIntros (tid q) "#LFT _ $ H". + iIntros (tid qmax qL) "#LFT _ $ H". rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[]|]) "[Hp1 Hown1]"; try done. iDestruct "Hp1" as %Hp1. iDestruct "H2" as ([[]|]) "[Hp2 Hown2]"; try done. @@ -303,11 +303,11 @@ End product_split. (* We do not want unification to try to unify the definition of these types with anything in order to try splitting or merging. *) -Hint Opaque own_ptr uniq_bor shr_bor tctx_extract_hasty : lrust_typing lrust_typing_merge. +Global Hint Opaque own_ptr uniq_bor shr_bor tctx_extract_hasty : lrust_typing lrust_typing_merge. (* We make sure that splitting is tried before borrowing, so that not the entire product is borrowed when only a part is needed. *) -Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extract_split_shr_prod +Global Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extract_split_shr_prod | 5 : lrust_typing. (* Merging is also tried after everything, except @@ -320,6 +320,6 @@ Hint Resolve tctx_extract_split_own_prod tctx_extract_split_uniq_prod tctx_extra solve_typing get slow because of that. See: https://coq.inria.fr/bugs/show_bug.cgi?id=5304 *) -Hint Resolve tctx_extract_merge_own_prod tctx_extract_merge_uniq_prod tctx_extract_merge_shr_prod +Global Hint Resolve tctx_extract_merge_own_prod tctx_extract_merge_uniq_prod tctx_extract_merge_shr_prod | 40 : lrust_typing_merge. -Hint Unfold extract_tyl : lrust_typing. +Global Hint Unfold extract_tyl : lrust_typing. diff --git a/theories/typing/programs.v b/theories/typing/programs.v index b7caa3674f2834773a47c28b914f37ad41b2d36b..faada6582cd2024af0d752b0d0f9eb0f59f227d5 100644 --- a/theories/typing/programs.v +++ b/theories/typing/programs.v @@ -12,8 +12,8 @@ Section typing. (* This is an iProp because it is also used by the function type. *) Definition typed_body (E : elctx) (L : llctx) (C : cctx) (T : tctx) (e : expr) : vPropI Σ := - (∀ tid, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ na_own tid ⊤ -∗ llctx_interp L 1 -∗ - cctx_interp tid C -∗ tctx_interp tid T -∗ + (∀ tid (qmax : Qp), ⎡lft_ctx⎤ -∗ elctx_interp E -∗ na_own tid ⊤ -∗ llctx_interp qmax L -∗ + cctx_interp tid qmax C -∗ tctx_interp tid T -∗ WP e in tid {{ _, cont_postcondition }})%I. Global Arguments typed_body _ _ _ _ _%E. @@ -36,12 +36,22 @@ Section typing. (typed_body E L). Proof. intros C1 C2 HC T1 T2 HT e ? <-. iIntros "H". - iIntros (tid) "#LFT #HE Htl HL HC HT". + iIntros (tid qmax) "#LFT #HE Htl HL HC HT". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (HT with "LFT HE HL HT") as "(HL & HT)". + iDestruct ("HLclose" with "HL") as "HL". iApply ("H" with "LFT HE Htl HL [HC] HT"). by iApply (HC with "LFT HE HC"). Qed. + Lemma typed_body_tctx_incl E L T2 T1 C e : + tctx_incl E L T1 T2 → + (⊢ typed_body E L C T2 e) → + ⊢ typed_body E L C T1 e. + Proof. + intros Hincl He2. iApply typed_body_mono; last done; done. + Qed. + Global Instance typed_body_mono_flip E L: Proper (cctx_incl E ==> tctx_incl E L ==> eq ==> flip (⊢)) (typed_body E L). @@ -50,19 +60,19 @@ Section typing. (** Instruction *) Definition typed_instruction (E : elctx) (L : llctx) (T1 : tctx) (e : expr) (T2 : val → tctx) : vPropI Σ := - (∀ tid, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ na_own tid ⊤ -∗ - llctx_interp L 1 -∗ tctx_interp tid T1 -∗ + (∀ tid qmax, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ na_own tid ⊤ -∗ + llctx_interp qmax L -∗ tctx_interp tid T1 -∗ WP e in tid {{ v, na_own tid ⊤ ∗ - llctx_interp L 1 ∗ tctx_interp tid (T2 v) }})%I. + llctx_interp qmax L ∗ tctx_interp tid (T2 v) }})%I. Global Arguments typed_instruction _ _ _ _%E _. (** Writing and Reading **) Definition typed_write_def (E : elctx) (L : llctx) (ty1 ty ty2 : type) : vPropI Σ := - (â–¡ ∀ v tid F qL, ⌜lftE ∪ (↑lrustN) ⊆ F⌠→ - ⎡lft_ctx⎤ -∗ elctx_interp E -∗ llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗ + (â–¡ ∀ v tid F qmax qL, ⌜lftE ∪ (↑lrustN) ⊆ F⌠→ + ⎡lft_ctx⎤ -∗ elctx_interp E -∗ llctx_interp_noend qmax L qL -∗ ty1.(ty_own) tid [v] ={F}=∗ ∃ (l : loc) vl, ⌜length vl = ty.(ty_size) ∧ v = #l⌠∗ l ↦∗ vl ∗ (â–· l ↦∗: ty.(ty_own) tid ={F}=∗ - llctx_interp L qL ∗ ty2.(ty_own) tid [v]))%I. + llctx_interp_noend qmax L qL ∗ ty2.(ty_own) tid [v]))%I. Definition typed_write_aux : seal (@typed_write_def). by eexists. Qed. Definition typed_write := typed_write_aux.(unseal). Definition typed_write_eq : @typed_write = @typed_write_def := typed_write_aux.(seal_eq). @@ -78,12 +88,12 @@ Section typing. that nobody could possibly have changed the vl (because only half the fraction was given). So we go with the definition that is easier to prove. *) Definition typed_read_def (E : elctx) (L : llctx) (ty1 ty ty2 : type) : vPropI Σ := - (â–¡ ∀ v tid F qL, ⌜lftE ∪ ↑lrustN ⊆ F⌠→ + (â–¡ ∀ v tid F qmax qL, ⌜lftE ∪ ↑lrustN ⊆ F⌠→ ⎡lft_ctx⎤ -∗ elctx_interp E -∗ na_own tid F -∗ - llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗ + llctx_interp_noend qmax L qL -∗ ty1.(ty_own) tid [v] ={F}=∗ ∃ (l : loc) vl q, ⌜v = #l⌠∗ l ↦∗{q} vl ∗ â–· ty.(ty_own) tid vl ∗ (l ↦∗{q} vl ={F}=∗ na_own tid F ∗ - llctx_interp L qL ∗ ty2.(ty_own) tid [v]))%I. + llctx_interp_noend qmax L qL ∗ ty2.(ty_own) tid [v]))%I. Definition typed_read_aux : seal (@typed_read_def). by eexists. Qed. Definition typed_read := typed_read_aux.(unseal). Definition typed_read_eq : @typed_read = @typed_read_def := typed_read_aux.(seal_eq). @@ -112,16 +122,40 @@ Section typing_rules. typed_body E L C T e -∗ typed_body E L C T e. Proof. done. Qed. - (* TODO: Proof a version of this that substitutes into a compatible context... - if we really want to do that. *) - Lemma type_equivalize_lft E L C T κ1 κ2 e : - (⊢ typed_body ((κ1 ⊑ₑ κ2) :: (κ2 ⊑ₑ κ1) :: E) L C T e) → - ⊢ typed_body E ((κ1 ⊑ₗ [κ2]) :: L) C T e. + (** This lemma can replace [κ1] by [κ2] and vice versa in positions that + respect "semantic lifetime equivalence"; in particular, lifetimes of + references can be adjusted this way. However, it cannot replace lifetimes in + other type constructors, as those might only respect *syntactic* lifetime + equivalence. This lemma is *weaker* than what was in the original paper where + lifetimes could be replaced everywhere; it had to be adjusted for GhostCell. + See [typing.lib.diverging_static] for an example of how + [type_equivalize_lft_static] without this restriction ciuld be used to subvert + branding. + + This is technically not a proper typing rule since the type system has no way + to express "subtyping wrt semantic lifetime inclusion". However, there is no + fundamental reason that we could not also reflect all these semantic facts on + the syntactic side, it would just be very clunky (and note that in Coq we do + not reflect this syntactic side anway). *) + Lemma type_equivalize_lft E L C T1 T2 κ1 κ2 e : + (∀ tid, ⎡lft_ctx⎤ -∗ κ1 ⊑ κ2 -∗ κ2 ⊑ κ1 -∗ tctx_interp tid T1 -∗ tctx_interp tid T2) → + (⊢ typed_body E L C T2 e) → + ⊢ typed_body E ((κ1 ⊑ₗ [κ2]) :: L) C T1 e. + Proof. + iIntros (Hswitch He tid qmax) "#LFT #HE Htl [Hκ HL] HC HT". + iMod (lctx_equalize_lft_sem with "LFT Hκ") as "[Hκ1 Hκ2]". + iApply (He with "LFT HE Htl HL HC [-]"). + iApply (Hswitch with "LFT Hκ1 Hκ2"). done. + Qed. + Lemma type_equivalize_lft_static E L C T1 T2 κ e : + (∀ tid, ⎡lft_ctx⎤ -∗ static ⊑ κ -∗ tctx_interp tid T1 -∗ tctx_interp tid T2) → + (⊢ typed_body E L C T2 e) → + ⊢ typed_body E ((κ ⊑ₗ []) :: L) C T1 e. Proof. - iIntros (He tid) "#LFT #HE Htl [Hκ HL] HC HT". - iMod (lctx_equalize_lft with "LFT Hκ") as "[Hκ1 Hκ2]". - iApply (He with "LFT [Hκ1 Hκ2 HE] Htl HL HC HT"). - rewrite /elctx_interp /=. by iFrame. + iIntros (Hswitch He tid qmax) "#LFT #HE Htl [Hκ HL] HC HT". + iMod (lctx_equalize_lft_sem_static with "LFT Hκ") as "Hκ". + iApply (He with "LFT HE Htl HL HC [-]"). + iApply (Hswitch with "LFT Hκ"). done. Qed. Lemma type_let' E L T1 T2 (T : tctx) C xb e e' : @@ -130,7 +164,7 @@ Section typing_rules. (∀ v : val, typed_body E L C (T2 v ++ T) (subst' xb v e')) -∗ typed_body E L C (T1 ++ T) (let: xb := e in e'). Proof. - iIntros (Hc) "He He'". iIntros (tid) "#LFT #HE Htl HL HC HT". rewrite tctx_interp_app. + iIntros (Hc) "He He'". iIntros (tid qmax) "#LFT #HE Htl HL HC HT". rewrite tctx_interp_app. iDestruct "HT" as "[HT1 HT]". wp_bind e. iApply (wp_wand with "[He HL HT1 Htl]"). { iApply ("He" with "LFT HE Htl HL HT1"). } iIntros (v) "/= (Htl & HL & HT2)". wp_let. @@ -167,12 +201,16 @@ Section typing_rules. (∀ κ, typed_body E ((κ ⊑ₗ κs) :: L) C T e) -∗ typed_body E L C T (Newlft ;; e). Proof. - iIntros (Hc) "He". iIntros (tid) "#LFT #HE Htl HL HC HT". + iIntros (Hc) "He". iIntros (tid qmax) "#LFT #HE Htl HL HC HT". iMod (lft_create with "LFT") as (Λ) "[Htk #Hinh]"; [done..|]. set (κ' := lft_intersect_list κs). wp_seq. iApply ("He" $! (κ' ⊓ Λ) with "LFT HE Htl [HL Htk] HC HT"). rewrite /llctx_interp /=. iFrame "HL". - iExists Λ. iSplit; first done. iFrame. iModIntro. done. + iExists Λ. iSplit; first done. + destruct (decide (1 ≤ qmax)%Qp) as [_|Hlt%Qp_lt_nge]. + - by iFrame "#∗". + - apply Qp_lt_sum in Hlt as [q' ->]. iDestruct "Htk" as "[$ Htk]". + iIntros "Htk'". iApply "Hinh". iFrame. Qed. (* TODO: It should be possible to show this while taking only one step. @@ -181,11 +219,11 @@ Section typing_rules. Closed [] e → UnblockTctx κ T1 T2 → typed_body E L C T2 e -∗ typed_body E ((κ ⊑ₗ κs) :: L) C T1 (Endlft ;; e). Proof. - iIntros (Hc Hub) "He". iIntros (tid) "#LFT #HE Htl [Hκ HL] HC HT". - iDestruct "Hκ" as (Λ) "(% & Htok & #Hend)". + iIntros (Hc Hub) "He". iIntros (tid qmax) "#LFT #HE Htl [Hκ HL] HC HT". + iDestruct "Hκ" as (Λ) "(% & Htok & Hend)". iSpecialize ("Hend" with "Htok"). wp_bind Endlft. - iApply (wp_mask_mono _ (↑lftN ∪ ↑histN)); [set_solver..|]. - iApply (wp_step_fupd with "Hend"). done. set_solver. wp_seq. + iApply (wp_mask_mono _ (↑lftN ∪ lft_userE)); first done. + iApply (wp_step_fupd with "Hend"); [done|set_solver-|]. wp_seq. iIntros "#Hdead !>". wp_seq. iApply ("He" with "LFT HE Htl HL HC [> -]"). iApply (Hub with "[] HT"). simpl in *. subst κ. rewrite -lft_dead_or. auto. Qed. @@ -193,7 +231,7 @@ Section typing_rules. Lemma type_path_instr {E L} p ty : ⊢ typed_instruction_ty E L [p â— ty] p ty. Proof. - iIntros (?) "_ _ $$ [? _]". + iIntros (??) "_ _ $$ [? _]". iApply (wp_hasty with "[-]"); [done..|]. iIntros (v) "_ Hv". rewrite tctx_interp_singleton. iExists v. iFrame. by rewrite eval_path_of_val. Qed. @@ -209,18 +247,20 @@ Section typing_rules. (⊢ typed_write E L ty1 ty ty1') → (⊢ typed_instruction E L [p1 â— ty1; p2 â— ty] (p1 <- p2) (λ _, [p1 â— ty1'])). Proof. - iIntros (Hwrt tid) "#LFT #HE $ HL". + iIntros (Hwrt tid ?) "#LFT #HE $ HL". rewrite tctx_interp_cons tctx_interp_singleton. iIntros "[Hp1 Hp2]". wp_bind p1. iApply (wp_hasty with "Hp1"); [done|]. iIntros (v1) "% Hown1". wp_bind p2. iApply (wp_hasty with "Hp2"); [done|]. iIntros (v2) "_ Hown2". rewrite typed_write_eq in Hwrt. - iMod (Hwrt with "[] LFT HE HL Hown1") as (l vl [?[= ->]]) "(Hl & Hclose)"; first done. - iDestruct (ty_size_eq with "Hown2") as "#Hsz". iDestruct "Hsz" as %Hsz. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". + iMod (Hwrt with "[] LFT HE HL Hown1") as (l vl) "([% %] & Hl & Hclose)"; first done. + subst v1. iDestruct (ty_size_eq with "Hown2") as "#Hsz". iDestruct "Hsz" as %Hsz. rewrite <-Hsz in *. destruct vl as [|v[|]]; try done. rewrite heap_mapsto_vec_singleton. iApply wp_fupd. wp_write. rewrite -heap_mapsto_vec_singleton. - iMod ("Hclose" with "[Hl Hown2]") as "($ & Hown)". + iMod ("Hclose" with "[Hl Hown2]") as "(HL & Hown)". { iExists _. iFrame. } + iDestruct ("HLclose" with "HL") as "$". rewrite tctx_interp_singleton tctx_hasty_val' //. Qed. @@ -236,16 +276,18 @@ Section typing_rules. ty.(ty_size) = 1%nat → (⊢ typed_read E L ty1 ty ty1') → (⊢ typed_instruction E L [p â— ty1] (!p) (λ v, [p â— ty1'; v â— ty])). Proof. - iIntros (Hsz Hread tid) "#LFT #HE Htl HL Hp". + iIntros (Hsz Hread tid qmax) "#LFT #HE Htl HL Hp". rewrite tctx_interp_singleton. wp_bind p. iApply (wp_hasty with "Hp"); [done|]. iIntros (v) "% Hown". rewrite typed_read_eq in Hread. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Hread with "[] LFT HE Htl HL Hown") as (l vl q [= ->]) "(Hl & Hown & Hclose)"; first done. iDestruct (ty_size_eq with "Hown") as "#>%". rewrite ->Hsz in *. destruct vl as [|v [|]]; try done. rewrite heap_mapsto_vec_singleton. iApply wp_fupd. wp_read. - iMod ("Hclose" with "Hl") as "($ & $ & Hown2)". + iMod ("Hclose" with "Hl") as "($ & HL & Hown2)". + iDestruct ("HLclose" with "HL") as "$". rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. by iFrame. Qed. @@ -259,16 +301,16 @@ Section typing_rules. typed_body E L C T (let: x := !p in e). Proof. iIntros. by iApply type_let; [apply type_deref_instr|solve_typing|]. Qed. - Lemma type_memcpy_iris E L qL tid ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 : + Lemma type_memcpy_iris E L qmax qL tid ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 : Z.of_nat (ty.(ty_size)) = n → typed_write E L ty1 ty ty1' -∗ typed_read E L ty2 ty ty2' -∗ - {{{ ⎡lft_ctx⎤ ∗ elctx_interp E ∗ na_own tid ⊤ ∗ llctx_interp L qL ∗ + {{{ ⎡lft_ctx⎤ ∗ elctx_interp E ∗ na_own tid ⊤ ∗ llctx_interp_noend qmax L qL ∗ tctx_elt_interp tid (p1 â— ty1) ∗ tctx_elt_interp tid (p2 â— ty2) }}} (p1 <-{n} !p2) in tid - {{{ RET #☠; na_own tid ⊤ ∗ llctx_interp L qL ∗ + {{{ RET #☠; na_own tid ⊤ ∗ llctx_interp_noend qmax L qL ∗ tctx_elt_interp tid (p1 â— ty1') ∗ tctx_elt_interp tid (p2 â— ty2') }}}. Proof. - iIntros (<-) "#Hwrt #Hread !#". + iIntros (<-) "#Hwrt #Hread !>". iIntros (Φ) "(#LFT & #HE & Htl & [HL1 HL2] & [Hp1 Hp2]) HΦ". wp_bind p1. iApply (wp_hasty with "Hp1"); [done|]. iIntros (v1) "% Hown1". wp_bind p2. iApply (wp_hasty with "Hp2"); [done|]. iIntros (v2) "% Hown2". @@ -292,10 +334,12 @@ Section typing_rules. ⊢ typed_instruction E L [p1 â— ty1; p2 â— ty2] (p1 <-{n} !p2) (λ _, [p1 â— ty1'; p2 â— ty2']). Proof. - iIntros (Hsz Hwrt Hread tid) "#LFT #HE Htl HL HT". + iIntros (Hsz Hwrt Hread tid qmax) "#LFT #HE Htl HL HT". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iApply (type_memcpy_iris with "[] [] [$LFT $Htl $HE $HL HT]"); try done. { by rewrite tctx_interp_cons tctx_interp_singleton. } - rewrite tctx_interp_cons tctx_interp_singleton. auto. + rewrite tctx_interp_cons tctx_interp_singleton. + iIntros "!> ($ & HL & $ & $)". by iApply "HLclose". Qed. Lemma type_memcpy {E L} ty ty1 ty2 (n : Z) C T T' ty1' ty2' p1 p2 e: diff --git a/theories/typing/shr_bor.v b/theories/typing/shr_bor.v index cf899f62b334f43b7c45a3458569cec94193e526..e6042d95081afc2f6a9713cb8924718d801f90f2 100644 --- a/theories/typing/shr_bor.v +++ b/theories/typing/shr_bor.v @@ -22,7 +22,7 @@ Section shr_bor. κ2 ⊑ κ1 -∗ type_incl ty1 ty2 -∗ type_incl (shr_bor κ1 ty1) (shr_bor κ2 ty2). Proof. iIntros "#Hκ #[_ [_ Hty]]". iApply type_incl_simple_type. simpl. - iIntros "!#" (tid [|[[]|][]]) "H"; try done. iApply "Hty". + iIntros "!>" (tid [|[[]|][]]) "H"; try done. iApply "Hty". iApply (ty1.(ty_shr_mono) with "Hκ"). done. Qed. @@ -30,10 +30,11 @@ Section shr_bor. Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) shr_bor. Proof. intros κ1 κ2 Hκ ty1 ty2 Hty. - iIntros (?) "/= HL". iDestruct (Hκ with "HL") as "#Hincl". - iDestruct (Hty with "HL") as "#Hty". iIntros "!# #HE". + iIntros (??) "/= HL". iDestruct (Hκ with "HL") as "#Hincl". + iDestruct (Hty with "HL") as "#Hty". iIntros "!> #HE". + iDestruct ("Hincl" with "HE") as "%". iApply shr_type_incl. - - by iApply "Hincl". + - by iApply lft_incl_syn_sem. - by iApply "Hty". Qed. Global Instance shr_mono_flip E L : @@ -66,26 +67,27 @@ Section typing. lctx_lft_incl E L κ2 κ1 → subtype E L ty1 ty2 → subtype E L (&shr{κ1}ty1) (&shr{κ2}ty2). Proof. by intros; apply shr_mono. Qed. - Lemma shr_proper' E L κ ty1 ty2 : - eqtype E L ty1 ty2 → eqtype E L (&shr{κ}ty1) (&shr{κ}ty2). + Lemma shr_proper' E L κ1 κ2 ty1 ty2 : + lctx_lft_eq E L κ1 κ2 → eqtype E L ty1 ty2 → eqtype E L (&shr{κ1}ty1) (&shr{κ2}ty2). Proof. by intros; apply shr_proper. Qed. Lemma tctx_reborrow_shr E L p ty κ κ' : lctx_lft_incl E L κ' κ → tctx_incl E L [p â— &shr{κ}ty] [p â— &shr{κ'}ty; p â—{κ} &shr{κ}ty]. Proof. - iIntros (Hκκ' tid ?) "#LFT #HE HL [H _]". iDestruct (Hκκ' with "HL HE") as "#Hκκ'". + iIntros (Hκκ' tid ??) "#LFT #HE HL [H _]". iDestruct (Hκκ' with "HL HE") as "%". iFrame. rewrite /tctx_interp /=. iDestruct "H" as ([[]|]) "[% #Hshr]"; try done. iModIntro. iSplit. - - iExists _. iSplit. done. by iApply (ty_shr_mono with "Hκκ' Hshr"). + - iExists _. iSplit. done. iApply (ty_shr_mono with "[] Hshr"). + by iApply lft_incl_syn_sem. - iSplit=> //. iExists _. auto. Qed. Lemma read_shr E L κ ty : Copy ty → lctx_lft_alive E L κ → ⊢ typed_read E L (&shr{κ}ty) ty (&shr{κ}ty). Proof. - rewrite typed_read_eq. iIntros (Hcopy Halive) "!#". - iIntros ([[]|] tid F qL ?) "#LFT #HE Htl HL #Hshr"; try done. + rewrite typed_read_eq. iIntros (Hcopy Halive) "!>". + iIntros ([[]|] tid F qmax qL ?) "#LFT #HE Htl HL #Hshr"; try done. iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first solve_ndisj. iMod (copy_shr_acc with "LFT Hshr Htl Hκ") as (q') "(Htl & H↦ & Hcl)"; first solve_ndisj. { rewrite ->shr_locsE_shrN. solve_ndisj. } @@ -96,4 +98,4 @@ Section typing. Qed. End typing. -Hint Resolve shr_mono' shr_proper' read_shr : lrust_typing. +Global Hint Resolve shr_mono' shr_proper' read_shr : lrust_typing. diff --git a/theories/typing/soundness.v b/theories/typing/soundness.v index 3e282b53ed3a4293e92457df7da1cc55c0a8f6c0..64a46e93d93aac23d7309b4cf2686f908e6fce40 100644 --- a/theories/typing/soundness.v +++ b/theories/typing/soundness.v @@ -10,7 +10,7 @@ Set Default Proof Using "Type". Class typePreG Σ := PreTypeG { type_heapG :> noprolPreG Σ; - type_lftG :> lftPreG view_Lat Σ; + type_lftG :> lftPreG Σ view_Lat; type_preG_na_invG :> na_invG view_Lat Σ; type_frac_borrowG :> frac_borG Σ }. @@ -40,22 +40,27 @@ Section type_soundness. - destruct toval as [v ?%nopro_lang.to_base_val]. eauto. - eauto. } apply: (noprol_adequacy' _ _ (λ x, True))=>? tid1. - iMod (lft_init _ (↑histN)) as (?) "#LFT"; [done|solve_ndisj|]. + iMod (lft_init _ lft_userE) as (?) "#LFT"; [done|solve_ndisj|]. iMod na_alloc as (tid2) "Htl". set (Htype := TypeG _ _ _ _ _). set (tid := {| tid_na_inv_pool := tid2; tid_tid := tid1 |}). wp_bind (of_val main). iApply (wp_wand with "[Htl]"). - { iApply (Hmain Htype [] [] $! tid with "LFT [] Htl [] []"); - by rewrite /elctx_interp /llctx_interp ?tctx_interp_nil. } + iApply (Hmain Htype [] [] $! tid 1%Qp with "LFT [] Htl [] []"). + { by rewrite /elctx_interp big_sepL_nil. } + { by rewrite /llctx_interp big_sepL_nil. } + { by rewrite tctx_interp_nil. } clear Hrtc Hmain main. iIntros (main) "(Htl & _ & Hmain & _)". iDestruct "Hmain" as (?) "[EQ Hmain]". rewrite eval_path_of_val. iDestruct "EQ" as %[= <-]. iDestruct "Hmain" as (f k x e ?) "(EQ & % & Hmain)". iDestruct "EQ" as %[= ->]. destruct x; try done. wp_rec. iMod (lft_create with "LFT") as (Ï) "HÏ"; first done. - iApply ("Hmain" $! () Ï exit_cont [#] tid with "LFT [] Htl [HÏ] []"); + iApply ("Hmain" $! () Ï exit_cont [#] tid 1%Qp with "LFT [] Htl [HÏ] []"); last by rewrite tctx_interp_nil. { by rewrite /elctx_interp /=. } - { rewrite /llctx_interp /=. iSplit; last done. iExists Ï. iFrame. by rewrite /= left_id. } + { rewrite /llctx_interp /=. iSplit; last done. iExists Ï. + rewrite /= left_id. iSplit; first done. + rewrite decide_True //. + by iDestruct "HÏ" as "[$ #$]". } rewrite cctx_interp_singleton. simpl. iIntros (args) "_ _". inv_vec args. iIntros (x) "_ /=". by wp_lam. Qed. diff --git a/theories/typing/sum.v b/theories/typing/sum.v index 8af91c2efb6c4c1c67815daf5c4b243ae8981f8c..4d42bb78362aae88eca95cb46c64f082a69da907 100644 --- a/theories/typing/sum.v +++ b/theories/typing/sum.v @@ -1,6 +1,7 @@ From iris.proofmode Require Import tactics. From iris.algebra Require Import list. From iris.bi Require Import fractional. +From lrust.typing Require Import lft_contexts. From lrust.typing Require Export type. Set Default Proof Using "Type". @@ -125,15 +126,15 @@ Section sum. Global Instance sum_mono E L : Proper (Forall2 (subtype E L) ==> subtype E L) sum. Proof. - iIntros (tyl1 tyl2 Htyl qL) "HL". - iAssert (â–¡ (lft_contexts.elctx_interp E -∗ ⌜max_list_with ty_size tyl1 = max_list_with ty_size tyl2âŒ))%I with "[#]" as "#Hleq". + iIntros (tyl1 tyl2 Htyl qmax qL) "HL". + iAssert (â–¡ (elctx_interp E -∗ ⌜max_list_with ty_size tyl1 = max_list_with ty_size tyl2âŒ))%I with "[#]" as "#Hleq". { iInduction Htyl as [|???? Hsub] "IH". - { iClear "∗". iIntros "!# _". done. } + { iClear "∗". iIntros "!> _". done. } iDestruct (Hsub with "HL") as "#Hsub". iDestruct ("IH" with "HL") as "{IH} #IH". iModIntro. iIntros "#HE". iDestruct ("Hsub" with "HE") as "(% & _ & _)". iDestruct ("IH" with "HE") as %IH. iPureIntro. simpl. inversion_clear IH. by f_equal. } - iDestruct (subtype_Forall2_llctx with "HL") as "#Htyl"; first done. - iClear "HL". iIntros "!# #HE". + iDestruct (subtype_Forall2_llctx_noend with "HL") as "#Htyl"; first done. + iClear "HL". iIntros "!> #HE". iDestruct ("Hleq" with "HE") as %Hleq. iSpecialize ("Htyl" with "HE"). iClear "Hleq HE". iAssert (∀ i, type_incl (nth i tyl1 emp0) (nth i tyl2 emp0))%I with "[#]" as "#Hty". { iIntros (i). edestruct (nth_lookup_or_length tyl1 i) as [Hl1|Hl]; last first. @@ -240,5 +241,5 @@ End sum. Notation "Σ[ ty1 ; .. ; tyn ]" := (sum (cons ty1%T (..(cons tyn%T nil)..))) : lrust_type_scope. -Hint Opaque sum : lrust_typing lrust_typing_merge. -Hint Resolve sum_mono' sum_proper' : lrust_typing. +Global Hint Opaque sum : lrust_typing lrust_typing_merge. +Global Hint Resolve sum_mono' sum_proper' : lrust_typing. diff --git a/theories/typing/type.v b/theories/typing/type.v index 746d8f078fc2a0899a7847b502717735a8aa3c51..8528f6d536a3940e089e75ca5f0f9b15fa46a14d 100644 --- a/theories/typing/type.v +++ b/theories/typing/type.v @@ -9,9 +9,9 @@ Set Default Proof Using "Type". Class typeG Σ := TypeG { type_noprolG :> noprolG Σ; - type_lftG :> lftG view_Lat (↑histN) Σ; + type_lftG :> lftG Σ view_Lat lft_userE; type_na_invG :> na_invG view_Lat Σ; - type_frac_borrowG :> frac_borG Σ; + type_frac_borrowG :> frac_borG Σ }. Definition lftE : coPset := ↑lftN. @@ -21,6 +21,8 @@ Definition shrN := lrustN .@ "shr". Record thread_id := { tid_na_inv_pool :> na_inv_pool_name; tid_tid :> history.thread_id }. +Instance thread_id_inhabited : Inhabited thread_id. +Proof. repeat constructor. Qed. Record type `{!typeG Σ} := { ty_size : nat; @@ -522,6 +524,18 @@ Section type. iIntros (Hsend κ tid1 tid2 l). iDestruct 1 as (vl) "[Hm Hown]". iExists vl. iFrame "Hm". iNext. by iApply Hsend. Qed. + + Lemma send_change_tid' t tid1 tid2 vl : + Send t → t.(ty_own) tid1 vl ≡ t.(ty_own) tid2 vl. + Proof. + intros ?. apply: anti_symm; apply send_change_tid. + Qed. + + Lemma sync_change_tid' t κ tid1 tid2 l : + Sync t → t.(ty_shr) κ tid1 l ≡ t.(ty_shr) κ tid2 l. + Proof. + intros ?. apply: anti_symm; apply sync_change_tid. + Qed. End type. Definition type_incl `{!typeG Σ} (ty1 ty2 : type) : vPropI Σ := @@ -531,8 +545,14 @@ Definition type_incl `{!typeG Σ} (ty1 ty2 : type) : vPropI Σ := Instance: Params (@type_incl) 2 := {}. (* Typeclasses Opaque type_incl. *) +Definition type_equal `{!typeG Σ} (ty1 ty2 : type) : vProp Σ := + (⌜ty1.(ty_size) = ty2.(ty_size)⌠∗ + (â–¡ ∀ tid vl, ty1.(ty_own) tid vl ∗-∗ ty2.(ty_own) tid vl) ∗ + (â–¡ ∀ κ tid l, ty1.(ty_shr) κ tid l ∗-∗ ty2.(ty_shr) κ tid l))%I. +Instance: Params (@type_equal) 2 := {}. + Definition subtype `{!typeG Σ} (E : elctx) (L : llctx) (ty1 ty2 : type) : Prop := - ∀ qL, llctx_interp L qL -∗ â–¡ (elctx_interp E -∗ type_incl ty1 ty2). + ∀ qmax qL, llctx_interp_noend qmax L qL -∗ â–¡ (elctx_interp E -∗ type_incl ty1 ty2). Instance: Params (@subtype) 4 := {}. (* TODO: The prelude should have a symmetric closure. *) @@ -565,20 +585,20 @@ Section subtyping. Qed. Lemma subtype_refl E L ty : subtype E L ty ty. - Proof. iIntros (?) "_ !# _". iApply type_incl_refl. Qed. + Proof. iIntros (??) "_ !> _". iApply type_incl_refl. Qed. Global Instance subtype_preorder E L : PreOrder (subtype E L). Proof. split; first by intros ?; apply subtype_refl. - intros ty1 ty2 ty3 H12 H23. iIntros (?) "HL". + iIntros (ty1 ty2 ty3 H12 H23 ??) "HL". iDestruct (H12 with "HL") as "#H12". iDestruct (H23 with "HL") as "#H23". - iClear "∗". iIntros "!# #HE". + iClear "∗". iIntros "!> #HE". iApply (type_incl_trans with "[#]"). by iApply "H12". by iApply "H23". Qed. - Lemma subtype_Forall2_llctx E L tys1 tys2 qL : + Lemma subtype_Forall2_llctx_noend E L tys1 tys2 qmax qL : Forall2 (subtype E L) tys1 tys2 → - llctx_interp L qL -∗ â–¡ (elctx_interp E -∗ + llctx_interp_noend qmax L qL -∗ â–¡ (elctx_interp E -∗ [∗ list] tys ∈ (zip tys1 tys2), type_incl (tys.1) (tys.2)). Proof. iIntros (Htys) "HL". @@ -587,42 +607,95 @@ Section subtyping. { iApply big_sepL_forall. iIntros (k [ty1 ty2] Hlookup). move:Htys => /Forall2_Forall /Forall_forall=>Htys. iApply (Htys (ty1, ty2)); first by exact: elem_of_list_lookup_2. done. } - iClear "∗". iIntros "!# #HE". iApply (big_sepL_impl with "[$Htys]"). - iIntros "!# * % #Hincl". by iApply "Hincl". + iClear "∗". iIntros "!> #HE". iApply (big_sepL_impl with "[$Htys]"). + iIntros "!> * % #Hincl". by iApply "Hincl". Qed. + Lemma subtype_Forall2_llctx E L tys1 tys2 qmax : + Forall2 (subtype E L) tys1 tys2 → + llctx_interp qmax L -∗ â–¡ (elctx_interp E -∗ + [∗ list] tys ∈ (zip tys1 tys2), type_incl (tys.1) (tys.2)). + Proof. + iIntros (?) "HL". iApply subtype_Forall2_llctx_noend; first done. + iDestruct (llctx_interp_acc_noend with "HL") as "[$ _]". + Qed. + + Lemma lft_invariant_subtype E L T : + Proper (lctx_lft_eq E L ==> subtype E L) T. + Proof. + iIntros (x y [Hxy Hyx] qmax qL) "L". + iPoseProof (Hxy with "L") as "#Hxy". + iPoseProof (Hyx with "L") as "#Hyx". + iIntros "!> #E". clear Hxy Hyx. + iDestruct ("Hxy" with "E") as %Hxy. + iDestruct ("Hyx" with "E") as %Hyx. + iClear "Hyx Hxy". + rewrite (anti_symm _ _ _ Hxy Hyx). + iApply type_incl_refl. + Qed. + + Lemma type_equal_incl ty1 ty2 : + type_equal ty1 ty2 ⊣⊢ type_incl ty1 ty2 ∗ type_incl ty2 ty1. + Proof. + iSplit. + - iIntros "#(% & Ho & Hs)". + iSplit; (iSplit; first done; iSplit; iModIntro). + + iIntros (??) "?". by iApply "Ho". + + iIntros (???) "?". by iApply "Hs". + + iIntros (??) "?". by iApply "Ho". + + iIntros (???) "?". by iApply "Hs". + - iIntros "#[(% & Ho1 & Hs1) (% & Ho2 & Hs2)]". + iSplit; first done. iSplit; iModIntro. + + iIntros (??). iSplit; [iApply "Ho1"|iApply "Ho2"]. + + iIntros (???). iSplit; [iApply "Hs1"|iApply "Hs2"]. + Qed. + + Lemma type_equal_refl ty : + ⊢ type_equal ty ty. + Proof. + iApply type_equal_incl. iSplit; iApply type_incl_refl. + Qed. + Lemma type_equal_trans ty1 ty2 ty3 : + type_equal ty1 ty2 -∗ type_equal ty2 ty3 -∗ type_equal ty1 ty3. + Proof. + rewrite !type_equal_incl. iIntros "#[??] #[??]". iSplit. + - iApply (type_incl_trans _ ty2); done. + - iApply (type_incl_trans _ ty2); done. + Qed. + + Lemma lft_invariant_eqtype E L T : + Proper (lctx_lft_eq E L ==> eqtype E L) T. + Proof. split; by apply lft_invariant_subtype. Qed. + Lemma equiv_subtype E L ty1 ty2 : ty1 ≡ ty2 → subtype E L ty1 ty2. Proof. unfold subtype, type_incl=>EQ. setoid_rewrite EQ. apply subtype_refl. Qed. Lemma eqtype_unfold E L ty1 ty2 : eqtype E L ty1 ty2 ↔ - (∀ qL, llctx_interp L qL -∗ â–¡ (elctx_interp E -∗ - (⌜ty1.(ty_size) = ty2.(ty_size)⌠∗ - (â–¡ ∀ tid vl, ty1.(ty_own) tid vl ↔ ty2.(ty_own) tid vl) ∗ - (â–¡ ∀ κ tid l, ty1.(ty_shr) κ tid l ↔ ty2.(ty_shr) κ tid l)))%I). + (∀ qmax qL, llctx_interp_noend qmax L qL -∗ â–¡ (elctx_interp E -∗ type_equal ty1 ty2)). Proof. split. - - iIntros ([EQ1 EQ2] qL) "HL". + - iIntros ([EQ1 EQ2] qmax qL) "HL". iDestruct (EQ1 with "HL") as "#EQ1". iDestruct (EQ2 with "HL") as "#EQ2". - iClear "∗". iIntros "!# #HE". + iClear "∗". iIntros "!> #HE". iDestruct ("EQ1" with "HE") as "[#Hsz [#H1own #H1shr]]". iDestruct ("EQ2" with "HE") as "[_ [#H2own #H2shr]]". iSplit; last iSplit. + done. - + by iIntros "!#*"; iSplit; iIntros "H"; [iApply "H1own"|iApply "H2own"]. - + by iIntros "!#*"; iSplit; iIntros "H"; [iApply "H1shr"|iApply "H2shr"]. - - intros EQ. split; (iIntros (qL) "HL"; + + by iIntros "!>*"; iSplit; iIntros "H"; [iApply "H1own"|iApply "H2own"]. + + by iIntros "!>*"; iSplit; iIntros "H"; [iApply "H1shr"|iApply "H2shr"]. + - intros EQ. split; (iIntros (qmax qL) "HL"; iDestruct (EQ with "HL") as "#EQ"; - iClear "∗"; iIntros "!# #HE"; + iClear "∗"; iIntros "!> #HE"; iDestruct ("EQ" with "HE") as "[% [#Hown #Hshr]]"; (iSplit; last iSplit)). + done. - + iIntros "!#* H". by iApply "Hown". - + iIntros "!#* H". by iApply "Hshr". + + iIntros "!>* H". by iApply "Hown". + + iIntros "!>* H". by iApply "Hshr". + done. - + iIntros "!#* H". by iApply "Hown". - + iIntros "!#* H". by iApply "Hshr". + + iIntros "!>* H". by iApply "Hown". + + iIntros "!>* H". by iApply "Hshr". Qed. Lemma eqtype_refl E L ty : eqtype E L ty ty. @@ -654,22 +727,22 @@ Section subtyping. Qed. Lemma subtype_simple_type E L (st1 st2 : simple_type) : - (∀ qL, llctx_interp L qL -∗ â–¡ (elctx_interp E -∗ + (∀ qmax qL, llctx_interp_noend qmax L qL -∗ â–¡ (elctx_interp E -∗ ∀ tid vl, st1.(st_own) tid vl -∗ st2.(st_own) tid vl)) → subtype E L st1 st2. Proof. - intros Hst. iIntros (qL) "HL". iDestruct (Hst with "HL") as "#Hst". - iClear "∗". iIntros "!# #HE". iApply type_incl_simple_type. - iIntros "!#" (??) "?". by iApply "Hst". + intros Hst. iIntros (qmax qL) "HL". iDestruct (Hst with "HL") as "#Hst". + iClear "∗". iIntros "!> #HE". iApply type_incl_simple_type. + iIntros "!>" (??) "?". by iApply "Hst". Qed. Lemma subtype_weaken E1 E2 L1 L2 ty1 ty2 : E1 ⊆+ E2 → L1 ⊆+ L2 → subtype E1 L1 ty1 ty2 → subtype E2 L2 ty1 ty2. Proof. - iIntros (HE12 ? Hsub qL) "HL". iDestruct (Hsub with "[HL]") as "#Hsub". + iIntros (HE12 ? Hsub qmax qL) "HL". iDestruct (Hsub with "[HL]") as "#Hsub". { rewrite /llctx_interp. by iApply big_sepL_submseteq. } - iClear "∗". iIntros "!# #HE". iApply "Hsub". + iClear "∗". iIntros "!> #HE". iApply "Hsub". rewrite /elctx_interp. by iApply big_sepL_submseteq. Qed. End subtyping. @@ -686,9 +759,8 @@ Section type_util. iExists (list_to_vec vl). rewrite vec_to_list_to_vec. iFrame. - iIntros "H". iDestruct "H" as (vl) "[Hl Hown]". eauto with iFrame. Qed. - End type_util. -Hint Resolve ty_outlives_E_elctx_sat tyl_outlives_E_elctx_sat : lrust_typing. -Hint Resolve subtype_refl eqtype_refl : lrust_typing. -Hint Opaque subtype eqtype : lrust_typing. +Global Hint Resolve ty_outlives_E_elctx_sat tyl_outlives_E_elctx_sat : lrust_typing. +Global Hint Resolve subtype_refl eqtype_refl : lrust_typing. +Global Hint Opaque subtype eqtype : lrust_typing. diff --git a/theories/typing/type_context.v b/theories/typing/type_context.v index 40fd8997d43816f7626d8f317fa97dec7dd32e74..f45b11d1a605e9e0863a1a5dbe6d0adbc7a3cbb9 100644 --- a/theories/typing/type_context.v +++ b/theories/typing/type_context.v @@ -165,28 +165,28 @@ Section type_context. (** Type context inclusion *) Definition tctx_incl (E : elctx) (L : llctx) (T1 T2 : tctx): Prop := - ∀ tid q2, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ llctx_interp L q2 -∗ - tctx_interp tid T1 ={⊤}=∗ llctx_interp L q2 ∗ tctx_interp tid T2. + ∀ tid qmax q2, ⎡lft_ctx⎤ -∗ elctx_interp E -∗ llctx_interp_noend qmax L q2 -∗ + tctx_interp tid T1 ={⊤}=∗ llctx_interp_noend qmax L q2 ∗ tctx_interp tid T2. Global Instance : ∀ E L, RewriteRelation (tctx_incl E L) := {}. Global Instance tctx_incl_preorder E L : PreOrder (tctx_incl E L). Proof. split. - - by iIntros (???) "_ _ $ $". - - iIntros (??? H1 H2 ??) "#LFT #HE HL H". + - by iIntros (????) "_ _ $ $". + - iIntros (??? H1 H2 ???) "#LFT #HE HL H". iMod (H1 with "LFT HE HL H") as "(HL & H)". by iMod (H2 with "LFT HE HL H") as "($ & $)". Qed. Lemma contains_tctx_incl E L T1 T2 : T1 ⊆+ T2 → tctx_incl E L T2 T1. Proof. - rewrite /tctx_incl. iIntros (Hc ??) "_ _ $ H". by iApply big_sepL_submseteq. + rewrite /tctx_incl. iIntros (Hc ???) "_ _ $ H". by iApply big_sepL_submseteq. Qed. Global Instance tctx_incl_app E L : Proper (tctx_incl E L ==> tctx_incl E L ==> tctx_incl E L) app. Proof. - intros ?? Hincl1 ?? Hincl2 ??. rewrite /tctx_interp !big_sepL_app. + intros ?? Hincl1 ?? Hincl2 ???. rewrite /tctx_interp !big_sepL_app. iIntros "#LFT #HE HL [H1 H2]". iMod (Hincl1 with "LFT HE HL H1") as "(HL & $)". iApply (Hincl2 with "LFT HE HL H2"). @@ -204,7 +204,7 @@ Section type_context. Lemma copy_tctx_incl E L p `{!Copy ty} : tctx_incl E L [p â— ty] [p â— ty; p â— ty]. - Proof. iIntros (??) "_ _ $ * [#$ $] //". Qed. + Proof. iIntros (???) "_ _ $ * [#$ $] //". Qed. Lemma copy_elem_of_tctx_incl E L T p `{!Copy ty} : p â— ty ∈ T → tctx_incl E L T ((p â— ty) :: T). @@ -215,12 +215,20 @@ Section type_context. apply contains_tctx_incl, submseteq_swap. Qed. + Lemma type_incl_tctx_incl tid p ty1 ty2 : + type_incl ty1 ty2 -∗ tctx_interp tid [p â— ty1] -∗ tctx_interp tid [p â— ty2]. + Proof. + iIntros "Hincl HT". rewrite !tctx_interp_singleton. + iDestruct "HT" as (v) "[% HT]". iExists _. iFrame "%". + iDestruct "Hincl" as "(_ & Hincl & _)". iApply "Hincl". done. + Qed. Lemma subtype_tctx_incl E L p ty1 ty2 : subtype E L ty1 ty2 → tctx_incl E L [p â— ty1] [p â— ty2]. Proof. - iIntros (Hst ??) "#LFT #HE HL [H _]". - iDestruct "H" as (v) "[% H]". iDestruct (Hst with "HL HE") as "#(_ & Ho & _)". - iFrame "HL". iApply big_sepL_singleton. iExists _. iFrame "%". by iApply "Ho". + iIntros (Hst ???) "#LFT #HE HL HT". + iDestruct (Hst with "HL") as "#Hst". iFrame "HL". + iApply type_incl_tctx_incl; last done. + by iApply "Hst". Qed. (* Extracting from a type context. *) @@ -312,17 +320,17 @@ Section type_context. Qed. End type_context. -Hint Resolve tctx_extract_hasty_here_copy | 1 : lrust_typing. -Hint Resolve tctx_extract_hasty_here | 20 : lrust_typing. -Hint Resolve tctx_extract_hasty_further | 50 : lrust_typing. -Hint Resolve tctx_extract_blocked_here tctx_extract_blocked_cons +Global Hint Resolve tctx_extract_hasty_here_copy | 1 : lrust_typing. +Global Hint Resolve tctx_extract_hasty_here | 20 : lrust_typing. +Global Hint Resolve tctx_extract_hasty_further | 50 : lrust_typing. +Global Hint Resolve tctx_extract_blocked_here tctx_extract_blocked_cons tctx_extract_ctx_nil tctx_extract_ctx_hasty tctx_extract_ctx_blocked tctx_extract_ctx_incl : lrust_typing. -Hint Opaque tctx_extract_ctx tctx_extract_hasty tctx_extract_blocked +Global Hint Opaque tctx_extract_ctx tctx_extract_hasty tctx_extract_blocked tctx_incl : lrust_typing. (* In general, we want reborrowing to be tried before subtyping, so that we get the extraction. However, in the case the types match exactly, we want to NOT use reborrowing. Therefore, we add [tctx_extract_hasty_here_eq] as a hint with a very low cost. *) -Hint Resolve tctx_extract_hasty_here_eq | 2 : lrust_typing. +Global Hint Resolve tctx_extract_hasty_here_eq | 2 : lrust_typing. diff --git a/theories/typing/type_sum.v b/theories/typing/type_sum.v index 3cfbd04e47100450b2cf0f21b6d1e3c0b166a9a8..f52ac69f24a049ba2a9406081144294cfe8571f8 100644 --- a/theories/typing/type_sum.v +++ b/theories/typing/type_sum.v @@ -7,7 +7,7 @@ Set Default Proof Using "Type". Section case. Context `{!typeG Σ}. - (* FIXME : have a iris version of Forall2. *) + (* FIXME : have an Iris version of Forall2. *) Lemma type_case_own' E L C T p n tyl el : Forall2 (λ ty e, (⊢ typed_body E L C ((p +â‚— #0 â— own_ptr n (uninit 1)) :: (p +â‚— #1 â— own_ptr n ty) :: @@ -17,7 +17,7 @@ Section case. tyl el → ⊢ typed_body E L C ((p â— own_ptr n (sum tyl)) :: T) (case: !p of el). Proof. - iIntros (Hel tid) "#LFT #HE Hna HL HC HT". wp_bind p. + iIntros (Hel tid qmax) "#LFT #HE Hna HL HC HT". wp_bind p. rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]". iApply (wp_hasty with "Hp"); [done|]. iIntros ([[]|] Hv) "Hp"; try done. iDestruct "Hp" as "[H↦ Hf]". iDestruct "H↦" as (vl) "[H↦ Hown]". @@ -63,10 +63,11 @@ Section case. (⊢ typed_body E L C ((p â— &uniq{κ}(sum tyl)) :: T) e)) tyl el → ⊢ typed_body E L C ((p â— &uniq{κ}(sum tyl)) :: T) (case: !p of el). Proof. - iIntros (Halive Hel tid) "#LFT #HE Hna HL HC HT". wp_bind p. + iIntros (Halive Hel tid qmax) "#LFT #HE Hna HL HC HT". wp_bind p. rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]". iApply (wp_hasty with "Hp"); [done|]. iIntros ([[]|] Hv) "Hp"; try iDestruct "Hp" as "[]". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Halive with "HE HL") as (q) "[Htok Hclose]". done. iMod (bor_acc_cons with "LFT Hp Htok") as "[H↦ Hclose']". done. iDestruct "H↦" as (vl) "[H↦ Hown]". @@ -89,6 +90,7 @@ Section case. rewrite /= -EQlen !app_length EQlenvl' EQlenvl'2 nth_lookup EQty /=. auto. } { iExists vl'. iFrame. } iMod ("Hclose" with "Htok") as "HL". + iDestruct ("HLclose" with "HL") as "HL". iApply (Hety with "LFT HE Hna HL HC"). rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //. iFrame. - iMod ("Hclose'" with "[] [H↦i H↦vl' H↦vl'' Hown]") as "[Hb Htok]"; @@ -97,6 +99,7 @@ Section case. rewrite heap_mapsto_vec_cons heap_mapsto_vec_app /= -EQlen. iFrame. iNext. iExists i, vl', vl''. rewrite nth_lookup EQty. auto. } iMod ("Hclose" with "Htok") as "HL". + iDestruct ("HLclose" with "HL") as "HL". iApply (Hety with "LFT HE Hna HL HC"). rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //. iFrame. Qed. @@ -117,10 +120,11 @@ Section case. (⊢ typed_body E L C ((p â— &shr{κ}(sum tyl)) :: T) e)) tyl el → ⊢ typed_body E L C ((p â— &shr{κ}(sum tyl)) :: T) (case: !p of el). Proof. - iIntros (Halive Hel tid) "#LFT #HE Hna HL HC HT". wp_bind p. + iIntros (Halive Hel tid qmax) "#LFT #HE Hna HL HC HT". wp_bind p. rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]". iApply (wp_hasty with "Hp"); [done|]. iIntros ([[]|] Hv) "Hp"; try done. iDestruct "Hp" as (i) "[#Hb Hshr]". + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Halive with "HE HL") as (q) "[Htok Hclose]". done. iMod (frac_bor_acc with "LFT Hb Htok") as (q') "[[H↦i H↦vl''] Hclose']". done. rewrite nth_lookup. @@ -129,6 +133,7 @@ Section case. wp_read. wp_case; first (split; [lia|by rewrite Nat2Z.id]). iMod ("Hclose'" with "[$H↦i $H↦vl'']") as "Htok". iMod ("Hclose" with "Htok") as "HL". + iDestruct ("HLclose" with "HL") as "HL". destruct Hety as [Hety|Hety]; iApply (Hety with "LFT HE Hna HL HC"); rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //; iFrame. iExists _. rewrite ->nth_lookup, EQty. auto. @@ -149,12 +154,13 @@ Section case. (⊢ typed_write E L ty1 (sum tyl) ty2) → ⊢ typed_instruction E L [p1 â— ty1; p2 â— ty] (p1 <-{Σ i} p2) (λ _, [p1 â— ty2]). Proof. - iIntros (Hty Hw tid) "#LFT #HE $ HL Hp". + iIntros (Hty Hw tid qmax) "#LFT #HE $ HL Hp". rewrite tctx_interp_cons tctx_interp_singleton. iDestruct "Hp" as "[Hp1 Hp2]". iDestruct (closed_hasty with "Hp1") as "%". iDestruct (closed_hasty with "Hp2") as "%". wp_bind p1. iApply (wp_hasty with "Hp1"); [done|]. iIntros (v1 Hv1) "Hty1". rewrite typed_write_eq in Hw. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". iMod (Hw with "[] LFT HE HL Hty1") as (l vl) "(H & H↦ & Hw)"=>//=. destruct vl as [|? vl]; iDestruct "H" as %[[= Hlen] [= ->]]. rewrite heap_mapsto_vec_cons. iDestruct "H↦" as "[H↦0 H↦vl]". @@ -166,7 +172,10 @@ Section case. - intros [= ->]. simpl in *. lia. - apply IHtyl. simpl in *. lia. } rewrite heap_mapsto_vec_cons -wp_fupd. iDestruct "H↦vl" as "[H↦ H↦vl]". wp_write. - rewrite tctx_interp_singleton tctx_hasty_val' //. iApply "Hw". iNext. + rewrite tctx_interp_singleton tctx_hasty_val' //. + iMod ("Hw" with "[-HLclose]") as "[HL $]"; last first. + { iApply "HLclose". done. } + iNext. iExists (_::_::_). rewrite !heap_mapsto_vec_cons. iFrame. iExists i, [_], _. rewrite -Hlen nth_lookup Hty. auto. Qed. @@ -190,15 +199,17 @@ Section case. (⊢ typed_write E L ty1 (sum tyl) ty2) → ⊢ typed_instruction E L [p â— ty1] (p <-{Σ i} ()) (λ _, [p â— ty2]). Proof. - iIntros (Hty Hw tid) "#LFT #HE $ HL Hp". rewrite tctx_interp_singleton. + iIntros (Hty Hw tid qmax) "#LFT #HE $ HL Hp". rewrite tctx_interp_singleton. wp_bind p. iApply (wp_hasty with "Hp"); [done|]. iIntros (v Hv) "Hty". rewrite typed_write_eq in Hw. - iMod (Hw with "[] LFT HE HL Hty") as (l vl) "(H & H↦ & Hw)"; first done. - clear Hw. - simpl. destruct vl as [|? vl]; iDestruct "H" as %[[= Hlen] [= ->]]. + iDestruct (llctx_interp_acc_noend with "HL") as "[HL HLclose]". + iMod (Hw with "[] LFT HE HL Hty") as (l vl) "(H & H↦ & Hw)". done. + simpl. destruct vl as [|? vl]; iDestruct "H" as %[[= Hlen] ->]. rewrite heap_mapsto_vec_cons -wp_fupd. iDestruct "H↦" as "[H↦0 H↦vl]". wp_write. rewrite tctx_interp_singleton tctx_hasty_val' //. - iApply "Hw". iModIntro. iExists (_::_). rewrite heap_mapsto_vec_cons. iFrame. + iMod ("Hw" with "[-HLclose]") as "[HL $]"; last first. + { iApply "HLclose". done. } + iModIntro. iExists (_::_). rewrite heap_mapsto_vec_cons. iFrame. iExists i, [], _. rewrite -Hlen nth_lookup Hty. auto. Qed. @@ -223,20 +234,21 @@ Section case. ⊢ typed_instruction E L [p1 â— ty1; p2 â— ty2] (p1 <-{ty.(ty_size),Σ i} !p2) (λ _, [p1 â— ty1'; p2 â— ty2']). Proof. - iIntros (Hty Hw Hr tid) "#LFT #HE Htl [HL1 HL2] Hp". + iIntros (Hty Hw Hr tid qmax) "#LFT #HE Htl HL Hp". + iDestruct (llctx_interp_acc_noend with "HL") as "[[HL1 HL2] HLclose]". rewrite tctx_interp_cons tctx_interp_singleton. iDestruct "Hp" as "[Hp1 Hp2]". iDestruct (closed_hasty with "Hp1") as "%". iDestruct (closed_hasty with "Hp2") as "%". wp_bind p1. iApply (wp_hasty with "Hp1"); [done|]. iIntros (v1 Hv1) "Hty1". rewrite typed_write_eq in Hw. iMod (Hw with "[] LFT HE HL1 Hty1") as (l1 vl1) "(H & H↦ & Hw)"=>//=. - destruct vl1 as [|? vl1]; iDestruct "H" as %[[= Hlen] [=->]]. + clear Hw. destruct vl1 as [|? vl1]; iDestruct "H" as %[[= Hlen] ->]. rewrite heap_mapsto_vec_cons -wp_fupd. iDestruct "H↦" as "[H↦0 H↦vl1]". wp_write. wp_bind p1. iApply (wp_wand with "[]"); first by iApply (wp_eval_path). iIntros (? ->). wp_op. wp_bind p2. iApply (wp_hasty with "Hp2"); [done|]. iIntros (v2 Hv2) "Hty2". rewrite typed_read_eq in Hr. - iMod (Hr with "[] LFT HE Htl HL2 Hty2") as (l2 vl2 q [= ->]) "(H↦2 & Hty & Hr)"=>//=. - clear Hr. assert (ty.(ty_size) ≤ length vl1). + iMod (Hr with "[] LFT HE Htl HL2 Hty2") as (l2 vl2 q) "(% & H↦2 & Hty & Hr)"=>//=. + clear Hr. subst. assert (ty.(ty_size) ≤ length vl1). { revert i Hty. rewrite Hlen. clear. induction tyl=>//= -[|i] /=. - intros [= ->]. lia. - specialize (IHtyl i). intuition lia. } @@ -247,7 +259,10 @@ Section case. { rewrite take_length. lia. } iNext; iIntros "[H↦vl1 H↦2]". rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. - iMod ("Hr" with "H↦2") as "($ & $ & $)". iApply "Hw". iNext. + iMod ("Hr" with "H↦2") as "($ & HL1 & $)". + iMod ("Hw" with "[-HLclose HL1]") as "[HL $]"; last first. + { iApply "HLclose". by iFrame. } + iNext. rewrite split_sum_mt /is_pad. iExists i. rewrite nth_lookup Hty. iFrame. iSplitL "H↦pad". - rewrite (shift_loc_assoc_nat _ 1) take_length Nat.min_l; last lia. @@ -280,4 +295,4 @@ Section case. Qed. End case. -Hint Resolve ty_outlives_E_elctx_sat_sum : lrust_typing. +Global Hint Resolve ty_outlives_E_elctx_sat_sum : lrust_typing. diff --git a/theories/typing/uninit.v b/theories/typing/uninit.v index 214c359f1aaae06ac98b3463810d04a8ae4a4a0f..e45318a9f54c9ed7b254e11c4122fc2b9800e55f 100644 --- a/theories/typing/uninit.v +++ b/theories/typing/uninit.v @@ -42,7 +42,7 @@ Section uninit. Next Obligation. iIntros (???) "%". done. Qed. Next Obligation. iIntros (???????) "LFT Hvl". iApply (ty_share (uninit0 n) with "LFT"); first done. - iApply (bor_iff with "[] Hvl"). iIntros "!> !#". setoid_rewrite uninit0_own. + iApply (bor_iff with "[] Hvl"). iIntros "!> !>". setoid_rewrite uninit0_own. iSplit; iIntros; done. Qed. Next Obligation. intros. by apply ty_shr_mono. Qed. @@ -132,7 +132,7 @@ Section uninit. Proof. rewrite -uninit_uninit0_eqtype=>->//. Qed. End uninit. -Hint Resolve uninit_product_eqtype_cons_l uninit_product_eqtype_cons_r +Global Hint Resolve uninit_product_eqtype_cons_l uninit_product_eqtype_cons_r uninit_product_subtype_cons_l uninit_product_subtype_cons_r uninit_unit_eqtype uninit_unit_eqtype' uninit_unit_subtype uninit_unit_subtype' : lrust_typing. diff --git a/theories/typing/uniq_bor.v b/theories/typing/uniq_bor.v index 2667230f1b6c2f1f0b00a42f10e3002520c73465..db21e0b94d7b04568eeddfc00d14676451258bea 100644 --- a/theories/typing/uniq_bor.v +++ b/theories/typing/uniq_bor.v @@ -33,7 +33,7 @@ Section uniq_bor. iDestruct "H" as (l') "[Hfb Hvs]". iAssert (κ0⊓κ' ⊑ κ0⊓κ)%I as "#Hκ0". { iApply lft_intersect_mono. iApply lft_incl_refl. done. } iExists l'. iSplit. by iApply (frac_bor_shorten with "[]"). - iIntros "!# * % Htok". iApply (step_fupd_mask_mono F _ (F∖↑shrN)); try solve_ndisj. + iIntros "!> * % Htok". iApply (step_fupd_mask_mono F _ (F∖↑shrN)); try solve_ndisj. iMod (lft_incl_acc with "Hκ0 Htok") as (q') "[Htok Hclose]"; first solve_ndisj. iMod ("Hvs" with "[%] Htok") as "Hvs'"; first solve_ndisj. iModIntro. iNext. iMod "Hvs'" as "[#Hshr Htok]". iMod ("Hclose" with "Htok") as "$". @@ -43,25 +43,40 @@ Section uniq_bor. Global Instance uniq_bor_wf κ ty `{!TyWf ty} : TyWf (uniq_bor κ ty) := { ty_lfts := [κ]; ty_wf_E := ty.(ty_wf_E) ++ ty_outlives_E ty κ }. - Global Instance uniq_mono E L : - Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) uniq_bor. + Lemma uniq_type_incl κ1 κ2 ty1 ty2 : + κ2 ⊑ κ1 -∗ + â–· type_equal ty1 ty2 -∗ + type_incl (uniq_bor κ1 ty1) (uniq_bor κ2 ty2). Proof. - intros κ1 κ2 Hκ ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (?) "HL". - iDestruct (Hty with "HL") as "#Hty". iDestruct (Hκ with "HL") as "#Hκ". - iIntros "!# #HE". iSplit; first done. - iDestruct ("Hty" with "HE") as "(_ & #Ho & #Hs)"; [done..|clear Hty]. - iSpecialize ("Hκ" with "HE"). iSplit; iModIntro. + iIntros "#Hlft #Hty". iSplit; first done. + iSplit; iModIntro. - iIntros (? [|[[]|][]]) "H"; try done. - iApply (bor_shorten with "Hκ"). iApply bor_iff; last done. - iNext. iModIntro. iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]"; - iExists vl; iFrame; by iApply "Ho". + iApply (bor_shorten with "Hlft"). iApply bor_iff; last done. + iNext. iModIntro. + iDestruct "Hty" as "(_ & Hty & _)". + iSplit; iIntros "H"; iDestruct "H" as (vl) "[??]"; + iExists vl; iFrame; by iApply "Hty". - iIntros (κ ??) "H". iAssert (κ2 ⊓ κ ⊑ κ1 ⊓ κ)%I as "#Hincl'". - { iApply lft_intersect_mono. done. iApply lft_incl_refl. } - iDestruct "H" as (l') "[Hbor #Hupd]". iExists l'. iIntros "{$Hbor}!# %%% Htok". + { iApply lft_intersect_mono; first done. iApply lft_incl_refl. } + iDestruct "H" as (l') "[Hbor #Hupd]". iExists l'. iIntros "{$Hbor}!> %%% Htok". iMod (lft_incl_acc with "Hincl' Htok") as (q') "[Htok Hclose]"; first solve_ndisj. iMod ("Hupd" with "[%] Htok") as "Hupd'"; try done. iModIntro. iNext. iMod "Hupd'" as "[H Htok]". iMod ("Hclose" with "Htok") as "$". - iApply ty_shr_mono; [done..|]. by iApply "Hs". + iDestruct "Hty" as "(_ & _ & Hty)". + iApply ty_shr_mono; last by iApply "Hty". + done. + Qed. + + Global Instance uniq_mono E L : + Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) uniq_bor. + Proof. + intros κ1 κ2 Hκ ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (??) "HL". + iDestruct (Hty with "HL") as "#Hty". iDestruct (Hκ with "HL") as "#Hκ". + iIntros "!> #HE". + iApply uniq_type_incl. + - iDestruct ("Hκ" with "HE") as %H. + apply lft_incl_syn_sem in H. iApply H. + - iNext. iApply "Hty". done. Qed. Global Instance uniq_mono_flip E L : Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) uniq_bor. @@ -103,15 +118,16 @@ Section typing. lctx_lft_incl E L κ2 κ1 → eqtype E L ty1 ty2 → subtype E L (&uniq{κ1}ty1) (&uniq{κ2}ty2). Proof. by intros; apply uniq_mono. Qed. - Lemma uniq_proper' E L κ ty1 ty2 : - eqtype E L ty1 ty2 → eqtype E L (&uniq{κ}ty1) (&uniq{κ}ty2). + Lemma uniq_proper' E L κ1 κ2 ty1 ty2 : + lctx_lft_eq E L κ1 κ2 → eqtype E L ty1 ty2 → eqtype E L (&uniq{κ1}ty1) (&uniq{κ2}ty2). Proof. by intros; apply uniq_proper. Qed. Lemma tctx_reborrow_uniq E L p ty κ κ' : lctx_lft_incl E L κ' κ → tctx_incl E L [p â— &uniq{κ}ty] [p â— &uniq{κ'}ty; p â—{κ'} &uniq{κ}ty]. Proof. - iIntros (Hκκ' tid ?) "#LFT HE HL H". iDestruct (Hκκ' with "HL HE") as "#Hκκ'". + iIntros (Hκκ' tid ??) "#LFT HE HL H". iDestruct (Hκκ' with "HL HE") as %H. + iDestruct (lft_incl_syn_sem κ' κ H) as "Hκκ'". iFrame. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iDestruct "H" as ([[]|]) "[% Hb]"; try done. iMod (rebor with "LFT Hκκ' Hb") as "[Hb Hext]". done. iModIntro. @@ -130,8 +146,8 @@ Section typing. Lemma read_uniq E L κ ty : Copy ty → lctx_lft_alive E L κ → ⊢ typed_read E L (&uniq{κ}ty) ty (&uniq{κ}ty). Proof. - rewrite typed_read_eq. iIntros (Hcopy Halive) "!#". - iIntros ([[]|] tid F qL ?) "#LFT #HE Htl HL Hown"; try done. + rewrite typed_read_eq. iIntros (Hcopy Halive) "!>". + iIntros ([[]|] tid F qmax qL ?) "#LFT #HE Htl HL Hown"; try done. iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first solve_ndisj. iMod (bor_acc with "LFT Hown Hκ") as "[H↦ Hclose']"; first solve_ndisj. iDestruct "H↦" as (vl) "[>H↦ #Hown]". @@ -144,8 +160,8 @@ Section typing. Lemma write_uniq E L κ ty : lctx_lft_alive E L κ → ⊢ typed_write E L (&uniq{κ}ty) ty (&uniq{κ}ty). Proof. - rewrite typed_write_eq. iIntros (Halive) "!#". - iIntros ([[]|] tid F qL ?) "#LFT HE HL Hown"; try done. + rewrite typed_write_eq. iIntros (Halive) "!>". + iIntros ([[]|] tid F qmax qL ?) "#LFT HE HL Hown"; try done. iMod (Halive with "HE HL") as (q) "[Htok Hclose]"; first solve_ndisj. iMod (bor_acc with "LFT Hown Htok") as "[H↦ Hclose']"; first solve_ndisj. iDestruct "H↦" as (vl) "[>H↦ Hown]". rewrite ty.(ty_size_eq). @@ -156,5 +172,5 @@ Section typing. Qed. End typing. -Hint Resolve uniq_mono' uniq_proper' write_uniq read_uniq : lrust_typing. -Hint Resolve tctx_extract_hasty_reborrow | 10 : lrust_typing. +Global Hint Resolve uniq_mono' uniq_proper' write_uniq read_uniq : lrust_typing. +Global Hint Resolve tctx_extract_hasty_reborrow | 10 : lrust_typing. diff --git a/theories/typing/util.v b/theories/typing/util.v index e9c5350c1fd7d00333a92f256888c58acf744ef7..424110d1b84983cc3ec297eb0c0c4a5825aeb113 100644 --- a/theories/typing/util.v +++ b/theories/typing/util.v @@ -35,7 +35,7 @@ Section util. iMod (inv_alloc shrN _ (idx_bor_own i V ∨ ty_shr ty κ tid l V)%I with "[Hpbown]") as "#Hinv". { iModIntro. eauto. } - iIntros "!> !# * % Htok". + iIntros "!> !> * % Htok". iMod (inv_acc with "Hinv") as "[INV Hclose]"; first solve_ndisj. iDestruct "INV" as "[>Hbtok|#Hshr]". - rewrite -(Qp_div_2 q). @@ -66,7 +66,7 @@ Section util. iMod (inv_alloc shrN _ (idx_bor_own i V ∨ ty_shr ty κ'' tid l V)%I with "[Hpbown]") as "#Hinv". { iModIntro. auto. } - iIntros "!> !# * % Htok". + iIntros "!> !> * % Htok". iMod (inv_acc with "Hinv") as "[INV Hclose]"; first solve_ndisj. iDestruct "INV" as "[>Hbtok|#Hshr]". - rewrite -(Qp_div_2 q).