From 9fda08ddc0dbe52e94bc249f6711f7145261ab58 Mon Sep 17 00:00:00 2001 From: Jacques-Henri Jourdan <jacques-henri.jourdan@cnrs.fr> Date: Sun, 30 Mar 2025 00:30:10 +0100 Subject: [PATCH] Bump Iris. --- _CoqProject | 21 - coq-lambda-rust.opam | 2 +- theories/lang/adequacy.v | 10 +- theories/lang/heap.v | 43 +- theories/lang/lang.v | 60 +- theories/lang/lib/arc.v | 672 --------- theories/lang/lib/spawn.v | 4 +- theories/lang/lifting.v | 93 +- theories/lang/proofmode.v | 57 +- theories/lang/races.v | 82 +- theories/lang/tactics.v | 8 +- theories/lang/time.v | 29 +- theories/lifetime/at_borrow.v | 6 +- theories/lifetime/frac_borrow.v | 14 +- theories/lifetime/lifetime_sig.v | 12 +- theories/lifetime/model/accessors.v | 28 +- theories/lifetime/model/borrow.v | 3 +- theories/lifetime/model/borrow_sep.v | 21 +- theories/lifetime/model/creation.v | 16 +- theories/lifetime/model/definitions.v | 30 +- theories/lifetime/model/faking.v | 4 +- theories/lifetime/model/primitive.v | 43 +- theories/lifetime/model/reborrow.v | 6 +- theories/prophecy/prophecy.v | 32 +- theories/prophecy/syn_type.v | 17 +- theories/typing/array.v | 6 +- theories/typing/array_idx.v | 2 +- theories/typing/array_util.v | 11 +- theories/typing/borrow.v | 3 +- theories/typing/examples/nonlexical.v | 132 -- theories/typing/fixpoint.v | 56 +- theories/typing/function.v | 20 +- theories/typing/lft_contexts.v | 18 +- theories/typing/lib/arc.v | 1202 ----------------- theories/typing/lib/cell.v | 13 +- theories/typing/lib/diverging_static.v | 55 - theories/typing/lib/fake_shared.v | 66 - theories/typing/lib/join.v | 96 -- theories/typing/lib/mutex/mutex.v | 12 +- theories/typing/lib/mutex/mutexguard.v | 6 +- theories/typing/lib/rc/rc.v | 1162 ---------------- theories/typing/lib/rc/weak.v | 504 ------- theories/typing/lib/refcell/ref.v | 125 -- theories/typing/lib/refcell/ref_code.v | 362 ----- theories/typing/lib/refcell/refcell.v | 212 --- theories/typing/lib/refcell/refcell_code.v | 319 ----- theories/typing/lib/refcell/refmut.v | 143 -- theories/typing/lib/refcell/refmut_code.v | 340 ----- theories/typing/lib/rwlock/rwlock.v | 226 ---- theories/typing/lib/rwlock/rwlock_code.v | 303 ----- theories/typing/lib/rwlock/rwlockreadguard.v | 152 --- .../typing/lib/rwlock/rwlockreadguard_code.v | 148 -- theories/typing/lib/rwlock/rwlockwriteguard.v | 162 --- .../typing/lib/rwlock/rwlockwriteguard_code.v | 148 -- theories/typing/lib/slice/slice_basic.v | 8 +- theories/typing/lib/slice/slice_split.v | 4 +- theories/typing/lib/smallvec/smallvec.v | 4 +- theories/typing/lib/smallvec/smallvec_basic.v | 6 +- theories/typing/lib/smallvec/smallvec_pop.v | 4 +- theories/typing/lib/smallvec/smallvec_push.v | 8 +- theories/typing/lib/take_mut.v | 77 -- theories/typing/lib/vec/vec_basic.v | 10 +- theories/typing/lib/vec/vec_pushpop.v | 2 +- theories/typing/lib/vec_util.v | 8 +- theories/typing/own.v | 12 +- theories/typing/product.v | 12 +- theories/typing/product_split.v | 7 +- theories/typing/programs.v | 16 +- theories/typing/shr_bor.v | 2 +- theories/typing/soundness.v | 12 +- theories/typing/sum.v | 8 +- theories/typing/type.v | 82 +- theories/typing/type_context.v | 7 +- theories/typing/type_sum.v | 4 +- theories/typing/uninit.v | 8 +- theories/typing/uniq_array_util.v | 20 +- theories/typing/uniq_bor.v | 4 +- theories/typing/uniq_cmra.v | 10 +- theories/util/basic.v | 6 +- theories/util/fancy_lists.v | 23 +- theories/util/update.v | 10 +- 81 files changed, 562 insertions(+), 7129 deletions(-) delete mode 100644 theories/lang/lib/arc.v delete mode 100644 theories/typing/examples/nonlexical.v delete mode 100644 theories/typing/lib/arc.v delete mode 100644 theories/typing/lib/diverging_static.v delete mode 100644 theories/typing/lib/fake_shared.v delete mode 100644 theories/typing/lib/join.v delete mode 100644 theories/typing/lib/rc/rc.v delete mode 100644 theories/typing/lib/rc/weak.v delete mode 100644 theories/typing/lib/refcell/ref.v delete mode 100644 theories/typing/lib/refcell/ref_code.v delete mode 100644 theories/typing/lib/refcell/refcell.v delete mode 100644 theories/typing/lib/refcell/refcell_code.v delete mode 100644 theories/typing/lib/refcell/refmut.v delete mode 100644 theories/typing/lib/refcell/refmut_code.v delete mode 100644 theories/typing/lib/rwlock/rwlock.v delete mode 100644 theories/typing/lib/rwlock/rwlock_code.v delete mode 100644 theories/typing/lib/rwlock/rwlockreadguard.v delete mode 100644 theories/typing/lib/rwlock/rwlockreadguard_code.v delete mode 100644 theories/typing/lib/rwlock/rwlockwriteguard.v delete mode 100644 theories/typing/lib/rwlock/rwlockwriteguard_code.v delete mode 100644 theories/typing/lib/take_mut.v diff --git a/_CoqProject b/_CoqProject index 603e4367..43d3a10e 100644 --- a/_CoqProject +++ b/_CoqProject @@ -39,7 +39,6 @@ theories/lang/lib/swap.v theories/lang/lib/new_delete.v theories/lang/lib/spawn.v theories/lang/lib/lock.v -theories/lang/lib/arc.v theories/lang/lib/tests.v theories/typing/base.v theories/typing/uniq_cmra.v @@ -95,30 +94,11 @@ theories/typing/lib/smallvec/smallvec_slice.v theories/typing/lib/smallvec/smallvec_index.v theories/typing/lib/smallvec/smallvec_push.v theories/typing/lib/smallvec/smallvec_pop.v -# theories/typing/lib/fake_shared.v theories/typing/lib/cell.v theories/typing/lib/spawn.v -# theories/typing/lib/join.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/mutex/mutex.v theories/typing/lib/mutex/mutexguard.v -# theories/typing/lib/refcell/refcell.v -# theories/typing/lib/refcell/ref.v -# theories/typing/lib/refcell/refmut.v -# theories/typing/lib/refcell/refcell_code.v -# theories/typing/lib/refcell/ref_code.v -# theories/typing/lib/refcell/refmut_code.v -# theories/typing/lib/rwlock/rwlock.v -# theories/typing/lib/rwlock/rwlockreadguard.v -# theories/typing/lib/rwlock/rwlockwriteguard.v -# theories/typing/lib/rwlock/rwlock_code.v -# theories/typing/lib/rwlock/rwlockreadguard_code.v -# theories/typing/lib/rwlock/rwlockwriteguard_code.v theories/typing/examples/inc_max.v theories/typing/examples/odd_sum.v theories/typing/examples/projection_toggle.v @@ -133,5 +113,4 @@ theories/typing/examples/rebor.v theories/typing/examples/unbox.v theories/typing/examples/init_prod.v theories/typing/examples/lazy_lft.v -# theories/typing/examples/nonlexical.v theories/typing/examples/derived_rules.v \ No newline at end of file diff --git a/coq-lambda-rust.opam b/coq-lambda-rust.opam index aa37621f..2916da23 100644 --- a/coq-lambda-rust.opam +++ b/coq-lambda-rust.opam @@ -13,7 +13,7 @@ it possible to easily prove the functionnal correctness of well-typed program. """ depends: [ - "coq-iris" { (= "dev.2022-04-12.0.a3bed7ea") | (= "dev") } + "coq-iris" { (= "dev.2025-03-25.0.79d33e24") | (= "dev") } ] build: [make "-j%{jobs}%"] diff --git a/theories/lang/adequacy.v b/theories/lang/adequacy.v index cc7600f2..2a82674c 100644 --- a/theories/lang/adequacy.v +++ b/theories/lang/adequacy.v @@ -6,10 +6,10 @@ From lrust.lang Require Import proofmode notation. Set Default Proof Using "Type". Class lrustGpreS Σ := HeapGpreS { - lrustGpreS_iris :> invGpreS Σ; - lrustGpreS_heap :> inG Σ (authR heapUR); - lrustGpreS_heap_freeable :> inG Σ (authR heap_freeableUR); - lrustGpreS_time :> timePreG Σ + lrustGpreS_iris :: invGpreS Σ; + lrustGpreS_heap :: inG Σ (authR heapUR); + lrustGpreS_heap_freeable :: inG Σ (authR heap_freeableUR); + lrustGpreS_time :: timePreG Σ }. Definition lrustΣ : gFunctors := @@ -31,7 +31,7 @@ Proof. first by apply auth_auth_valid. iMod time_init as (Htime) "[TIME Htime]"; [done|]. set (Hheap := HeapGS _ _ _ vγ fγ). - iModIntro. iExists NotStuck, _, [_], _, _. simpl. + iModIntro. iExists _, [_], _, _. simpl. iDestruct (Hwp (LRustGS _ _ Hheap Htime) with "TIME") as "$". iSplitL; first by auto with iFrame. iIntros ([|e' [|]]? -> ??) "//". iIntros "[??] [?_] _". iApply fupd_mask_weaken; [|iIntros "_ !>"]; [done|]. diff --git a/theories/lang/heap.v b/theories/lang/heap.v index 80758c14..249af9a5 100644 --- a/theories/lang/heap.v +++ b/theories/lang/heap.v @@ -1,4 +1,3 @@ -From Coq Require Import Min. From stdpp Require Import coPset. From iris.algebra Require Import big_op gmap frac agree numbers. From iris.algebra Require Import csum excl auth cmra_big_op. @@ -19,8 +18,8 @@ Definition heap_freeableUR : ucmra := gmapUR block (prodR fracR (gmapR Z (exclR unitO))). Class heapGS Σ := HeapGS { - heap_inG :> inG Σ (authR heapUR); - heap_freeable_inG :> inG Σ (authR heap_freeableUR); + heap_inG :: inG Σ (authR heapUR); + heap_freeable_inG :: inG Σ (authR heap_freeableUR); heap_name : gname; heap_freeable_name : gname }. @@ -91,7 +90,7 @@ Section to_heap. Implicit Types σ : state. Lemma to_heap_valid σ : ✓ to_heap σ. - Proof. intros l. rewrite lookup_fmap. case (σ !! l)=> [[[|n] v]|] //=. Qed. + Proof. intros l. rewrite lookup_fmap. destruct (σ !! l) as [[[|n] v]|] eqn:EQ; rewrite EQ //. Qed. Lemma lookup_to_heap_None σ l : σ !! l = None → to_heap σ !! l = None. Proof. by rewrite /to_heap lookup_fmap=> ->. Qed. @@ -123,10 +122,10 @@ Section heap. Global Instance heap_mapsto_as_fractional l q v: AsFractional (l ↦{q} v) (λ q, l ↦{q} v)%I q. Proof. split; first done. apply _. Qed. - Global Instance frame_heap_mapsto p l v q1 q2 RES : - FrameFractionalHyps p (l ↦{q1} v) (λ q, l ↦{q} v)%I RES q1 q2 → - Frame p (l ↦{q1} v) (l ↦{q2} v) RES | 5. - Proof. apply: frame_fractional. Qed. + Global Instance frame_heap_mapsto p l v q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p (l ↦{q1} v) (l ↦{q2} v) (l ↦{q} v) | 5. + Proof. apply: (frame_fractional (λ q, l ↦{q} v)%I). Qed. Global Instance heap_mapsto_vec_timeless l q vl : Timeless (l ↦∗{q} vl). Proof. rewrite /heap_mapsto_vec. apply _. Qed. @@ -139,10 +138,10 @@ Section heap. Global Instance heap_mapsto_vec_as_fractional l q vl: AsFractional (l ↦∗{q} vl) (λ q, l ↦∗{q} vl)%I q. Proof. split; first done. apply _. Qed. - Global Instance frame_heap_mapsto_vec p l vl q1 q2 RES : - FrameFractionalHyps p (l ↦∗{q1} vl) (λ q, l ↦∗{q} vl)%I RES q1 q2 → - Frame p (l ↦∗{q1} vl) (l ↦∗{q2} vl) RES | 5. - Proof. apply: frame_fractional. Qed. + Global Instance frame_heap_mapsto_vec p l vl q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p (l ↦∗{q1} vl) (l ↦∗{q2} vl) (l ↦∗{q} vl) | 5. + Proof. apply: (frame_fractional (λ q, l ↦∗{q} vl)%I). Qed. Global Instance heap_freeable_timeless q l n : Timeless (†{q}l…n). Proof. rewrite heap_freeable_eq /heap_freeable_def. apply _. Qed. @@ -238,10 +237,10 @@ Section heap. rewrite -[vl1](firstn_skipn ll) -[vl2](firstn_skipn ll) 2!heap_mapsto_vec_app. iDestruct "Hown1" as "[Hown1 _]". iDestruct "Hown2" as "[Hown2 _]". iCombine "Hown1" "Hown2" as "Hown". rewrite heap_mapsto_vec_op; last first. - { rewrite !firstn_length. subst ll. - rewrite -!min_assoc min_idempotent min_comm -min_assoc min_idempotent min_comm. done. } + { rewrite !length_firstn. subst ll. + rewrite -!Nat.min_assoc Nat.min_idempotent Nat.min_comm -Nat.min_assoc Nat.min_idempotent Nat.min_comm. done. } iDestruct "Hown" as "[H Hown]". iDestruct "H" as %Hl. iExists (take ll vl1). iFrame. - destruct (min_spec (length vl1) (length vl2)) as [[Hne Heq]|[Hne Heq]]. + destruct (Nat.min_spec (length vl1) (length vl2)) as [[Hne Heq]|[Hne Heq]]. + iClear "HP2". rewrite take_ge; last first. { rewrite -Heq /ll. done. } rewrite drop_ge; first by rewrite app_nil_r. by rewrite -Heq. @@ -277,7 +276,7 @@ Section heap. Proof. revert i. induction n as [|n IH]=>i; first done. by apply insert_valid. Qed. Lemma heap_freeable_op_eq l q1 q2 n n' : - †{q1}l…n ∗ †{q2}l+ₗn … n' ⊣⊢ †{q1+q2}l…(n+n'). + †{q1}l…n ∗ †{q2}l+ₗ n … n' ⊣⊢ †{q1+q2}l…(n+n'). Proof. by rewrite heap_freeable_eq /heap_freeable_def -own_op -auth_frag_op singleton_op -pair_op inter_op. @@ -351,7 +350,7 @@ Section heap. ∗ own heap_name (◯ [^op list] i ↦ v ∈ (repeat (LitV LitPoison) n), {[l +ₗ i := (1%Qp, Cinr 0%nat, to_agree v)]}). Proof. - intros FREE. rewrite -own_op. apply own_update, auth_update_alloc. + intros FREE. rewrite -own_op own_update; [by auto|]. apply auth_update_alloc. revert l FREE. induction n as [|n IH]=> l FRESH; [done|]. rewrite (big_opL_consZ_l (λ k _, _ (_ k) _ )) /=. etrans; first apply (IH (l +ₗ 1)). @@ -392,7 +391,7 @@ Section heap. {[l +ₗ i := (1%Qp, Cinr 0%nat, to_agree v)]}) ==∗ own heap_name (● to_heap (free_mem l (length vl) σ)). Proof. - rewrite -own_op. apply own_update, auth_update_dealloc. + rewrite -own_op own_update; [by auto|]. apply auth_update_dealloc. revert σ l. induction vl as [|v vl IH]=> σ l; [done|]. rewrite (big_opL_consZ_l (λ k _, _ (_ k) _ )) /= shift_loc_0. apply local_update_total_valid=> _ Hvalid _. @@ -472,8 +471,8 @@ Section heap. ==∗ own heap_name (● to_heap (<[l:=(RSt (n2 + nf), v)]> σ)) ∗ heap_mapsto_st (RSt n2) l q v. Proof. - intros Hσv. apply wand_intro_r. rewrite -!own_op to_heap_insert. - eapply own_update, auth_update, singleton_local_update. + intros Hσv. apply entails_wand, wand_intro_r. rewrite -!own_op to_heap_insert. + iApply own_update. eapply auth_update, singleton_local_update. { by rewrite /to_heap lookup_fmap Hσv. } apply prod_local_update_1, prod_local_update_2, csum_local_update_r. apply nat_local_update; lia. @@ -521,8 +520,8 @@ Section heap. ==∗ own heap_name (● to_heap (<[l:=(st2, v')]> σ)) ∗ heap_mapsto_st st2 l 1%Qp v'. Proof. - intros Hσv. apply wand_intro_r. rewrite -!own_op to_heap_insert. - eapply own_update, auth_update, singleton_local_update. + intros Hσv. apply entails_wand, wand_intro_r. rewrite -!own_op to_heap_insert. + iApply own_update. eapply auth_update, singleton_local_update. { by rewrite /to_heap lookup_fmap Hσv. } apply exclusive_local_update. by destruct st2. Qed. diff --git a/theories/lang/lang.v b/theories/lang/lang.v index d2b59983..ff962f36 100644 --- a/theories/lang/lang.v +++ b/theories/lang/lang.v @@ -44,8 +44,8 @@ Inductive expr := | Case (e : expr) (el : list expr) | Fork (e : expr). -Arguments App _%E _%E. -Arguments Case _%E _%E. +Arguments App _%_E _%_E. +Arguments Case _%_E _%_E. Fixpoint is_closed (X : list string) (e : expr) : bool := match e with @@ -153,12 +153,12 @@ Fixpoint subst_l (xl : list binder) (esl : list expr) (e : expr) : option expr : | x::xl, es::esl => subst' x es <$> subst_l xl esl e | _, _ => None end. -Arguments subst_l _%binder _ _%E. +Arguments subst_l _%_binder _ _%_E. Definition subst_v (xl : list binder) (vsl : vec val (length xl)) (e : expr) : expr := Vector.fold_right2 (λ b, subst' b ∘ of_val) e _ (list_to_vec xl) vsl. -Arguments subst_v _%binder _ _%E. +Arguments subst_v _%_binder _ _%_E. Lemma subst_v_eq (xl : list binder) (vsl : vec val (length xl)) e : Some $ subst_v xl vsl e = subst_l xl (of_val <$> vec_to_list vsl) e. @@ -238,50 +238,50 @@ Inductive bin_op_eval (σ : state) : bin_op → base_lit → base_lit → base_l Notation stuck_term := (App (Lit (LitInt 0)) []). -Inductive head_step : expr → state → list Empty_set → expr → state → list expr → Prop := +Inductive base_step : expr → state → list Empty_set → expr → state → list expr → Prop := | BinOpS op l1 l2 l' σ : bin_op_eval σ op l1 l2 l' → - head_step (BinOp op (Lit l1) (Lit l2)) σ [] (Lit l') σ [] + base_step (BinOp op (Lit l1) (Lit l2)) σ [] (Lit l') σ [] | NdIntS z σ : - head_step NdInt σ [] (Lit (LitInt z)) σ [] + base_step NdInt σ [] (Lit (LitInt z)) σ [] | BetaS f xl e e' el σ: Forall (λ ei, is_Some (to_val ei)) el → Closed (f :b: xl +b+ []) e → subst_l (f::xl) (Rec f xl e :: el) e = Some e' → - head_step (App (Rec f xl e) el) σ [] e' σ [] + base_step (App (Rec f xl e) el) σ [] e' σ [] | ReadScS l n v σ: σ !! l = Some (RSt n, v) → - head_step (Read ScOrd (Lit $ LitLoc l)) σ [] (of_val v) σ [] + base_step (Read ScOrd (Lit $ LitLoc l)) σ [] (of_val v) σ [] | ReadNa1S l n v σ: σ !! l = Some (RSt n, v) → - head_step (Read Na1Ord (Lit $ LitLoc l)) σ + base_step (Read Na1Ord (Lit $ LitLoc l)) σ [] (Read Na2Ord (Lit $ LitLoc l)) (<[l:=(RSt $ S n, v)]>σ) [] | ReadNa2S l n v σ: σ !! l = Some (RSt $ S n, v) → - head_step (Read Na2Ord (Lit $ LitLoc l)) σ + base_step (Read Na2Ord (Lit $ LitLoc l)) σ [] (of_val v) (<[l:=(RSt n, v)]>σ) [] | WriteScS l e v v' σ: to_val e = Some v → σ !! l = Some (RSt 0, v') → - head_step (Write ScOrd (Lit $ LitLoc l) e) σ + base_step (Write ScOrd (Lit $ LitLoc l) e) σ [] (Lit LitPoison) (<[l:=(RSt 0, v)]>σ) [] | WriteNa1S l e v v' σ: to_val e = Some v → σ !! l = Some (RSt 0, v') → - head_step (Write Na1Ord (Lit $ LitLoc l) e) σ + base_step (Write Na1Ord (Lit $ LitLoc l) e) σ [] (Write Na2Ord (Lit $ LitLoc l) e) (<[l:=(WSt, v')]>σ) [] | WriteNa2S l e v v' σ: to_val e = Some v → σ !! l = Some (WSt, v') → - head_step (Write Na2Ord (Lit $ LitLoc l) e) σ + base_step (Write Na2Ord (Lit $ LitLoc l) e) σ [] (Lit LitPoison) (<[l:=(RSt 0, v)]>σ) [] @@ -289,12 +289,12 @@ Inductive head_step : expr → state → list Empty_set → expr → state → l to_val e1 = Some $ LitV lit1 → to_val e2 = Some $ LitV lit2 → σ !! l = Some (RSt n, LitV litl) → lit_neq lit1 litl → - head_step (CAS (Lit $ LitLoc l) e1 e2) σ [] (Lit $ lit_of_bool false) σ [] + base_step (CAS (Lit $ LitLoc l) e1 e2) σ [] (Lit $ lit_of_bool false) σ [] | CasSucS l e1 lit1 e2 lit2 litl σ : to_val e1 = Some $ LitV lit1 → to_val e2 = Some $ LitV lit2 → σ !! l = Some (RSt 0, LitV litl) → lit_eq σ lit1 litl → - head_step (CAS (Lit $ LitLoc l) e1 e2) σ + base_step (CAS (Lit $ LitLoc l) e1 e2) σ [] (Lit $ lit_of_bool true) (<[l:=(RSt 0, LitV lit2)]>σ) [] @@ -316,30 +316,30 @@ Inductive head_step : expr → state → list Empty_set → expr → state → l to_val e1 = Some $ LitV lit1 → to_val e2 = Some $ LitV lit2 → σ !! l = Some (RSt n, LitV litl) → 0 < n → lit_eq σ lit1 litl → - head_step (CAS (Lit $ LitLoc l) e1 e2) σ + base_step (CAS (Lit $ LitLoc l) e1 e2) σ [] stuck_term σ [] | AllocS n l σ : 0 < n → (∀ m, σ !! (l +ₗ m) = None) → - head_step (Alloc $ Lit $ LitInt n) σ + base_step (Alloc $ Lit $ LitInt n) σ [] (Lit $ LitLoc l) (init_mem l (Z.to_nat n) σ) [] | FreeS n l σ : 0 < n → (∀ m, is_Some (σ !! (l +ₗ m)) ↔ 0 ≤ m < n) → - head_step (Free (Lit $ LitInt n) (Lit $ LitLoc l)) σ + base_step (Free (Lit $ LitInt n) (Lit $ LitLoc l)) σ [] (Lit LitPoison) (free_mem l (Z.to_nat n) σ) [] | CaseS i el e σ : 0 ≤ i → el !! (Z.to_nat i) = Some e → - head_step (Case (Lit $ LitInt i) el) σ [] e σ [] + base_step (Case (Lit $ LitInt i) el) σ [] e σ [] | ForkS e σ: - head_step (Fork e) σ [] (Lit LitPoison) σ [e]. + base_step (Fork e) σ [] (Lit LitPoison) σ [e]. (** Basic properties about the language *) Lemma to_of_val v : to_val (of_val v) = Some v. @@ -363,11 +363,11 @@ Lemma fill_item_val Ki e : Proof. intros [v ?]. destruct Ki; simplify_option_eq; eauto. Qed. Lemma val_stuck e1 σ1 κ e2 σ2 ef : - head_step e1 σ1 κ e2 σ2 ef → to_val e1 = None. + base_step e1 σ1 κ e2 σ2 ef → to_val e1 = None. Proof. destruct 1; naive_solver. Qed. Lemma head_ctx_step_val Ki e σ1 κ e2 σ2 ef : - head_step (fill_item Ki e) σ1 κ e2 σ2 ef → is_Some (to_val e). + base_step (fill_item Ki e) σ1 κ e2 σ2 ef → is_Some (to_val e). Proof. destruct Ki; inversion_clear 1; decompose_Forall_hyps; simplify_option_eq; by eauto. @@ -409,7 +409,7 @@ Lemma shift_loc_0_nat l : l +ₗ 0%nat = l. Proof. destruct l as [b o]. rewrite /shift_loc /=. f_equal. lia. Qed. Global Instance shift_loc_inj l : Inj (=) (=) (shift_loc l). -Proof. destruct l as [b o]; intros n n' [=?]; lia. Qed. +Proof. destruct l as [b o]; intros n n' [=]; lia. Qed. Lemma shift_loc_block l n : (l +ₗ n).1 = l.1. Proof. done. Qed. @@ -436,7 +436,7 @@ Proof. Qed. Definition fresh_block (σ : state) : block := - let loclst : list loc := elements (dom _ σ : gset loc) in + let loclst : list loc := elements (dom σ : gset loc) in let blockset : gset block := foldr (λ l, ({[l.1]} ∪.)) ∅ loclst in fresh blockset. @@ -453,7 +453,7 @@ Lemma alloc_fresh n σ : let l := (fresh_block σ, 0) in let init := repeat (LitV $ LitInt 0) (Z.to_nat n) in 0 < n → - head_step (Alloc $ Lit $ LitInt n) σ [] (Lit $ LitLoc l) (init_mem l (Z.to_nat n) σ) []. + base_step (Alloc $ Lit $ LitInt n) σ [] (Lit $ LitLoc l) (init_mem l (Z.to_nat n) σ) []. Proof. intros l init Hn. apply AllocS. auto. - intros i. apply (is_fresh_block _ i). @@ -537,8 +537,8 @@ Proof. Qed. (* Misc *) -Lemma stuck_not_head_step σ e' κ σ' ef : - ¬head_step stuck_term σ e' κ σ' ef. +Lemma stuck_not_base_step σ e' κ σ' ef : + ¬base_step stuck_term σ e' κ σ' ef. Proof. inversion 1. Qed. (** Equality and other typeclass stuff *) @@ -609,7 +609,7 @@ Canonical Structure valO := leibnizO val. Canonical Structure exprO := leibnizO expr. (** Language *) -Lemma lrust_lang_mixin : EctxiLanguageMixin of_val to_val fill_item head_step. +Lemma lrust_lang_mixin : EctxiLanguageMixin of_val to_val fill_item base_step. Proof. split; apply _ || eauto using to_of_val, of_to_val, val_stuck, fill_item_val, fill_item_no_val_inj, head_ctx_step_val. @@ -622,7 +622,7 @@ Canonical Structure lrust_lang := LanguageOfEctx lrust_ectx_lang. Lemma stuck_irreducible K σ : irreducible (fill K stuck_term) σ. Proof. apply: (irreducible_fill (K:=ectx_language.fill K)); first done. - apply prim_head_irreducible. + apply prim_base_irreducible. - inversion 1. - apply ectxi_language_sub_redexes_are_values. intros [] ??; simplify_eq/=; eauto; discriminate_list. diff --git a/theories/lang/lib/arc.v b/theories/lang/lib/arc.v deleted file mode 100644 index 63d8295c..00000000 --- a/theories/lang/lib/arc.v +++ /dev/null @@ -1,672 +0,0 @@ -From iris.base_logic.lib Require Import invariants. -From iris.program_logic Require Import weakestpre. -From iris.proofmode Require Import proofmode. -From iris.bi Require Import fractional. -From iris.algebra Require Import excl csum frac auth numbers. -From lrust.lang Require Import lang proofmode notation new_delete. -From iris.prelude Require Import options. - -(* JH: while working on Arc, I think figured that the "weak count -locking" mechanism that Rust is using and that is verified below may -not be necessary. - -Instead, what can be done in get_mut is the following: - - First, do a CAS on the strong count to put it to 0 from the expected -value 1. This has the effect, at the same time, of ensuring that no one -else has a strong reference, and of forbidding other weak references to -be upgraded (since the strong count is now at 0). If the CAS fail, -simply return None. - - Second, check the weak count. If it is at 1, then we are sure that -we are the only owner. - - Third, in any case write 1 in the strong count to cancel the CAS. - -What's wrong with this protocol? The "only" problem I can see is that if -someone tries to upgrade a weak after we did the CAS, then this will -fail even though this could be possible. - -RJ: Upgrade failing spuriously sounds like a problem severe enough to -justify the locking protocol. -*) - -Definition strong_count : val := - λ: ["l"], !ˢᶜ"l". - -Definition weak_count : val := - λ: ["l"], !ˢᶜ("l" +ₗ #1). - -Definition clone_arc : val := - rec: "clone" ["l"] := - let: "strong" := !ˢᶜ"l" in - if: CAS "l" "strong" ("strong" + #1) then #☠ else "clone" ["l"]. - -Definition clone_weak : val := - rec: "clone" ["l"] := - let: "weak" := !ˢᶜ("l" +ₗ #1) in - if: CAS ("l" +ₗ #1) "weak" ("weak" + #1) then #☠ else "clone" ["l"]. - -Definition downgrade : val := - rec: "downgrade" ["l"] := - let: "weak" := !ˢᶜ("l" +ₗ #1) in - if: "weak" = #(-1) then - (* -1 in weak indicates that somebody locked the counts; let's spin. *) - "downgrade" ["l"] - else - if: CAS ("l" +ₗ #1) "weak" ("weak" + #1) then #☠ - else "downgrade" ["l"]. - -Definition upgrade : val := - rec: "upgrade" ["l"] := - let: "strong" := !ˢᶜ"l" in - if: "strong" = #0 then #false - else - if: CAS "l" "strong" ("strong" + #1) then #true - else "upgrade" ["l"]. - -Definition drop_weak : val := - rec: "drop" ["l"] := - (* This can ignore the lock because the lock is only held when there - are no other weak refs in existence, so this code will never be called - with the lock held. *) - let: "weak" := !ˢᶜ("l" +ₗ #1) in - if: CAS ("l" +ₗ #1) "weak" ("weak" - #1) then "weak" = #1 - else "drop" ["l"]. - -Definition drop_arc : val := - rec: "drop" ["l"] := - let: "strong" := !ˢᶜ"l" in - if: CAS "l" "strong" ("strong" - #1) then "strong" = #1 - else "drop" ["l"]. - -Definition try_unwrap : val := - λ: ["l"], CAS "l" #1 #0. - -Definition is_unique : val := - λ: ["l"], - (* Try to acquire the lock for the last ref by CASing weak count from 1 to -1. *) - if: CAS ("l" +ₗ #1) #1 #(-1) then - let: "strong" := !ˢᶜ"l" in - let: "unique" := "strong" = #1 in - "l" +ₗ #1 <-ˢᶜ #1;; - "unique" - else - #false. - -Definition try_unwrap_full : val := - λ: ["l"], - if: CAS "l" #1 #0 then - if: !ˢᶜ("l" +ₗ #1) = #1 then "l" <- #1;; #0 - else #1 - else #2. - -(** The CMRA we need. *) -(* Not bundling heapGS, as it may be shared with other users. *) - -(* See rc.v for understanding the structure of this CMRA. - The only additional thing is the [optionR (exclR unitO))], used to handle - properly the locking mechanisme for the weak count. *) -Definition arc_stR := - prodUR (optionUR (csumR (prodR (prodR fracR positiveR) (optionR (exclR unitO))) - (exclR unitO))) natUR. -Class arcG Σ := - RcG :> inG Σ (authR arc_stR). -Definition arcΣ : gFunctors := #[GFunctor (authR arc_stR)]. -Global Instance subG_arcΣ {Σ} : subG arcΣ Σ → arcG Σ. -Proof. solve_inG. Qed. - -Section def. - Context `{!lrustGS Σ, !arcG Σ} (P1 : Qp → iProp Σ) (P2 : iProp Σ) (N : namespace). - - Definition arc_tok γ q : iProp Σ := - own γ (◯ (Some $ Cinl (q, 1%positive, None), 0%nat)). - Definition weak_tok γ : iProp Σ := - own γ (◯ (None, 1%nat)). - - Global Instance arc_tok_timeless γ q : Timeless (arc_tok γ q) := _. - Global Instance weak_tok_timeless γ : Timeless (weak_tok γ) := _. - - Definition arc_inv (γ : gname) (l : loc) : iProp Σ := - (∃ st : arc_stR, own γ (● st) ∗ - match st with - | (Some (Cinl (q, strong, wlock)), weak) => - ∃ q', ⌜(q + q')%Qp = 1%Qp⌝ ∗ P1 q' ∗ l ↦ #(Zpos strong) ∗ - if wlock then (l +ₗ 1) ↦ #(-1) ∗ ⌜weak = 0%nat⌝ - else (l +ₗ 1) ↦ #(S weak) - | (Some (Cinr _), weak) => - l ↦ #0 ∗ (l +ₗ 1) ↦ #(S weak) - | (None, (S _) as weak) => - l ↦ #0 ∗ (l +ₗ 1) ↦ #weak ∗ P2 - | _ => True - end)%I. - - Definition is_arc (γ : gname) (l : loc) : iProp Σ := - inv N (arc_inv γ l). - - Global Instance is_arc_persistent γ l : Persistent (is_arc γ l) := _. - - Definition arc_tok_acc (γ : gname) P E : iProp Σ := - (□ (P ={E,∅}=∗ ∃ q, arc_tok γ q ∗ (arc_tok γ q ={∅,E}=∗ P)))%I. - - Definition weak_tok_acc (γ : gname) P E : iProp Σ := - (□ (P ={E,∅}=∗ weak_tok γ ∗ (weak_tok γ ={∅,E}=∗ P)))%I. - - Definition drop_spec (drop : val) P : iProp Σ := - ({{{ P ∗ P1 1 }}} drop [] {{{ RET #☠; P ∗ P2 }}})%I. -End def. - -Section arc. - (* P1 is the thing that is shared by strong reference owners (in practice, - this is the lifetime token), and P2 is the thing that is owned by the - protocol when only weak references are left (in practice, P2 is the - ownership of the underlying memory incl. deallocation). *) - Context `{!lrustGS Σ, !arcG Σ} (P1 : Qp → iProp Σ) {HP1:Fractional P1} - (P2 : iProp Σ) (N : namespace). - - Local Instance P1_AsFractional q : AsFractional (P1 q) P1 q. - Proof using HP1. done. Qed. - - Global Instance arc_inv_ne n : - Proper (pointwise_relation _ (dist n) ==> dist n ==> eq ==> eq ==> dist n) - arc_inv. - Proof. solve_proper. Qed. - Global Instance arc_inv_proper : - Proper (pointwise_relation _ (≡) ==> (≡) ==> eq ==> eq ==> (≡)) - arc_inv. - Proof. solve_proper. Qed. - - Global Instance is_arc_contractive n : - Proper (pointwise_relation _ (dist_later n) ==> dist_later n ==> eq ==> eq ==> eq ==> dist n) - is_arc. - Proof. solve_contractive. Qed. - Global Instance is_arc_proper : - Proper (pointwise_relation _ (≡) ==> (≡) ==> eq ==> eq ==> eq ==> (≡)) is_arc. - Proof. solve_proper. Qed. - - Lemma create_arc E l : - l ↦ #1 -∗ (l +ₗ 1) ↦ #1 -∗ P1 1%Qp ={E}=∗ - ∃ γ q, is_arc P1 P2 N γ l ∗ P1 q ∗ arc_tok γ q. - Proof using HP1. - iIntros "Hl1 Hl2 [HP1 HP1']". - iMod (own_alloc ((● (Some $ Cinl ((1/2)%Qp, xH, None), O) ⋅ - ◯ (Some $ Cinl ((1/2)%Qp, xH, None), O)))) - as (γ) "[H● H◯]"; first by apply auth_both_valid_discrete. - iExists _, _. iFrame. iApply inv_alloc. iExists _. iFrame. iExists _. iFrame. - rewrite Qp_div_2. auto. - Qed. - - Lemma create_weak E l : - l ↦ #0 -∗ (l +ₗ 1) ↦ #1 -∗ P2 ={E}=∗ ∃ γ, is_arc P1 P2 N γ l ∗ weak_tok γ. - Proof. - iIntros "Hl1 Hl2 HP2". - iMod (own_alloc ((● (None, 1%nat) ⋅ ◯ (None, 1%nat)))) as (γ) "[H● H◯]"; - first by apply auth_both_valid_discrete. - iExists _. iFrame. iApply inv_alloc. iExists _. iFrame. - Qed. - - Lemma arc_tok_auth_val st γ q : - own γ (● st) -∗ arc_tok γ q -∗ - ⌜∃ q' strong wlock weak, st = (Some $ Cinl (q', strong, wlock), weak) ∧ - if decide (strong = xH) then q = q' ∧ wlock = None - else ∃ q'', q' = (q + q'')%Qp⌝. - Proof. - iIntros "H● Htok". iDestruct (own_valid_2 with "H● Htok") as - %[[Hincl%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete. - destruct st, Hincl as [[=]|(?&?&[= <-]&?&[Hincl|Hincl%csum_included])]; - simpl in *; subst. - - setoid_subst. iExists _, _, _, _. by iSplit. - - destruct Hincl as [->|[(?&[[??]?]&[=<-]&->&[[[??]%frac_included%Qp_lt_sum - ?%pos_included]%prod_included _]%prod_included)|(?&?&[=]&?)]]; first done. - iExists _, _, _, _. iSplit=>//. simpl in *. destruct decide; [subst;lia|eauto]. - Qed. - - Lemma strong_count_spec (γ : gname) (l : loc) (P : iProp Σ) : - is_arc P1 P2 N γ l -∗ arc_tok_acc γ P (⊤ ∖ ↑N) -∗ - {{{ P }}} strong_count [ #l] {{{ (c : nat), RET #c; P ∗ ⌜(c > 0)%nat⌝ }}}. - Proof using HP1. - iIntros "#INV #Hacc !> * HP HΦ". iLöb as "IH". wp_rec. - iInv N as (st) "[>H● H]" "Hclose1". - iMod ("Hacc" with "HP") as (?) "[Hown Hclose2]". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?& strong &?&?&[-> _]). - iDestruct "H" as (?) "(Hq & HP1 & H↦ & Hw)". wp_read. - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iMod ("Hclose1" with "[-HP HΦ]") as "_"; [iExists _; auto with iFrame|]. - iModIntro. rewrite -positive_nat_Z. iApply "HΦ". auto with iFrame lia. - Qed. - - (* FIXME : in the case the weak count is locked, we can possibly - return -1. This problem already exists in Rust. *) - Lemma weak_count_spec (γ : gname) (l : loc) (P : iProp Σ) : - is_arc P1 P2 N γ l -∗ arc_tok_acc γ P (⊤ ∖ ↑N) -∗ - {{{ P }}} weak_count [ #l] {{{ (c : Z), RET #c; P ∗ ⌜c >= -1⌝ }}}. - Proof using HP1. - iIntros "#INV #Hacc !> * HP HΦ". iLöb as "IH". wp_rec. wp_op. - iInv N as (st) "[>H● H]" "Hclose1". - iMod ("Hacc" with "HP") as (?) "[Hown Hclose2]". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?& strong &wl&?&[-> _]). - iDestruct "H" as (?) "(Hq & HP1 & H↦ & Hw)". destruct wl. - - iDestruct "Hw" as ">[Hw %]". wp_read. - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iMod ("Hclose1" with "[-HP HΦ]") as "_"; [iExists _; auto with iFrame|]. - iApply "HΦ". auto with iFrame lia. - - wp_read. iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iMod ("Hclose1" with "[-HP HΦ]") as "_"; [iExists _; auto with iFrame|]. - iApply "HΦ". auto with iFrame lia. - Qed. - - Lemma clone_arc_spec (γ : gname) (l : loc) (P : iProp Σ) : - is_arc P1 P2 N γ l -∗ arc_tok_acc γ P (⊤ ∖ ↑N) -∗ - {{{ P }}} clone_arc [ #l] - {{{ q', RET #☠; P ∗ arc_tok γ q' ∗ P1 q' }}}. - Proof using HP1. - iIntros "#INV #Hacc !> * HP HΦ". iLöb as "IH". wp_rec. wp_bind (!ˢᶜ_)%E. - iInv N as (st) "[>H● H]" "Hclose1". - iMod ("Hacc" with "HP") as (?) "[Hown Hclose2]". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?& strong &?&?&[-> _]). - iDestruct "H" as (?) "(Hq & HP1 & H↦ & Hw)". wp_read. - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iMod ("Hclose1" with "[-HP HΦ]") as "_"; [iExists _; auto with iFrame|]. - iModIntro. wp_let. wp_op. wp_bind (CAS _ _ _). iInv N as (st) "[>H● H]" "Hclose1". - iMod ("Hacc" with "HP") as (?) "[Hown Hclose2]". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?&strong'&?&?&[-> _]). - iDestruct "H" as (qq) "(>#Hq & [HP1 HP1'] & Hl & Hw)". iDestruct "Hq" as %Hq. - destruct (decide (strong = strong')) as [->|?]. - - wp_apply (wp_cas_int_suc with "Hl"). iIntros "Hl". - iMod (own_update with "H●") as "[H● Hown']". - { apply auth_update_alloc, prod_local_update_1, - (op_local_update_discrete _ _ (Some (Cinl ((qq/2)%Qp, 1%positive, None)))) - =>-[/= Hqa ?]. split;[split|]=>//=; last by rewrite left_id. - apply frac_valid. rewrite -Hq comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. } - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iMod ("Hclose1" with "[Hl Hw H● HP1']") as "_". - { iExists _. iFrame. iExists _. rewrite /= [xH ⋅ _]comm_L. iFrame. - rewrite [(qq / 2)%Qp ⋅ _]comm frac_op -[(_ + _)%Qp]assoc Qp_div_2 left_id_L. auto. } - iModIntro. wp_case. iApply "HΦ". iFrame. - - wp_apply (wp_cas_int_fail with "Hl"); [congruence|]. iIntros "Hl". - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iMod ("Hclose1" with "[-HP HΦ]") as "_". - { iExists _. iFrame. iExists qq. iCombine "HP1 HP1'" as "$". auto with iFrame. } - iModIntro. wp_case. iApply ("IH" with "HP HΦ"). - Qed. - - Lemma downgrade_spec (γ : gname) (l : loc) (P : iProp Σ) : - is_arc P1 P2 N γ l -∗ arc_tok_acc γ P (⊤ ∖ ↑N) -∗ - {{{ P }}} downgrade [ #l] {{{ RET #☠; P ∗ weak_tok γ }}}. - Proof. - iIntros "#INV #Hacc !> * HP HΦ". iLöb as "IH". wp_rec. wp_op. wp_bind (!ˢᶜ_)%E. - iInv N as (st) "[>H● H]" "Hclose1". - iApply fupd_wp. iMod ("Hacc" with "HP") as (?) "[Hown Hclose2]". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?&?& wlock & weak &[-> _]). - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iDestruct "H" as (?) "(? & ? & ? & Hw)". destruct wlock. - { iDestruct "Hw" as "(Hw & >%)". subst. wp_read. - iMod ("Hclose1" with "[-HP HΦ]") as "_"; first by iExists _; auto with iFrame. - iModIntro. wp_let. wp_op. wp_if. iApply ("IH" with "HP HΦ"). } - wp_read. iMod ("Hclose1" with "[-HP HΦ]") as "_"; first by iExists _; auto with iFrame. - iModIntro. wp_let. wp_op. wp_if. wp_op. wp_op. wp_bind (CAS _ _ _). - iInv N as (st) "[>H● H]" "Hclose1". - iApply fupd_wp. iMod ("Hacc" with "HP") as (?) "[Hown Hclose2]". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?&?& wlock & weak' &[-> _]). - iMod ("Hclose2" with "Hown") as "HP". iModIntro. - iDestruct "H" as (qq) "(>Heq & HP1 & Hl & Hl1)". iDestruct "Heq" as %Heq. - destruct (decide (weak = weak' ∧ wlock = None)) as [[<- ->]|Hw]. - - wp_apply (wp_cas_int_suc with "Hl1"). iIntros "Hl1". - iMod (own_update with "H●") as "[H● Hown']". - { by apply auth_update_alloc, prod_local_update_2, - (op_local_update_discrete _ _ (1%nat)). } - iMod ("Hclose1" with "[-Hown' HP HΦ]") as "_". - { iExists _. iFrame. iExists _. - rewrite Z.add_comm -(Nat2Z.inj_add 1) /=. auto with iFrame. } - iModIntro. wp_case. iApply "HΦ". iFrame. - - destruct wlock. - + iDestruct "Hl1" as "[Hl1 >%]". subst. - wp_apply (wp_cas_int_fail with "Hl1"); [done..|]. - iIntros "Hl1". iMod ("Hclose1" with "[-HP HΦ]") as "_". - { iExists _. auto with iFrame. } - iModIntro. wp_case. iApply ("IH" with "HP HΦ"). - + wp_apply (wp_cas_int_fail with "Hl1"). - { contradict Hw. split=>//. apply SuccNat2Pos.inj. lia. } - iIntros "Hl1". iMod ("Hclose1" with "[-HP HΦ]") as "_". - { iExists _. auto with iFrame. } - iModIntro. wp_case. iApply ("IH" with "HP HΦ"). - Qed. - - Lemma weak_tok_auth_val γ st : - own γ (● st) -∗ weak_tok γ -∗ ⌜∃ st' weak, st = (st', S weak) ∧ ✓ st'⌝. - Proof. - iIntros "H● Htok". iDestruct (own_valid_2 with "H● Htok") as - %[[Hincl%option_included Hincl'%nat_included]%prod_included [Hval _]] - %auth_both_valid_discrete. - destruct st as [?[]], Hincl as [_|(?&?&[=]&?)]; simpl in *; try lia. eauto. - Qed. - - Lemma clone_weak_spec (γ : gname) (l : loc) (P : iProp Σ) : - is_arc P1 P2 N γ l -∗ weak_tok_acc γ P (⊤ ∖ ↑N) -∗ - {{{ P }}} clone_weak [ #l] {{{ RET #☠; P ∗ weak_tok γ }}}. - Proof. - iIntros "#INV #Hacc !> * HP HΦ". iLöb as "IH". wp_rec. wp_op. wp_bind (!ˢᶜ_)%E. - iAssert (□ (P ={⊤,⊤∖↑N}=∗ ∃ w : Z, (l +ₗ 1) ↦ #w ∗ - ((l +ₗ 1) ↦ #(w + 1) ={⊤∖↑N,⊤}=∗ P ∗ weak_tok γ) ∧ - ((l +ₗ 1) ↦ #w ={⊤∖↑N,⊤}=∗ P)))%I as "#Hproto". - { iIntros "!> HP". iInv N as (st) "[>H● H]" "Hclose1". - iMod ("Hacc" with "HP") as "[Hown Hclose2]". - iDestruct (weak_tok_auth_val with "H● Hown") as %(st' & weak & -> & Hval). - iMod ("Hclose2" with "Hown") as "HP". - destruct st' as [[[[??][?|]]| |]|]; try done; [|iExists _..]. - - by iDestruct "H" as (?) "(_ & _ & _ & _ & >%)". - - iDestruct "H" as (?) "(? & ? & ? & >$)". iIntros "!>"; iSplit; iIntros "?". - + iMod (own_update with "H●") as "[H● $]". - { by apply auth_update_alloc, prod_local_update_2, - (op_local_update_discrete _ _ (1%nat)). } - rewrite Z.add_comm -(Nat2Z.inj_add 1). iFrame. - iApply "Hclose1". iExists _. auto with iFrame. - + iFrame. iApply ("Hclose1" with "[-]"). iExists _. auto with iFrame. - - iDestruct "H" as "[? >$]". iIntros "!>"; iSplit; iIntros "?". - + iMod (own_update with "H●") as "[H● $]". - { by apply auth_update_alloc, prod_local_update_2, - (op_local_update_discrete _ _ (1%nat)). } - rewrite Z.add_comm -(Nat2Z.inj_add 1). iFrame. iApply "Hclose1". - iExists _. auto with iFrame. - + iFrame. iApply ("Hclose1" with "[-]"). iExists _. auto with iFrame. - - iDestruct "H" as "(? & >$ & ?)". iIntros "!>"; iSplit; iIntros "?". - + iMod (own_update with "H●") as "[H● $]". - { by apply auth_update_alloc, prod_local_update_2, - (op_local_update_discrete _ _ (1%nat)). } - rewrite Z.add_comm -(Nat2Z.inj_add 1). iFrame. iApply "Hclose1". - iExists _. auto with iFrame. - + iFrame. iApply ("Hclose1" with "[-]"). iExists _. auto with iFrame. } - iMod ("Hproto" with "HP") as (w) "[Hw [_ Hclose]]". wp_read. - iMod ("Hclose" with "Hw") as "HP". iModIntro. wp_let. wp_op. wp_op. - wp_bind (CAS _ _ _). iMod ("Hproto" with "HP") as (w') "[H↦ Hclose]". - destruct (decide (w = w')) as [<-|]. - - wp_apply (wp_cas_int_suc with "H↦"). iIntros "H↦". - iDestruct "Hclose" as "[Hclose _]". iMod ("Hclose" with "H↦"). iModIntro. - wp_case. by iApply "HΦ". - - wp_apply (wp_cas_int_fail with "H↦"); try done. iIntros "H↦". - iDestruct "Hclose" as "[_ Hclose]". iMod ("Hclose" with "H↦") as "Hown". - iModIntro. wp_case. by iApply ("IH" with "Hown"). - Qed. - - Lemma upgrade_spec (γ : gname) (l : loc) (P : iProp Σ) : - is_arc P1 P2 N γ l -∗ weak_tok_acc γ P (⊤ ∖ ↑N) -∗ - {{{ P }}} upgrade [ #l] - {{{ (b : bool) q, RET #b; P ∗ if b then arc_tok γ q ∗ P1 q else True }}}. - Proof using HP1. - iIntros "#INV #Hacc !> * HP HΦ". iLöb as "IH". wp_rec. wp_bind (!ˢᶜ_)%E. - iAssert (□ (P ={⊤,∅}=∗ ∃ s : Z, l ↦ #s ∗ - (⌜s ≠ 0⌝ -∗ l ↦ #(s + 1) ={∅,⊤}=∗ ∃ q, P ∗ arc_tok γ q ∗ ▷ P1 q) ∧ - (l ↦ #s ={∅,⊤}=∗ P)))%I as "#Hproto". - { iIntros "!> HP". iInv N as (st) "[>H● H]" "Hclose1". - iMod ("Hacc" with "HP") as "[Hown Hclose2]". - iDestruct (weak_tok_auth_val with "H● Hown") as %(st' & weak & -> & Hval). - destruct st' as [[[[q' c]?]| |]|]; try done; iExists _. - - iDestruct "H" as (q) "(>Heq & [HP1 HP1'] & >$ & ?)". iDestruct "Heq" as %Heq. - iIntros "!>"; iSplit; iMod ("Hclose2" with "Hown") as "HP". - + iIntros "_ Hl". iExists (q/2)%Qp. iMod (own_update with "H●") as "[H● $]". - { apply auth_update_alloc. rewrite -[(_,0%nat)]right_id. - apply op_local_update_discrete=>Hv. constructor; last done. - split; last by rewrite /= left_id; apply Hv. split=>[|//]. - apply frac_valid. rewrite /= -Heq comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. } - iFrame. iApply "Hclose1". iExists _. iFrame. iExists _. iFrame. - rewrite /= [xH ⋅ _]comm_L frac_op [(_ + q')%Qp]comm -[(_ + _)%Qp]assoc - Qp_div_2 left_id_L. auto with iFrame. - + iIntros "Hl". iFrame. iApply ("Hclose1" with "[-]"). iExists _. iFrame. - iExists q. iCombine "HP1 HP1'" as "$". auto with iFrame. - - iDestruct "H" as "[>$ ?]". iIntros "!>"; iSplit; first by auto with congruence. - iIntros "Hl". iMod ("Hclose2" with "Hown") as "$". iApply "Hclose1". - iExists _. auto with iFrame. - - iDestruct "H" as "[>$ ?]". iIntros "!>"; iSplit; first by auto with congruence. - iIntros "Hl". iMod ("Hclose2" with "Hown") as "$". iApply "Hclose1". - iExists _. auto with iFrame. } - iMod ("Hproto" with "HP") as (s) "[Hs [_ Hclose]]". wp_read. - iMod ("Hclose" with "Hs") as "HP". iModIntro. wp_let. wp_op; wp_if; case_bool_decide. - - iApply wp_value. iApply ("HΦ" $! _ 1%Qp). auto. - - wp_op. wp_bind (CAS _ _ _). iMod ("Hproto" with "HP") as (s') "[H↦ Hclose]". - destruct (decide (s = s')) as [<-|]. - + wp_apply (wp_cas_int_suc with "H↦"). iIntros "H↦". - iDestruct "Hclose" as "[Hclose _]". - iMod ("Hclose" with "[//] H↦") as (q) "(?&?&?)". iModIntro. - wp_case. iApply "HΦ". iFrame. - + wp_apply (wp_cas_int_fail with "H↦"); try done. iIntros "H↦". - iDestruct "Hclose" as "[_ Hclose]". iMod ("Hclose" with "H↦") as "Hown". - iModIntro. wp_case. by iApply ("IH" with "Hown"). - Qed. - - Lemma drop_weak_spec (γ : gname) (l : loc) : - is_arc P1 P2 N γ l -∗ - {{{ weak_tok γ }}} drop_weak [ #l] - {{{ (b : bool), RET #b ; if b then P2 ∗ l ↦ #0 ∗ (l +ₗ 1) ↦ #0 else True }}}. - Proof. - iIntros "#INV !> * Hown HΦ". iLöb as "IH". wp_rec. wp_op. wp_bind (!ˢᶜ_)%E. - iAssert (□ (weak_tok γ ={⊤,⊤ ∖ ↑N}=∗ ∃ w : Z, (l +ₗ 1) ↦ #w ∗ - ((l +ₗ 1) ↦ #(w - 1) ={⊤ ∖ ↑N,⊤}=∗ ⌜w ≠ 1⌝ ∨ - ▷ P2 ∗ l ↦ #0 ∗ (l +ₗ 1) ↦ #0) ∧ - ((l +ₗ 1) ↦ #w ={⊤ ∖ ↑N,⊤}=∗ weak_tok γ)))%I as "#Hproto". - { iIntros "!> Hown". iInv N as (st) "[>H● H]" "Hclose1". - iDestruct (weak_tok_auth_val with "H● Hown") as %(st' & weak & -> & Hval). - destruct st' as [[[[??][?|]]| |]|]; try done; [|iExists _..]. - - by iDestruct "H" as (?) "(_ & _ & _ & _ & >%)". - - iDestruct "H" as (q) "(>Heq & HP1 & ? & >$)". iIntros "!>"; iSplit; iIntros "Hl1". - + iMod ("Hclose1" with "[>-]") as "_"; last iLeft; auto with lia. - iExists _. iMod (own_update_2 with "H● Hown") as "$". - { apply auth_update_dealloc, prod_local_update_2, - (cancel_local_update_unit 1%nat), _. } - iExists _. iFrame. by replace (S (S weak) - 1) with (S weak:Z) by lia. - + iFrame. iApply "Hclose1". iExists _. auto with iFrame. - - iDestruct "H" as "[? >$]". iIntros "!>"; iSplit; iIntros "Hl1". - + iMod ("Hclose1" with "[>-]") as "_"; last iLeft; auto with lia. - iExists _. iMod (own_update_2 with "H● Hown") as "$". - { apply auth_update_dealloc, prod_local_update_2, - (cancel_local_update_unit 1%nat), _. } - iFrame. by replace (S (S weak) - 1) with (S weak:Z) by lia. - + iFrame. iApply "Hclose1". iExists _. auto with iFrame. - - iDestruct "H" as "(>? & >$ & HP2)". iIntros "!>"; iSplit; iIntros "Hl1". - + iMod (own_update_2 with "H● Hown") as "H●". - { apply auth_update_dealloc, prod_local_update_2, - (cancel_local_update_unit 1%nat), _. } - destruct weak as [|weak]. - * iMod ("Hclose1" with "[H●]") as "_"; last by auto with iFrame. - iExists _. iFrame. - * iMod ("Hclose1" with "[>-]") as "_"; last iLeft; auto with lia. - iExists _. iFrame. by replace (S (S weak) - 1) with (S weak:Z) by lia. - + iFrame. iApply "Hclose1". iExists _. auto with iFrame. } - iMod ("Hproto" with "Hown") as (w) "[Hw [_ Hclose]]". wp_read. - iMod ("Hclose" with "Hw") as "Hown". iModIntro. wp_let. wp_op. wp_op. - wp_bind (CAS _ _ _). - iMod ("Hproto" with "Hown") as (w') "[Hw Hclose]". destruct (decide (w = w')) as [<-|?]. - - wp_apply (wp_cas_int_suc with "Hw"). iIntros "Hw". - iDestruct "Hclose" as "[Hclose _]". iMod ("Hclose" with "Hw") as "HP2". iModIntro. - wp_case. wp_op; case_bool_decide; subst; iApply "HΦ"=>//. - by iDestruct "HP2" as "[%|$]". - - wp_apply (wp_cas_int_fail with "Hw"); try done. iIntros "Hw". - iDestruct "Hclose" as "[_ Hclose]". iMod ("Hclose" with "Hw") as "Hown". - iModIntro. wp_case. by iApply ("IH" with "Hown"). - Qed. - - Lemma close_last_strong γ l : - is_arc P1 P2 N γ l -∗ own γ (◯ (Some (Cinr (Excl ())), 0%nat)) -∗ P2 - ={⊤}=∗ weak_tok γ. - Proof. - iIntros "INV H◯ HP2". iInv N as ([st ?]) "[>H● H]" "Hclose". - iDestruct (own_valid_2 with "H● H◯") - as %[[[[=]|Hincl]%option_included _]%prod_included [? _]]%auth_both_valid_discrete. - simpl in Hincl. destruct Hincl as (?&?&[=<-]&->&[?|[]%exclusive_included]); - try done; try apply _. setoid_subst. - iMod (own_update_2 with "H● H◯") as "[H● $]". - { apply auth_update, prod_local_update', (op_local_update _ _ 1%nat)=>//. - apply delete_option_local_update, _. } - iApply "Hclose". iExists _. by iFrame. - Qed. - - Lemma drop_arc_spec (γ : gname) (q: Qp) (l : loc) : - is_arc P1 P2 N γ l -∗ - {{{ arc_tok γ q ∗ P1 q }}} drop_arc [ #l] - {{{ (b : bool), RET #b ; if b then P1 1 ∗ (P2 ={⊤}=∗ weak_tok γ) else True }}}. - Proof using HP1. - iIntros "#INV !> * [Hown HP1] HΦ". iLöb as "IH". - wp_rec. wp_bind (!ˢᶜ_)%E. iInv N as (st) "[>H● H]" "Hclose". - iDestruct (arc_tok_auth_val with "H● Hown") as %(?& s &?&?&[-> _]). - iDestruct "H" as (x') "(? & ? & ? & ?)". wp_read. - iMod ("Hclose" with "[-Hown HP1 HΦ]") as "_". { iExists _. auto with iFrame. } - iModIntro. wp_let. wp_op. wp_bind (CAS _ _ _). iInv N as (st) "[>H● H]" "Hclose". - iDestruct (arc_tok_auth_val with "H● Hown") as %(q' & s' & wl & w &[-> Hqq']). - iDestruct "H" as (q'') "(>Hq'' & HP1' & Hs & Hw)". iDestruct "Hq''" as %Hq''. - destruct (decide (s = s')) as [<-|?]. - - wp_apply (wp_cas_int_suc with "Hs"). iIntros "Hs". - destruct decide as [->|?]. - + destruct Hqq' as [<- ->]. - iMod (own_update_2 with "H● Hown") as "[H● H◯]". - { apply auth_update, prod_local_update_1. rewrite -[x in (x, _)]right_id. - etrans; first apply: cancel_local_update_unit. - by apply (op_local_update _ _ (Some (Cinr (Excl ())))). } - iMod ("Hclose" with "[H● Hs Hw]") as "_"; first by iExists _; do 2 iFrame. - iModIntro. wp_case. iApply wp_fupd. wp_op. - iApply ("HΦ"). rewrite -{2}Hq''. iCombine "HP1 HP1'" as "$". - by iApply close_last_strong. - + destruct Hqq' as [? ->]. - rewrite -[in (_, _)](Pos.succ_pred s) // -[wl in Cinl (_, wl)]left_id - -Pos.add_1_l 2!pair_op Cinl_op Some_op. - iMod (own_update_2 _ _ _ _ with "H● Hown") as "H●". - { apply auth_update_dealloc, prod_local_update_1, @cancel_local_update_unit, _. } - iMod ("Hclose" with "[- HΦ]") as "_". - { iExists _. iFrame. iExists (q + q'')%Qp. iCombine "HP1 HP1'" as "$". - iSplit; last by destruct s. - iIntros "!> !%". rewrite assoc -Hq''. f_equal. apply comm, _. } - iModIntro. wp_case. wp_op; case_bool_decide; simplify_eq. by iApply "HΦ". - - wp_apply (wp_cas_int_fail with "Hs"); [congruence|]. iIntros "Hs". - iSpecialize ("IH" with "Hown HP1 HΦ"). - iMod ("Hclose" with "[- IH]") as "_"; first by iExists _; auto with iFrame. - iModIntro. by wp_case. - Qed. - - Lemma try_unwrap_spec (γ : gname) (q: Qp) (l : loc) : - is_arc P1 P2 N γ l -∗ - {{{ arc_tok γ q ∗ P1 q }}} try_unwrap [ #l] - {{{ (b : bool), RET #b ; - if b then P1 1 ∗ (P2 ={⊤}=∗ weak_tok γ) else arc_tok γ q ∗ P1 q }}}. - Proof using HP1. - iIntros "#INV !> * [Hown HP1] HΦ". wp_rec. iInv N as (st) "[>H● H]" "Hclose". - iDestruct (arc_tok_auth_val with "H● Hown") as %(q' & s & wl & w &[-> Hqq']). - iDestruct "H" as (q'') "(>Hq'' & HP1' & Hs & Hw)". iDestruct "Hq''" as %Hq''. - destruct (decide (s = xH)) as [->|?]. - - wp_apply (wp_cas_int_suc with "Hs"). iIntros "Hs". - destruct Hqq' as [<- ->]. iMod (own_update_2 with "H● Hown") as "[H● H◯]". - { apply auth_update, prod_local_update_1. rewrite -[x in (x, _)]right_id. - etrans; first apply: cancel_local_update_unit. - by apply (op_local_update _ _ (Some (Cinr (Excl ())))). } - iMod ("Hclose" with "[H● Hs Hw]") as "_"; first by iExists _; do 2 iFrame. - iApply ("HΦ" $! true). rewrite -{1}Hq''. iCombine "HP1 HP1'" as "$". - by iApply close_last_strong. - - wp_apply (wp_cas_int_fail with "Hs"); [congruence|]. iIntros "Hs". - iMod ("Hclose" with "[-Hown HP1 HΦ]") as "_"; first by iExists _; auto with iFrame. - iApply ("HΦ" $! false). by iFrame. - Qed. - - Lemma is_unique_spec (γ : gname) (q: Qp) (l : loc) : - is_arc P1 P2 N γ l -∗ - {{{ arc_tok γ q ∗ P1 q }}} is_unique [ #l] - {{{ (b : bool), RET #b ; - if b then l ↦ #1 ∗ (l +ₗ 1) ↦ #1 ∗ P1 1 - else arc_tok γ q ∗ P1 q }}}. - Proof using HP1. - iIntros "#INV !> * [Hown HP1] HΦ". wp_rec. wp_bind (CAS _ _ _). wp_op. - iInv N as (st) "[>H● H]" "Hclose". - iDestruct (arc_tok_auth_val with "H● Hown") as %(? & ? & wl & w &[-> _]). - iDestruct "H" as (?) "(>Hq'' & HP1' & >Hs & >Hw)". - destruct wl; last (destruct w; last first). - - iDestruct "Hw" as "[Hw %]". subst. - iApply (wp_cas_int_fail with "Hw")=>//. iNext. iIntros "Hw". - iMod ("Hclose" with "[-HΦ Hown HP1]") as "_"; first by iExists _; eauto with iFrame. - iModIntro. wp_case. iApply "HΦ". iFrame. - - iApply (wp_cas_int_fail with "Hw"); [lia|]. iNext. iIntros "Hw". - iMod ("Hclose" with "[-HΦ Hown HP1]") as "_"; first by iExists _; eauto with iFrame. - iModIntro. wp_case. iApply "HΦ". iFrame. - - iApply (wp_cas_int_suc with "Hw")=>//. iNext. iIntros "Hw". - iMod (own_update_2 with "H● Hown") as "[H● H◯]". - { by apply auth_update, prod_local_update_1, option_local_update, - csum_local_update_l, prod_local_update_2, - (alloc_option_local_update (Excl ())). } - iMod ("Hclose" with "[-HΦ HP1 H◯]") as "_"; first by iExists _; eauto with iFrame. - iModIntro. wp_case. wp_bind (!ˢᶜ_)%E. iInv N as ([st ?]) "[>H● H]" "Hclose". - iDestruct (own_valid_2 with "H● H◯") as - %[[[[=]|Hincl]%option_included _]%prod_included [? _]]%auth_both_valid_discrete. - simpl in Hincl. destruct Hincl as (? & ? & [=<-] & -> & [|Hincl]); last first. - + apply csum_included in Hincl. destruct Hincl as [->|[Hincl|(?&?&[=]&?)]]=>//. - destruct Hincl as (?&[[??]?]&[=<-]&->&[[_ Hlt]%prod_included _]%prod_included). - simpl in Hlt. apply pos_included in Hlt. - iDestruct "H" as (?) "(? & ? & ? & ?)". wp_read. - iMod ("Hclose" with "[-HΦ H◯ HP1]") as "_"; first by iExists _; auto with iFrame. - iModIntro. wp_let. wp_op; case_bool_decide; [lia|]. wp_let. wp_op. wp_bind (_ <-ˢᶜ _)%E. - iInv N as ([st w]) "[>H● H]" "Hclose". - iDestruct (own_valid_2 with "H● H◯") as - %[[[[=]|Hincl]%option_included _]%prod_included [Hval _]]%auth_both_valid_discrete. - simpl in Hincl. destruct Hincl as (x1 & x2 & [=<-] & -> & Hincl); last first. - assert (∃ q p, x2 = Cinl (q, p, Excl' ())) as (? & ? & ->). - { destruct Hincl as [|Hincl]; first by setoid_subst; eauto. - apply csum_included in Hincl. destruct Hincl as [->|[Hincl|(?&?&[=]&?)]]=>//. - destruct Hincl as (?&[[??]?]&[=<-]&->&[_ Hincl%option_included]%prod_included). - simpl in Hincl. destruct Hincl as [[=]|(?&?&[=<-]&->&[?|[]%exclusive_included])]; - [|by apply _|by apply Hval]. setoid_subst. eauto. } - iDestruct "H" as (?) "(? & ? & ? & >? & >%)". subst. wp_write. - iMod (own_update_2 with "H● H◯") as "[H● H◯]". - { apply auth_update, prod_local_update_1, option_local_update, - csum_local_update_l, prod_local_update_2, delete_option_local_update, _. } - iMod ("Hclose" with "[-HΦ H◯ HP1]") as "_"; first by iExists _; auto with iFrame. - iModIntro. wp_seq. iApply "HΦ". iFrame. - + setoid_subst. iDestruct "H" as (?) "(Hq & HP1' & ? & >? & >%)". subst. wp_read. - iMod (own_update_2 with "H● H◯") as "H●". - { apply auth_update_dealloc. rewrite -{1}[(_, 0%nat)]right_id. - apply cancel_local_update_unit, _. } - iMod ("Hclose" with "[H●]") as "_"; first by iExists _; iFrame. - iModIntro. wp_seq. wp_op. wp_let. wp_op. wp_write. iApply "HΦ". - iDestruct "Hq" as %<-. iCombine "HP1 HP1'" as "$". iFrame. - Qed. - - Lemma try_unwrap_full_spec (γ : gname) (q: Qp) (l : loc) : - is_arc P1 P2 N γ l -∗ - {{{ arc_tok γ q ∗ P1 q }}} try_unwrap_full [ #l] - {{{ (x : fin 3), RET #x ; - match x : nat with - (* No other reference : we get everything. *) - | 0%nat => l ↦ #1 ∗ (l +ₗ 1) ↦ #1 ∗ P1 1 - (* No other strong, but there are weaks : we get P1, - plus the option to get a weak if we pay P2. *) - | 1%nat => P1 1 ∗ (P2 ={⊤}=∗ weak_tok γ) - (* There are other strong : we get back what we gave at the first place. *) - | _ (* 2 *) => arc_tok γ q ∗ P1 q - end }}}. - Proof using HP1. - iIntros "#INV !> * [Hown HP1] HΦ". wp_rec. wp_bind (CAS _ _ _). - iInv N as (st) "[>H● H]" "Hclose". - iDestruct (arc_tok_auth_val with "H● Hown") as %(q' & s & wl & w &[-> Hqq']). - iDestruct "H" as (q'') "(>Hq'' & HP1' & Hs & Hw)". iDestruct "Hq''" as %Hq''. - destruct (decide (s = xH)) as [->|?]. - - wp_apply (wp_cas_int_suc with "Hs")=>//. iIntros "Hs". - destruct Hqq' as [<- ->]. iMod (own_update_2 with "H● Hown") as "[H● H◯]". - { apply auth_update, prod_local_update_1. rewrite -[x in (x, _)]right_id. - etrans; first apply: cancel_local_update_unit. - by apply (op_local_update _ _ (Some (Cinr (Excl ())))). } - iCombine "HP1" "HP1'" as "HP1". rewrite Hq''. clear Hq'' q''. - iMod ("Hclose" with "[H● Hs Hw]") as "_"; first by iExists _; do 2 iFrame. - clear w. iModIntro. wp_case. wp_op. wp_bind (!ˢᶜ_)%E. - iInv N as ([st w']) "[>H● H]" "Hclose". - iDestruct (own_valid_2 with "H● H◯") - as %[[[[=]|Hincl]%option_included _]%prod_included [? _]]%auth_both_valid_discrete. - simpl in Hincl. destruct Hincl as (?&?&[=<-]&->&[?|[]%exclusive_included]); - [|by apply _|done]. setoid_subst. iDestruct "H" as "[Hl Hl1]". - wp_read. destruct w'. - + iMod (own_update_2 with "H● H◯") as "H●". - { apply auth_update_dealloc, prod_local_update_1, - delete_option_local_update, _. } - iMod ("Hclose" with "[H●]") as "_"; first by iExists _; iFrame. iModIntro. - wp_op. wp_if. wp_write. iApply ("HΦ" $! 0%fin). iFrame. - + iMod ("Hclose" with "[H● Hl Hl1]") as "_"; first by iExists _; do 2 iFrame. - iModIntro. wp_op; case_bool_decide; [lia|]. wp_if. iApply ("HΦ" $! 1%fin). iFrame. - by iApply close_last_strong. - - wp_apply (wp_cas_int_fail with "Hs"); [congruence|]. iIntros "Hs". - iMod ("Hclose" with "[H● Hs Hw HP1']") as "_"; first by iExists _; auto with iFrame. - iModIntro. wp_case. iApply ("HΦ" $! 2%fin). iFrame. - Qed. -End arc. - -Global Typeclasses Opaque is_arc arc_tok weak_tok. diff --git a/theories/lang/lib/spawn.v b/theories/lang/lib/spawn.v index 3b21e980..b1cd6aee 100644 --- a/theories/lang/lib/spawn.v +++ b/theories/lang/lib/spawn.v @@ -28,7 +28,7 @@ Definition join : val := "join" ["c"]. (** The CMRA & functor we need. *) -Class spawnG Σ := SpawnG { spawn_tokG :> inG Σ (exclR unitO) }. +Class spawnG Σ := SpawnG { spawn_tokG :: inG Σ (exclR unitO) }. Definition spawnΣ : gFunctors := #[GFunctor (exclR unitO)]. Global Instance subG_spawnΣ {Σ} : subG spawnΣ Σ → spawnG Σ. @@ -90,7 +90,7 @@ Proof. auto. } iDestruct "Hinv" as (b) "[>Hc _]". wp_write. iMod ("Hclose" with "[-HΦ]"). - - iNext. iRight. iExists true. iFrame. iExists _. iFrame. + - iNext. iRight. iExists true. iFrame. - iModIntro. wp_seq. by iApply "HΦ". Qed. diff --git a/theories/lang/lifting.v b/theories/lang/lifting.v index 4c263e3d..6172956a 100644 --- a/theories/lang/lifting.v +++ b/theories/lang/lifting.v @@ -8,8 +8,8 @@ Import uPred. Class lrustGS Σ := LRustGS { lrustGS_invGS : invGS Σ; - lrustGS_gen_heapGS :> heapGS Σ; - lrustGS_gen_timeGS :> timeGS Σ + lrustGS_gen_heapGS :: heapGS Σ; + lrustGS_gen_timeGS :: timeGS Σ }. Global Program Instance lrustGS_irisGS `{!lrustGS Σ} : irisGS lrust_lang Σ := { @@ -36,9 +36,9 @@ Ltac inv_bin_op_eval := end. Local Hint Extern 0 (atomic _) => solve_atomic : core. -Local Hint Extern 0 (head_reducible _ _) => eexists _, _, _, _; simpl : core. +Local Hint Extern 0 (base_reducible _ _) => eexists _, _, _, _; simpl : core. -Local Hint Constructors head_step bin_op_eval lit_neq lit_eq : core. +Local Hint Constructors base_step bin_op_eval lit_neq lit_eq : core. Local Hint Resolve alloc_fresh : core. Local Hint Resolve to_of_val : core. @@ -109,7 +109,7 @@ Proof. iIntros (?????) "[Hσ Ht]". iMod (step_cumulative_time_receipt with "TIME Ht") as "[Ht Hclose]"=>//. iMod ("Hwp" $! _ _ _ [] 0%nat with "[$]") as "[$ Hwp]". - iIntros "!>" (e2 σ2 efs stp). iMod ("Hwp" $! e2 σ2 efs stp) as "Hwp". + iIntros "!>" (e2 σ2 efs stp) "H£". iMod ("Hwp" $! e2 σ2 efs stp with "H£") as "Hwp". iIntros "!> !>". iMod "Hwp". iModIntro. iApply (step_fupdN_wand with "Hwp"). iIntros ">([$ Ht] & Hwp & $)". iMod ("Hclose" with "Ht") as "[$ ?]". @@ -132,9 +132,9 @@ Qed. Lemma wp_fork E e : {{{ ▷ WP e {{ _, True }} }}} Fork e @ E {{{ RET LitV LitPoison; True }}}. Proof. - iIntros (?) "?HΦ". iApply wp_lift_atomic_head_step; [done|]. + iIntros (?) "?HΦ". iApply wp_lift_atomic_base_step; [done|]. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht] !>"; iSplit; first by eauto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. iFrame. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step. iFrame. iMod (time_interp_step with "Ht") as "$". by iApply "HΦ". Qed. @@ -142,9 +142,9 @@ Qed. Local Ltac solve_exec_safe := intros; destruct_and?; subst; do 3 eexists; econstructor; simpl; eauto with lia. Local Ltac solve_exec_puredet := - simpl; intros; destruct_and?; inv_head_step; inv_bin_op_eval; inv_lit; done. + simpl; intros; destruct_and?; inv_base_step; inv_bin_op_eval; inv_lit; done. Local Ltac solve_pure_exec := - intros ?; apply nsteps_once, pure_head_step_pure_step; + intros ?; apply nsteps_once, pure_base_step_pure_step; constructor; [solve_exec_safe | solve_exec_puredet]. Global Instance pure_rec e f xl erec erec' el : @@ -203,9 +203,9 @@ Lemma wp_nd_int E : {{{ True }}} NdInt @ E {{{ z, RET LitV $ LitInt z; True }}}. Proof. - iIntros (? _) "Φ". iApply wp_lift_atomic_head_step_no_fork; auto. + iIntros (? _) "Φ". iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n') "[σ t] !>"; iSplit. { unshelve auto. apply 0. } - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step. iMod (time_interp_step with "t") as "$". iFrame "σ". by iDestruct ("Φ" with "[//]") as "$". Qed. @@ -216,9 +216,9 @@ Lemma wp_alloc E (n : Z) : {{{ True }}} Alloc (Lit $ LitInt n) @ E {{{ l (sz: nat), RET LitV $ LitLoc l; ⌜n = sz⌝ ∗ †l…sz ∗ l ↦∗ repeat (LitV LitPoison) sz }}}. Proof. - iIntros (? Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. + iIntros (? Φ) "_ HΦ". iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n') "[Hσ Ht] !>"; iSplit; first by auto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step. iMod (heap_alloc with "Hσ") as "[$ Hl]"; [done..|]. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit=> //. iApply ("HΦ" $! _ (Z.to_nat n)). auto with lia iFrame. @@ -230,12 +230,12 @@ Lemma wp_free E (n:Z) l vl : Free (Lit $ LitInt n) (Lit $ LitLoc l) @ E {{{ RET LitV LitPoison; True }}}. Proof. - iIntros (? Φ) "[>Hmt >Hf] HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. + iIntros (? Φ) "[>Hmt >Hf] HΦ". iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n') "[Hσ Ht]". iMod (heap_free _ _ _ n with "Hσ Hmt Hf") as "(% & % & Hσ)"=>//. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first by auto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step. iModIntro; iSplit=> //. iFrame. iApply "HΦ"; auto. Qed. @@ -243,12 +243,12 @@ Lemma wp_read_sc E l q v : {{{ ▷ l ↦{q} v }}} Read ScOrd (Lit $ LitLoc l) @ E {{{ RET v; l ↦{q} v }}}. Proof. - iIntros (?) ">Hv HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. + iIntros (?) ">Hv HΦ". iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iDestruct (heap_read with "Hσ Hv") as %[m ?]. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first by eauto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step. iModIntro; iSplit=> //. iFrame. by iApply "HΦ". Qed. @@ -256,18 +256,18 @@ Lemma wp_read_na E l q v : {{{ ▷ l ↦{q} v }}} Read Na1Ord (Lit $ LitLoc l) @ E {{{ RET v; l ↦{q} v }}}. Proof. - iIntros (Φ) ">Hv HΦ". iApply wp_lift_head_step; auto. + iIntros (Φ) ">Hv HΦ". iApply wp_lift_base_step; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iMod (heap_read_na with "Hσ Hv") as (m) "(% & Hσ & Hσclose)". iApply fupd_mask_intro; first set_solver. iIntros "Hclose". iSplit; first by eauto. - iNext; iIntros (e2 σ2 efs Hstep); inv_head_step. iMod "Hclose" as "_". + iNext; iIntros (e2 σ2 efs Hstep) "_"; inv_base_step. iMod "Hclose" as "_". iMod (time_interp_step with "Ht") as "$". iIntros "!> {$Hσ}". iSplit; last done. - clear dependent σ1 n. iApply wp_lift_atomic_head_step_no_fork; auto. + clear dependent σ1 n. iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt' κ' κs' n') "[Hσ Ht]". iMod (time_interp_step with "Ht") as "$". iMod ("Hσclose" with "Hσ") as (n) "(% & Hσ & Hv)". iModIntro; iSplit; first by eauto. - iNext; iIntros (e2 σ2 efs Hstep) "!>"; inv_head_step. + iNext; iIntros (e2 σ2 efs Hstep) "_ !>"; inv_base_step. iFrame "Hσ". iSplit; [done|]. by iApply "HΦ". Qed. @@ -276,11 +276,11 @@ Lemma wp_write_sc E l e v v' : {{{ ▷ l ↦ v' }}} Write ScOrd (Lit $ LitLoc l) e @ E {{{ RET LitV LitPoison; l ↦ v }}}. Proof. - iIntros (<- Φ) ">Hv HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. + iIntros (<- Φ) ">Hv HΦ". iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iDestruct (heap_read_1 with "Hσ Hv") as %?. iMod (heap_write _ _ _ v with "Hσ Hv") as "[Hσ Hv]". iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first by eauto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step. iModIntro; iSplit=> //. iFrame. by iApply "HΦ". Qed. @@ -290,16 +290,16 @@ Lemma wp_write_na E l e v v' : {{{ RET LitV LitPoison; l ↦ v }}}. Proof. iIntros (<- Φ) ">Hv HΦ". - iApply wp_lift_head_step; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". + iApply wp_lift_base_step; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iMod (heap_write_na with "Hσ Hv") as "(% & Hσ & Hσclose)". iApply (fupd_mask_intro _ ∅); first set_solver. iIntros "Hclose". iSplit; first by eauto. - iNext; iIntros (e2 σ2 efs Hstep); inv_head_step. iMod "Hclose" as "_". + iNext; iIntros (e2 σ2 efs Hstep) "_"; inv_base_step. iMod "Hclose" as "_". iMod (time_interp_step with "Ht") as "$". iModIntro. iFrame "Hσ". iSplit; last done. - clear dependent σ1. iApply wp_lift_atomic_head_step_no_fork; auto. + clear dependent σ1. iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt' κ' κs' n') "[Hσ Ht]". iMod ("Hσclose" with "Hσ") as "(% & Hσ & Hv)". iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first by eauto. - iNext; iIntros (e2 σ2 efs Hstep) "!>"; inv_head_step. + iNext; iIntros (e2 σ2 efs Hstep) "_ !>"; inv_base_step. iFrame "Hσ". iSplit; [done|]. by iApply "HΦ". Qed. @@ -310,12 +310,12 @@ Lemma wp_cas_int_fail E l q z1 e2 lit2 zl : {{{ RET LitV $ LitInt 0; l ↦{q} LitV (LitInt zl) }}}. Proof. iIntros (<- ? Φ) ">Hv HΦ". - iApply wp_lift_atomic_head_step_no_fork; auto. + iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iDestruct (heap_read with "Hσ Hv") as %[m ?]. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first by eauto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step; inv_lit. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step; inv_lit. iModIntro; iSplit=> //. iFrame. by iApply "HΦ". Qed. @@ -326,12 +326,12 @@ Lemma wp_cas_suc E l lit1 e2 lit2 : {{{ RET LitV (LitInt 1); l ↦ LitV lit2 }}}. Proof. iIntros (<- ? Φ) ">Hv HΦ". - iApply wp_lift_atomic_head_step_no_fork; auto. + iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iDestruct (heap_read_1 with "Hσ Hv") as %?. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first (destruct lit1; by eauto). - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step; [inv_lit|]. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step; [inv_lit|]. iMod (heap_write with "Hσ Hv") as "[$ Hv]". iModIntro; iSplit=> //. iFrame. by iApply "HΦ". Qed. @@ -358,14 +358,14 @@ Lemma wp_cas_loc_fail E l q q' q1 l1 v1' e2 lit2 l' vl' : l ↦{q} LitV (LitLoc l') ∗ l' ↦{q'} vl' ∗ l1 ↦{q1} v1' }}}. Proof. iIntros (<- ? Φ) "(>Hl & >Hl' & >Hl1) HΦ". - iApply wp_lift_atomic_head_step_no_fork; auto. + iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iDestruct (heap_read with "Hσ Hl") as %[nl ?]. iDestruct (heap_read with "Hσ Hl'") as %[nl' ?]. iDestruct (heap_read with "Hσ Hl1") as %[nl1 ?]. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first by eauto. - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step; inv_lit. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step; inv_lit. iModIntro; iSplit=> //. iFrame. iApply "HΦ"; iFrame. Qed. @@ -378,12 +378,12 @@ Lemma wp_cas_loc_nondet E l l1 e2 l2 ll : else ⌜l1 ≠ ll⌝ ∗ l ↦ LitV (LitLoc ll) }}}. Proof. iIntros (<- Φ) ">Hv HΦ". - iApply wp_lift_atomic_head_step_no_fork; auto. + iApply wp_lift_atomic_base_step_no_fork; auto. iIntros (σ1 stepcnt κ κs n) "[Hσ Ht]". iDestruct (heap_read_1 with "Hσ Hv") as %?. iMod (time_interp_step with "Ht") as "$". iModIntro; iSplit; first (destruct (decide (ll = l1)) as [->|]; by eauto). - iNext; iIntros (v2 σ2 efs Hstep); inv_head_step; last lia. + iNext; iIntros (v2 σ2 efs Hstep) "_"; inv_base_step; last lia. - inv_lit. iModIntro; iSplit; [done|]; iFrame "Hσ". iApply "HΦ"; simpl; auto. - iMod (heap_write with "Hσ Hv") as "[$ Hv]". @@ -391,20 +391,21 @@ Proof. Qed. Lemma wp_eq_loc E (l1 : loc) (l2: loc) q1 q2 v1 v2 P Φ : - (P -∗ ▷ l1 ↦{q1} v1) → - (P -∗ ▷ l2 ↦{q2} v2) → - (P -∗ ▷ Φ (LitV (bool_decide (l1 = l2)))) → - P -∗ WP BinOp EqOp (Lit (LitLoc l1)) (Lit (LitLoc l2)) @ E {{ Φ }}. + (P ⊢ ▷ l1 ↦{q1} v1) → + (P ⊢ ▷ l2 ↦{q2} v2) → + (P ⊢ ▷ Φ (LitV (bool_decide (l1 = l2)))) → + P ⊢ WP BinOp EqOp (Lit (LitLoc l1)) (Lit (LitLoc l2)) @ E {{ Φ }}. Proof. iIntros (Hl1 Hl2 Hpost) "HP". destruct (bool_decide_reflect (l1 = l2)) as [->|]. - - iApply wp_lift_pure_det_head_step_no_fork'; + - iApply wp_lift_pure_det_base_step_no_fork'; [done|solve_exec_safe|solve_exec_puredet|]. - iApply wp_value. by iApply Hpost. - - iApply wp_lift_atomic_head_step_no_fork; subst=>//. + iPoseProof (Hpost with "HP") as "?". + iIntros "!> _". by iApply wp_value. + - iApply wp_lift_atomic_base_step_no_fork; subst=>//. iIntros (σ1 stepcnt κ κs n') "[Hσ1 Ht]". iMod (time_interp_step with "Ht") as "$". - iModIntro. inv_head_step. iSplitR. + iModIntro. inv_base_step. iSplitR. { iPureIntro. repeat eexists. econstructor. eapply BinOpEqFalse. by auto. } (* We need to do a little gymnastics here to apply Hne now and strip away a ▷ but also have the ↦s. *) @@ -414,8 +415,8 @@ Proof. + iExists _, _. by iApply Hl1. + iExists _, _. by iApply Hl2. + by iApply Hpost. } - clear Hl1 Hl2. iNext. iIntros (e2 σ2 efs Hs) "!>". - inv_head_step. iSplitR=>//. inv_bin_op_eval; inv_lit. + clear Hl1 Hl2. iNext. iIntros (e2 σ2 efs Hs) "_ !>". + inv_base_step. iSplitR=>//. inv_bin_op_eval; inv_lit. + iExFalso. iDestruct "HP" as "[Hl1 _]". iDestruct "Hl1" as (??) "Hl1". iDestruct (heap_read σ2 with "Hσ1 Hl1") as %[??]; simplify_eq. @@ -464,6 +465,6 @@ Proof. iIntros (Hlen Hf) "Hel HΦ". rewrite -(vec_to_list_to_vec Ql). generalize (list_to_vec Ql). rewrite Hlen. clear Hlen Ql=>Ql. iApply (wp_app_vec with "Hel"). iIntros (vl) "Hvl". - iApply ("HΦ" with "[%] Hvl"). by rewrite vec_to_list_length. + iApply ("HΦ" with "[%] Hvl"). by rewrite length_vec_to_list. Qed. End lifting. diff --git a/theories/lang/proofmode.v b/theories/lang/proofmode.v index 14b227b6..b56256e5 100644 --- a/theories/lang/proofmode.v +++ b/theories/lang/proofmode.v @@ -9,9 +9,9 @@ Import uPred. Lemma tac_wp_value `{!lrustGS Σ} Δ E Φ e v : IntoVal e v → envs_entails Δ (Φ v) → envs_entails Δ (WP e @ E {{ Φ }}). -Proof. rewrite envs_entails_eq=> ? ->. by apply wp_value. Qed. +Proof. rewrite envs_entails_unseal=> ? ->. by apply wp_value. Qed. -Ltac wp_value_head := eapply tac_wp_value; [iSolveTC|reduction.pm_prettify]. +Ltac wp_value_head := eapply tac_wp_value; [tc_solve|reduction.pm_prettify]. Lemma tac_wp_pure `{!lrustGS Σ} K Δ Δ' E e1 e2 φ n Φ : PureExec φ n e1 e2 → @@ -20,8 +20,9 @@ Lemma tac_wp_pure `{!lrustGS Σ} K Δ Δ' E e1 e2 φ n Φ : envs_entails Δ' (WP fill K e2 @ E {{ Φ }}) → envs_entails Δ (WP fill K e1 @ E {{ Φ }}). Proof. - rewrite envs_entails_eq=> ??? HΔ'. rewrite into_laterN_env_sound /=. - rewrite -wp_bind HΔ' -wp_pure_step_later //. by rewrite -wp_bind_inv. + rewrite envs_entails_unseal=> ??? HΔ'. rewrite into_laterN_env_sound /=. + rewrite -wp_bind HΔ' -wp_pure_step_later //. rewrite -wp_bind_inv. + f_equiv. apply wand_intro_l. by rewrite sep_elim_r. Qed. Tactic Notation "wp_pure" open_constr(efoc) := @@ -30,9 +31,9 @@ Tactic Notation "wp_pure" open_constr(efoc) := | |- envs_entails _ (wp ?s ?E ?e ?Q) => reshape_expr e ltac:(fun K e' => unify e' efoc; eapply (tac_wp_pure K); - [simpl; iSolveTC (* PureExec *) + [simpl; tc_solve (* PureExec *) |try done (* The pure condition for PureExec *) - |iSolveTC (* IntoLaters *) + |tc_solve (* IntoLaters *) |simpl_subst; try wp_value_head (* new goal *)]) || fail "wp_pure: cannot find" efoc "in" e "or" efoc "is not a reduct" | _ => fail "wp_pure: not a 'wp'" @@ -45,7 +46,7 @@ Lemma tac_wp_eq_loc `{!lrustGS Σ} K Δ Δ' E i1 i2 l1 l2 q1 q2 v1 v2 Φ : envs_entails Δ' (WP fill K (Lit (bool_decide (l1 = l2))) @ E {{ Φ }}) → envs_entails Δ (WP fill K (BinOp EqOp (Lit (LitLoc l1)) (Lit (LitLoc l2))) @ E {{ Φ }}). Proof. - rewrite envs_entails_eq=> ? /envs_lookup_sound /=. rewrite sep_elim_l=> ?. + rewrite envs_entails_unseal=> ? /envs_lookup_sound /=. rewrite sep_elim_l=> ?. move /envs_lookup_sound; rewrite sep_elim_l=> ? HΔ. rewrite -wp_bind. rewrite into_laterN_env_sound /=. eapply wp_eq_loc; eauto using later_mono. Qed. @@ -55,7 +56,7 @@ Tactic Notation "wp_eq_loc" := lazymatch goal with | |- envs_entails _ (wp ?s ?E ?e ?Q) => reshape_expr e ltac:(fun K e' => eapply (tac_wp_eq_loc K)); - [iSolveTC|iAssumptionCore|iAssumptionCore|simpl; try wp_value_head] + [tc_solve|iAssumptionCore|iAssumptionCore|simpl; try wp_value_head] | _ => fail "wp_pure: not a 'wp'" end. @@ -70,7 +71,7 @@ Tactic Notation "wp_case" := wp_pure (Case _ _); try wp_value_head. Lemma tac_wp_bind `{!lrustGS Σ} K Δ E Φ e : envs_entails Δ (WP e @ E {{ v, WP fill K (of_val v) @ E {{ Φ }} }})%I → envs_entails Δ (WP fill K e @ E {{ Φ }}). -Proof. rewrite envs_entails_eq=> ->. apply: wp_bind. Qed. +Proof. rewrite envs_entails_unseal=> ->. apply: wp_bind. Qed. Ltac wp_bind_core K := lazymatch eval hnf in K with @@ -99,9 +100,9 @@ Lemma tac_wp_nd_int `{!lrustGS Σ} K Δ Δ' E Φ : (∀z, envs_entails Δ' (WP fill K (Lit (LitInt z)) @ E {{ Φ }})) → envs_entails Δ (WP fill K NdInt @ E {{ Φ }}). Proof. - rewrite envs_entails_eq=> ? HΔ'. rewrite into_laterN_env_sound /=. - rewrite -wp_bind. eapply wand_apply; [by apply wp_nd_int|]. - iIntros "Δ'". iSplit; [done|]. iIntros "!>% _". by iApply HΔ'. + rewrite envs_entails_unseal=> ? HΔ'. rewrite into_laterN_env_sound /=. + rewrite -wp_bind. iIntros "Δ'". iApply (wp_nd_int with "[//] [-]"). + iIntros (z) "_ !>". by iApply HΔ'. Qed. Lemma tac_wp_alloc K Δ Δ' E j1 j2 n Φ : @@ -113,8 +114,8 @@ Lemma tac_wp_alloc K Δ Δ' E j1 j2 n Φ : envs_entails Δ'' (WP fill K (Lit $ LitLoc l) @ E {{ Φ }})) → envs_entails Δ (WP fill K (Alloc (Lit $ LitInt n)) @ E {{ Φ }}). Proof. - rewrite envs_entails_eq=> ?? HΔ. rewrite -wp_bind. - eapply wand_apply; first exact:wp_alloc. + rewrite envs_entails_unseal=> ?? HΔ. rewrite -wp_bind. + eapply wand_apply; first by apply wand_entails, wp_alloc. rewrite -persistent_and_sep. apply and_intro; first by auto. rewrite into_laterN_env_sound; apply later_mono, forall_intro=> l. apply forall_intro=>sz. apply wand_intro_l. rewrite -assoc. @@ -134,8 +135,8 @@ Lemma tac_wp_free K Δ Δ' Δ'' Δ''' E i1 i2 vl (n : Z) (n' : nat) l Φ : envs_entails Δ''' (WP fill K (Lit LitPoison) @ E {{ Φ }}) → envs_entails Δ (WP fill K (Free (Lit $ LitInt n) (Lit $ LitLoc l)) @ E {{ Φ }}). Proof. - rewrite envs_entails_eq; intros -> ?? <- ? <- -> HΔ. rewrite -wp_bind. - eapply wand_apply; first exact:wp_free; simpl. + rewrite envs_entails_unseal; intros -> ?? <- ? <- -> HΔ. rewrite -wp_bind. + eapply wand_apply; first by apply wand_entails, wp_free. rewrite into_laterN_env_sound -!later_sep; apply later_mono. do 2 (rewrite envs_lookup_sound //). by rewrite HΔ True_emp emp_wand -assoc. Qed. @@ -147,11 +148,11 @@ Lemma tac_wp_read K Δ Δ' E i l q v o Φ : envs_entails Δ' (WP fill K (of_val v) @ E {{ Φ }}) → envs_entails Δ (WP fill K (Read o (Lit $ LitLoc l)) @ E {{ Φ }}). Proof. - rewrite envs_entails_eq; intros [->| ->] ???. - - rewrite -wp_bind. eapply wand_apply; first exact:wp_read_na. + rewrite envs_entails_unseal; intros [->| ->] ???. + - rewrite -wp_bind. eapply wand_apply; first by apply wand_entails, wp_read_na. rewrite into_laterN_env_sound -later_sep envs_lookup_split //; simpl. by apply later_mono, sep_mono_r, wand_mono. - - rewrite -wp_bind. eapply wand_apply; first exact:wp_read_sc. + - rewrite -wp_bind. eapply wand_apply; first by apply wand_entails, wp_read_sc. rewrite into_laterN_env_sound -later_sep envs_lookup_split //; simpl. by apply later_mono, sep_mono_r, wand_mono. Qed. @@ -165,11 +166,11 @@ Lemma tac_wp_write K Δ Δ' Δ'' E i l v e v' o Φ : envs_entails Δ'' (WP fill K (Lit LitPoison) @ E {{ Φ }}) → envs_entails Δ (WP fill K (Write o (Lit $ LitLoc l) e) @ E {{ Φ }}). Proof. - rewrite envs_entails_eq; intros ? [->| ->] ????. - - rewrite -wp_bind. eapply wand_apply; first by apply wp_write_na. + rewrite envs_entails_unseal; intros ? [->| ->] ????. + - rewrite -wp_bind. eapply wand_apply; first by apply wand_entails, wp_write_na. rewrite into_laterN_env_sound -later_sep envs_simple_replace_sound //; simpl. rewrite right_id. by apply later_mono, sep_mono_r, wand_mono. - - rewrite -wp_bind. eapply wand_apply; first by apply wp_write_sc. + - rewrite -wp_bind. eapply wand_apply; first by apply wand_entails, wp_write_sc. rewrite into_laterN_env_sound -later_sep envs_simple_replace_sound //; simpl. rewrite right_id. by apply later_mono, sep_mono_r, wand_mono. Qed. @@ -193,7 +194,7 @@ Tactic Notation "wp_nd_int" ident(z) := | |- envs_entails _ (wp ?s ?E ?e ?Q) => reshape_expr e ltac:(fun K e' => unify e' NdInt; eapply (tac_wp_nd_int K); - [iSolveTC (* IntoLaters *) + [tc_solve (* IntoLaters *) |iIntros (z); simpl_subst; try wp_value_head (* new goal *)]) || fail "wp_nd_int: cannot find NdInt in" e | _ => fail "wp_nd_int: not a 'wp'" @@ -207,7 +208,7 @@ Tactic Notation "wp_alloc" ident(l) "as" constr(H) constr(Hf) := [reshape_expr e ltac:(fun K e' => eapply (tac_wp_alloc K _ _ _ H Hf)) |fail 1 "wp_alloc: cannot find 'Alloc' in" e]; [try fast_done - |iSolveTC + |tc_solve |let sz := fresh "sz" in let Hsz := fresh "Hsz" in first [intros l sz Hsz | fail 1 "wp_alloc:" l "not fresh"]; (* If Hsz is "constant Z = nat", change that to an equation on nat and @@ -235,7 +236,7 @@ Tactic Notation "wp_free" := [reshape_expr e ltac:(fun K e' => eapply (tac_wp_free K)) |fail 1 "wp_free: cannot find 'Free' in" e]; [try fast_done - |iSolveTC + |tc_solve |let l := match goal with |- _ = Some (_, (?l ↦∗ _)%I) => l end in iAssumptionCore || fail "wp_free: cannot find" l "↦∗ ?" |pm_reflexivity @@ -256,7 +257,7 @@ Tactic Notation "wp_read" := |fail 1 "wp_read: cannot find 'Read' in" e]; [(right; fast_done) || (left; fast_done) || fail "wp_read: order is neither Na2Ord nor ScOrd" - |iSolveTC + |tc_solve |let l := match goal with |- _ = Some (_, (?l ↦{_} _)%I) => l end in iAssumptionCore || fail "wp_read: cannot find" l "↦ ?" |simpl; try wp_value_head] @@ -268,11 +269,11 @@ Tactic Notation "wp_write" := lazymatch goal with | |- envs_entails _ (wp ?s ?E ?e ?Q) => first - [reshape_expr e ltac:(fun K e' => eapply (tac_wp_write K); [iSolveTC|..]) + [reshape_expr e ltac:(fun K e' => eapply (tac_wp_write K); [tc_solve|..]) |fail 1 "wp_write: cannot find 'Write' in" e]; [(right; fast_done) || (left; fast_done) || fail "wp_write: order is neither Na2Ord nor ScOrd" - |iSolveTC + |tc_solve |let l := match goal with |- _ = Some (_, (?l ↦{_} _)%I) => l end in iAssumptionCore || fail "wp_write: cannot find" l "↦ ?" |pm_reflexivity diff --git a/theories/lang/races.v b/theories/lang/races.v index 9c923d16..dcdba7f9 100644 --- a/theories/lang/races.v +++ b/theories/lang/races.v @@ -5,7 +5,7 @@ Set Default Proof Using "Type". Inductive access_kind : Type := ReadAcc | WriteAcc | FreeAcc. -(* This is a crucial definition; if we forget to sync it with head_step, +(* This is a crucial definition; if we forget to sync it with base_step, the results proven here are worthless. *) Inductive next_access_head : expr → state → access_kind * order → loc → Prop := | Access_read ord l σ : @@ -27,24 +27,24 @@ Inductive next_access_head : expr → state → access_kind * order → loc → σ (FreeAcc, Na2Ord) (l +ₗ i). (* Some sanity checks to make sure the definition above is not entirely insensible. *) -Goal ∀ e1 e2 e3 σ, head_reducible (CAS e1 e2 e3) σ → +Goal ∀ e1 e2 e3 σ, base_reducible (CAS e1 e2 e3) σ → ∃ a l, next_access_head (CAS e1 e2 e3) σ a l. Proof. intros ??? σ (?&?&?&?&?). inversion H; do 2 eexists; (eapply Access_cas_fail; done) || eapply Access_cas_suc; done. Qed. -Goal ∀ o e σ, head_reducible (Read o e) σ → +Goal ∀ o e σ, base_reducible (Read o e) σ → ∃ a l, next_access_head (Read o e) σ a l. Proof. intros ?? σ (?&?&?&?&?). inversion H; do 2 eexists; eapply Access_read; done. Qed. -Goal ∀ o e1 e2 σ, head_reducible (Write o e1 e2) σ → +Goal ∀ o e1 e2 σ, base_reducible (Write o e1 e2) σ → ∃ a l, next_access_head (Write o e1 e2) σ a l. Proof. intros ??? σ (?&?&?&?&?). inversion H; do 2 eexists; eapply Access_write; try done; eexists; done. Qed. -Goal ∀ e1 e2 σ, head_reducible (Free e1 e2) σ → +Goal ∀ e1 e2 σ, base_reducible (Free e1 e2) σ → ∃ a l, next_access_head (Free e1 e2) σ a l. Proof. intros ?? σ (?&?&?&?&?). inversion H; do 2 eexists; eapply Access_free; done. @@ -70,19 +70,19 @@ Definition nonracing_threadpool (el : list expr) (σ : state) : Prop := ∀ l a1 a2, next_accesses_threadpool el σ a1 a2 l → nonracing_accesses a1 a2. -Lemma next_access_head_reductible_ctx e σ σ' a l K : - next_access_head e σ' a l → reducible (fill K e) σ → head_reducible e σ. +Lemma next_access_base_reducible_ctx e σ σ' a l K : + next_access_head e σ' a l → reducible (fill K e) σ → base_reducible e σ. Proof. - intros Hhead Hred. apply prim_head_reducible. + intros Hhead Hred. apply prim_base_reducible. - eapply (reducible_fill_inv (K:=ectx_language.fill K)), Hred. destruct Hhead; eauto. - apply ectxi_language_sub_redexes_are_values. intros [] ? ->; inversion Hhead; eauto. Qed. Definition head_reduce_not_to_stuck (e : expr) (σ : state) := - ∀ κ e' σ' efs, ectx_language.head_step e σ κ e' σ' efs → e' ≠ App (Lit $ LitInt 0) []. + ∀ κ e' σ' efs, ectx_language.base_step e σ κ e' σ' efs → e' ≠ App (Lit $ LitInt 0) []. -Lemma next_access_head_reducible_state e σ a l : - next_access_head e σ a l → head_reducible e σ → +Lemma next_access_base_reducible_state e σ a l : + next_access_head e σ a l → base_reducible e σ → head_reduce_not_to_stuck e σ → match a with | (ReadAcc, (ScOrd | Na1Ord)) => ∃ v n, σ !! l = Some (RSt n, v) @@ -92,7 +92,7 @@ Lemma next_access_head_reducible_state e σ a l : | (FreeAcc, _) => ∃ v ls, σ !! l = Some (ls, v) end. Proof. - intros Ha (κ&e'&σ'&ef&Hstep) Hrednonstuck. destruct Ha; inv_head_step; eauto. + intros Ha (κ&e'&σ'&ef&Hstep) Hrednonstuck. destruct Ha; inv_base_step; eauto. - destruct n; first by eexists. exfalso. eapply Hrednonstuck; last done. by eapply CasStuckS. - exfalso. eapply Hrednonstuck; last done. @@ -102,19 +102,19 @@ Qed. Lemma next_access_head_Na1Ord_step e1 e2 ef σ1 σ2 κ a l : next_access_head e1 σ1 (a, Na1Ord) l → - head_step e1 σ1 κ e2 σ2 ef → + base_step e1 σ1 κ e2 σ2 ef → next_access_head e2 σ2 (a, Na2Ord) l. Proof. - intros Ha Hstep. inversion Ha; subst; clear Ha; inv_head_step; constructor; by rewrite to_of_val. + intros Ha Hstep. inversion Ha; subst; clear Ha; inv_base_step; constructor; by rewrite to_of_val. Qed. Lemma next_access_head_Na1Ord_concurent_step e1 e1' e2 e'f σ σ' κ a1 a2 l : next_access_head e1 σ (a1, Na1Ord) l → - head_step e1 σ κ e1' σ' e'f → + base_step e1 σ κ e1' σ' e'f → next_access_head e2 σ a2 l → next_access_head e2 σ' a2 l. Proof. - intros Ha1 Hstep Ha2. inversion Ha1; subst; clear Ha1; inv_head_step; + intros Ha1 Hstep Ha2. inversion Ha1; subst; clear Ha1; inv_base_step; destruct Ha2; simplify_eq; econstructor; eauto; try apply lookup_insert. (* Oh my. FIXME. *) - eapply lit_eq_state; last done. @@ -126,8 +126,8 @@ Proof. Qed. Lemma next_access_head_Free_concurent_step e1 e1' e2 e'f σ σ' κ o1 a2 l : - next_access_head e1 σ (FreeAcc, o1) l → head_step e1 σ κ e1' σ' e'f → - next_access_head e2 σ a2 l → head_reducible e2 σ' → + next_access_head e1 σ (FreeAcc, o1) l → base_step e1 σ κ e1' σ' e'f → + next_access_head e2 σ a2 l → base_reducible e2 σ' → False. Proof. intros Ha1 Hstep Ha2 Hred2. @@ -139,18 +139,18 @@ Proof. - subst. by rewrite /shift_loc Z.add_0_r -surjective_pairing lookup_delete. - replace i with (1+(i-1)) by lia. rewrite lookup_delete_ne -shift_loc_assoc ?IH //. lia. - destruct l; intros [=?]. lia. } + destruct l; intros [=]. lia. } assert (FREE' : σ' !! l = None). - { inversion Ha1; clear Ha1; inv_head_step. auto. } + { inversion Ha1; clear Ha1; inv_base_step. auto. } destruct Hred2 as (κ'&e2'&σ''&ef&?). - inversion Ha2; clear Ha2; inv_head_step. + inversion Ha2; clear Ha2; inv_base_step. eapply (@is_Some_None (lock_state * val)). rewrite -FREE'. naive_solver. Qed. Lemma safe_step_not_reduce_to_stuck el σ K e e' σ' ef κ : (∀ el' σ' e', rtc erased_step (el, σ) (el', σ') → e' ∈ el' → to_val e' = None → reducible e' σ') → - fill K e ∈ el → head_step e σ κ e' σ' ef → head_reduce_not_to_stuck e' σ'. + fill K e ∈ el → base_step e σ κ e' σ' ef → head_reduce_not_to_stuck e' σ'. Proof. intros Hsafe Hi Hstep κ1 e1 σ1 ? Hstep1 Hstuck. cut (reducible (fill K e1) σ1). @@ -187,12 +187,12 @@ Proof. intros Hsafe l a1 a2 (t1&?&t2&?&t3&->&(K1&e1&Ha1&->)&(K2&e2&Ha2&->)). assert (to_val e1 = None). by destruct Ha1. - assert (Hrede1:head_reducible e1 σ). - { eapply (next_access_head_reductible_ctx e1 σ σ a1 l K1), Hsafe, fill_not_val=>//. + assert (Hrede1:base_reducible e1 σ). + { eapply (next_access_base_reducible_ctx e1 σ σ a1 l K1), Hsafe, fill_not_val=>//. abstract set_solver. } assert (Hnse1:head_reduce_not_to_stuck e1 σ). { eapply safe_not_reduce_to_stuck with (K:=K1); first done. set_solver+. } - assert (Hσe1:=next_access_head_reducible_state _ _ _ _ Ha1 Hrede1 Hnse1). + assert (Hσe1:=next_access_base_reducible_state _ _ _ _ Ha1 Hrede1 Hnse1). destruct Hrede1 as (κ1'&e1'&σ1'&ef1&Hstep1). clear Hnse1. assert (Ha1' : a1.2 = Na1Ord → next_access_head e1' σ1' (a1.1, Na2Ord) l). { intros. eapply next_access_head_Na1Ord_step, Hstep1. @@ -202,21 +202,21 @@ Proof. (t1 ++ fill K1 e1' :: t2 ++ fill K2 e2 :: t3 ++ ef1, σ1')). { eapply rtc_l, rtc_refl. eexists. eapply step_atomic, Ectx_step, Hstep1; try done. by rewrite -assoc. } - assert (Hrede1: a1.2 = Na1Ord → head_reducible e1' σ1'). + assert (Hrede1: a1.2 = Na1Ord → base_reducible e1' σ1'). { destruct a1 as [a1 []]=>// _. - eapply (next_access_head_reductible_ctx e1' σ1' σ1' (a1, Na2Ord) l K1), Hsafe, + eapply (next_access_base_reducible_ctx e1' σ1' σ1' (a1, Na2Ord) l K1), Hsafe, fill_not_val=>//. by auto. abstract set_solver. by destruct Hstep1; inversion Ha1. } assert (Hnse1: head_reduce_not_to_stuck e1' σ1'). { eapply safe_step_not_reduce_to_stuck with (K:=K1); first done; last done. set_solver+. } assert (to_val e2 = None). by destruct Ha2. - assert (Hrede2:head_reducible e2 σ). - { eapply (next_access_head_reductible_ctx e2 σ σ a2 l K2), Hsafe, fill_not_val=>//. + assert (Hrede2:base_reducible e2 σ). + { eapply (next_access_base_reducible_ctx e2 σ σ a2 l K2), Hsafe, fill_not_val=>//. abstract set_solver. } assert (Hnse2:head_reduce_not_to_stuck e2 σ). { eapply safe_not_reduce_to_stuck with (K:=K2); first done. set_solver+. } - assert (Hσe2:=next_access_head_reducible_state _ _ _ _ Ha2 Hrede2 Hnse2). + assert (Hσe2:=next_access_base_reducible_state _ _ _ _ Ha2 Hrede2 Hnse2). destruct Hrede2 as (κ2'&e2'&σ2'&ef2&Hstep2). assert (Ha2' : a2.2 = Na1Ord → next_access_head e2' σ2' (a2.1, Na2Ord) l). { intros. eapply next_access_head_Na1Ord_step, Hstep2. @@ -226,9 +226,9 @@ Proof. (t1 ++ fill K1 e1 :: t2 ++ fill K2 e2' :: t3 ++ ef2, σ2')). { eapply rtc_l, rtc_refl. rewrite app_comm_cons assoc. eexists. eapply step_atomic, Ectx_step, Hstep2; try done. by rewrite -assoc. } - assert (Hrede2: a2.2 = Na1Ord → head_reducible e2' σ2'). + assert (Hrede2: a2.2 = Na1Ord → base_reducible e2' σ2'). { destruct a2 as [a2 []]=>// _. - eapply (next_access_head_reductible_ctx e2' σ2' σ2' (a2, Na2Ord) l K2), Hsafe, + eapply (next_access_base_reducible_ctx e2' σ2' σ2' (a2, Na2Ord) l K2), Hsafe, fill_not_val=>//. by auto. abstract set_solver. by destruct Hstep2; inversion Ha2. } assert (Hnse2':head_reduce_not_to_stuck e2' σ2'). @@ -237,32 +237,32 @@ Proof. assert (Ha1'2 : a1.2 = Na1Ord → next_access_head e2 σ1' a2 l). { intros HNa. eapply next_access_head_Na1Ord_concurent_step=>//. by rewrite <-HNa, <-surjective_pairing. } - assert (Hrede1_2: head_reducible e2 σ1'). - { intros. eapply (next_access_head_reductible_ctx e2 σ1' σ a2 l K2), Hsafe, fill_not_val=>//. + assert (Hrede1_2: base_reducible e2 σ1'). + { intros. eapply (next_access_base_reducible_ctx e2 σ1' σ a2 l K2), Hsafe, fill_not_val=>//. abstract set_solver. } assert (Hnse1_2:head_reduce_not_to_stuck e2 σ1'). { eapply safe_not_reduce_to_stuck with (K:=K2). - intros. eapply Hsafe. etrans; last done. done. done. done. - set_solver+. } assert (Hσe1'1:= - λ Hord, next_access_head_reducible_state _ _ _ _ (Ha1' Hord) (Hrede1 Hord) Hnse1). + λ Hord, next_access_base_reducible_state _ _ _ _ (Ha1' Hord) (Hrede1 Hord) Hnse1). assert (Hσe1'2:= - λ Hord, next_access_head_reducible_state _ _ _ _ (Ha1'2 Hord) Hrede1_2 Hnse1_2). + λ Hord, next_access_base_reducible_state _ _ _ _ (Ha1'2 Hord) Hrede1_2 Hnse1_2). assert (Ha2'1 : a2.2 = Na1Ord → next_access_head e1 σ2' a1 l). { intros HNa. eapply next_access_head_Na1Ord_concurent_step=>//. by rewrite <-HNa, <-surjective_pairing. } - assert (Hrede2_1: head_reducible e1 σ2'). - { intros. eapply (next_access_head_reductible_ctx e1 σ2' σ a1 l K1), Hsafe, fill_not_val=>//. + assert (Hrede2_1: base_reducible e1 σ2'). + { intros. eapply (next_access_base_reducible_ctx e1 σ2' σ a1 l K1), Hsafe, fill_not_val=>//. abstract set_solver. } assert (Hnse2_1:head_reduce_not_to_stuck e1 σ2'). { eapply safe_not_reduce_to_stuck with (K:=K1). - intros. eapply Hsafe. etrans; last done. done. done. done. - set_solver+. } assert (Hσe2'1:= - λ Hord, next_access_head_reducible_state _ _ _ _ (Ha2'1 Hord) Hrede2_1 Hnse2_1). + λ Hord, next_access_base_reducible_state _ _ _ _ (Ha2'1 Hord) Hrede2_1 Hnse2_1). assert (Hσe2'2:= - λ Hord, next_access_head_reducible_state _ _ _ _ (Ha2' Hord) (Hrede2 Hord) Hnse2'). + λ Hord, next_access_base_reducible_state _ _ _ _ (Ha2' Hord) (Hrede2 Hord) Hnse2'). assert (a1.1 = FreeAcc → False). { destruct a1 as [[]?]; inversion 1. @@ -281,7 +281,7 @@ Proof. clear κ2' e2' Hnse2' Hnse2_1 Hstep2 σ2' Hrtc_step2 Hrede2_1. destruct Hrede1_2 as (κ2'&e2'&σ'&ef&?). - inv_head_step. + inv_base_step. match goal with | H : <[l:=_]> σ !! l = _ |- _ => by rewrite lookup_insert in H end. diff --git a/theories/lang/tactics.v b/theories/lang/tactics.v index b0a40627..ecc110cc 100644 --- a/theories/lang/tactics.v +++ b/theories/lang/tactics.v @@ -219,12 +219,12 @@ Ltac simpl_subst := Arguments W.to_expr : simpl never. Arguments subst : simpl never. -(** The tactic [inv_head_step] performs inversion on hypotheses of the -shape [head_step]. The tactic will discharge head-reductions starting +(** The tactic [inv_base_step] performs inversion on hypotheses of the +shape [base_step]. The tactic will discharge head-reductions starting from values, and simplifies hypothesis related to conversions from and to values, and finite map operations. This tactic is slightly ad-hoc and tuned for proving our lifting lemmas. *) -Ltac inv_head_step := +Ltac inv_base_step := repeat match goal with | _ => progress simplify_map_eq/= (* simplify memory stuff *) | H : to_val _ = Some _ |- _ => apply of_to_val in H @@ -232,7 +232,7 @@ Ltac inv_head_step := apply (f_equal (to_val)) in H; rewrite to_of_val in H; injection H; clear H; intro | H : context [to_val (of_val _)] |- _ => rewrite to_of_val in H - | H : head_step ?e _ _ _ _ _ |- _ => + | H : base_step ?e _ _ _ _ _ |- _ => try (is_var e; fail 1); (* inversion yields many goals if [e] is a variable and can thus better be avoided. *) inversion H; subst; clear H diff --git a/theories/lang/time.v b/theories/lang/time.v index c40fd603..62e74abe 100644 --- a/theories/lang/time.v +++ b/theories/lang/time.v @@ -7,16 +7,16 @@ Set Default Proof Using "Type". Import uPred. Class timeGS Σ := TimeG { - time_mono_nat_inG :> inG Σ mono_natR; - time_nat_inG :> inG Σ (authR natUR); + time_mono_nat_inG :: inG Σ mono_natR; + time_nat_inG :: inG Σ (authR natUR); time_global_name : gname; time_persistent_name : gname; time_cumulative_name : gname; }. Class timePreG Σ := TimePreG { - time_preG_mono_nat_inG :> inG Σ mono_natR; - time_preG_nat_inG :> inG Σ (authR natUR); + time_preG_mono_nat_inG :: inG Σ mono_natR; + time_preG_nat_inG :: inG Σ (authR natUR); }. Definition timeΣ : gFunctors := @@ -68,18 +68,18 @@ Section time. Lemma time_interp_step n : time_interp n ==∗ time_interp (S n). - Proof. eapply own_update, mono_nat_update. lia. Qed. + Proof. iApply own_update. apply mono_nat_update. lia. Qed. Lemma persistent_time_receipt_mono n m : (n ≤ m)%nat → ⧖m -∗ ⧖n. Proof. - intros ?. unfold persistent_time_receipt. by apply own_mono, mono_nat_lb_mono. + intros ?. unfold persistent_time_receipt. iApply own_mono. by apply mono_nat_lb_mono. Qed. Lemma cumulative_time_receipt_mono n m : (n ≤ m)%nat → ⧗m -∗ ⧗n. Proof. intros ?. unfold cumulative_time_receipt. - by apply own_mono, auth_frag_mono, nat_included. + iApply own_mono. by apply auth_frag_mono, nat_included. Qed. Lemma persistent_time_receipt_sep n m : ⧖(n `max` m) ⊣⊢ ⧖n ∗ ⧖m. @@ -87,14 +87,19 @@ Section time. Lemma cumulative_time_receipt_sep n m : ⧗(n + m) ⊣⊢ ⧗n ∗ ⧗m. Proof. by rewrite /cumulative_time_receipt -nat_op auth_frag_op own_op. Qed. - Global Instance persistent_time_receipt_from_sep n m : FromSep ⧖(n `max` m) ⧖n ⧖m. - Proof. rewrite /FromSep. by rewrite persistent_time_receipt_sep. Qed. Global Instance persistent_time_receipt_into_sep n m : IntoSep ⧖(n `max` m) ⧖n ⧖m. Proof. rewrite /IntoSep. by rewrite persistent_time_receipt_sep. Qed. - Global Instance cumulative_time_receipt_from_sep n m : FromSep ⧗(n + m) ⧗n ⧗m. - Proof. rewrite /FromSep. by rewrite cumulative_time_receipt_sep. Qed. + Global Instance persistent_time_receipt_from_sep n m : FromSep ⧖(n `max` m) ⧖n ⧖m. + Proof. rewrite /FromSep. by rewrite persistent_time_receipt_sep. Qed. + Global Instance persistent_time_receipt_combine_sep_as n m : CombineSepAs ⧖n ⧖m ⧖(n `max` m). + Proof. rewrite /CombineSepAs. by rewrite persistent_time_receipt_sep. Qed. + Global Instance cumulative_time_receipt_into_sep n m : IntoSep ⧗(n + m) ⧗n ⧗m. Proof. rewrite /IntoSep. by rewrite cumulative_time_receipt_sep. Qed. + Global Instance cumulative_time_receipt_from_sep n m : FromSep ⧗(n + m) ⧗n ⧗m. + Proof. rewrite /FromSep. by rewrite cumulative_time_receipt_sep. Qed. + Global Instance cumulative_time_receipt_combine_sep_as n m : CombineSepAs ⧗n ⧗m ⧗(n + m). + Proof. rewrite /CombineSepAs. by rewrite cumulative_time_receipt_sep. Qed. Lemma persistent_time_receipt_0 : ⊢ |==> ⧖0. Proof. rewrite /persistent_time_receipt. apply own_unit. Qed. @@ -147,7 +152,7 @@ Section time. iDestruct (own_valid_2 with "Hm0 Hm") as %?%mono_nat_both_valid. iDestruct (own_valid_2 with "Hm'0 Hm'") as %[?%nat_included _]%auth_both_valid_discrete. - iModIntro. iFrame. iSplitL; auto with iFrame lia. + iModIntro. iFrame. auto with lia. Qed. End time. diff --git a/theories/lifetime/at_borrow.v b/theories/lifetime/at_borrow.v index a0c5e574..8d8c29cf 100644 --- a/theories/lifetime/at_borrow.v +++ b/theories/lifetime/at_borrow.v @@ -101,6 +101,10 @@ Lemma at_bor_acc_lftN `{!invGS Σ, !lftGS Σ} (P : iProp Σ) E κ : ↑lftN ⊆ E → lft_ctx -∗ &at{κ,lftN}P ={E,E∖↑lftN}=∗ ▷P ∗ (▷P ={E∖↑lftN,E}=∗ True) ∨ [†κ] ∗ |={E∖↑lftN,E}=> True. -Proof. intros. by rewrite (at_bor_acc _ lftN) // difference_twice_L. Qed. +Proof. + iIntros (?) "H". + iDestruct (at_bor_acc _ lftN with "H") as "H"; [done..|]. + by rewrite difference_twice_L. +Qed. Global Typeclasses Opaque at_bor. diff --git a/theories/lifetime/frac_borrow.v b/theories/lifetime/frac_borrow.v index 2d377558..58ba554c 100644 --- a/theories/lifetime/frac_borrow.v +++ b/theories/lifetime/frac_borrow.v @@ -4,7 +4,7 @@ From iris.algebra Require Import frac. From lrust.lifetime Require Export at_borrow. Set Default Proof Using "Type". -Class frac_borG Σ := frac_borG_inG :> inG Σ fracR. +Class frac_borG Σ := frac_borG_inG :: inG Σ fracR. Local Definition frac_bor_inv `{!invGS Σ, !lftGS Σ, !frac_borG Σ} (φ : Qp → iProp Σ) γ κ' := (∃ q, φ q ∗ own γ q ∗ (⌜q = 1%Qp⌝ ∨ ∃ q', ⌜(q + q' = 1)%Qp⌝ ∗ q'.[κ']))%I. @@ -75,16 +75,16 @@ Section frac_bor. iIntros "#Hφ (H & Hown & Hφ1)". iNext. iDestruct "H" as (qφ) "(Hφqφ & Hown' & [%|Hq])". { subst. iCombine "Hown'" "Hown" as "Hown". - by iDestruct (own_valid with "Hown") as %Hval%Qp_not_add_le_l. } + by iDestruct (own_valid with "Hown") as %Hval%Qp.not_add_le_l. } rewrite /frac_bor_inv. iApply bi.sep_exist_r. iExists (q + qφ)%Qp. iDestruct "Hq" as (q' Hqφq') "Hq'κ". iCombine "Hown'" "Hown" as "Hown". iDestruct (own_valid with "Hown") as %Hval. rewrite comm_L. iFrame "Hown". iCombine "Hφ1 Hφqφ" as "Hφq". iDestruct ("Hφ" with "Hφq") as "$". - assert (q ≤ q')%Qp as [[r ->]%Qp_lt_sum|<-]%Qp_le_lteq. - { apply (Qp_add_le_mono_l _ _ qφ). by rewrite Hqφq'. } + assert (q ≤ q')%Qp as [[r ->]%Qp.lt_sum|<-]%Qp.le_lteq. + { apply (Qp.add_le_mono_l _ _ qφ). by rewrite Hqφq'. } - iDestruct "Hq'κ" as "[$ Hr]". iRight. iExists _. iIntros "{$Hr} !%". - by rewrite (comm_L Qp_add q) -assoc_L. + by rewrite (comm_L Qp.add q) -assoc_L. - iFrame "Hq'κ". iLeft. iPureIntro. rewrite comm_L. done. Qed. @@ -95,7 +95,7 @@ Section frac_bor. Proof. iIntros "#Hφ [H Hκ']". iNext. iDestruct "H" as (qφ) "(Hφqφ & Hown & Hq)". - destruct (Qp_lower_bound qκ' qφ) as (qq & qκ'0 & qφ0 & Hqκ' & Hqφ). + destruct (Qp.lower_bound qκ' qφ) as (qq & qκ'0 & qφ0 & Hqκ' & Hqφ). iApply bi.sep_exist_l. iExists qq. iApply bi.sep_exist_l. iExists qκ'0. subst qκ' qφ. rewrite /frac_bor_inv. @@ -160,6 +160,8 @@ Lemma frac_bor_lft_incl `{!invGS Σ, !lftGS Σ, !frac_borG Σ} κ κ' q: Proof. iIntros "#LFT#Hbor". iApply lft_incl_intro. iModIntro. iSplitR. - iIntros (q') "Hκ'". + assert (Fractional (λ q' : Qp, (q * q').[κ']))%I. + { intros ??. rewrite Qp.mul_add_distr_l lft_tok_fractional //. } iMod (frac_bor_acc with "LFT Hbor Hκ'") as (q'') "[>? Hclose]". done. iExists _. iFrame. iIntros "!>Hκ'". iApply "Hclose". auto. - iIntros "H†'". diff --git a/theories/lifetime/lifetime_sig.v b/theories/lifetime/lifetime_sig.v index 5feadc30..426975fb 100644 --- a/theories/lifetime/lifetime_sig.v +++ b/theories/lifetime/lifetime_sig.v @@ -85,16 +85,16 @@ Module Type lifetime_sig. Global Declare Instance lft_tok_fractional κ : Fractional (λ q, q.[κ])%I. Global Declare Instance lft_tok_as_fractional κ q : AsFractional q.[κ] (λ q, q.[κ])%I q. - Global Declare Instance frame_lft_tok p κ q1 q2 RES : - FrameFractionalHyps p q1.[κ] (λ q, q.[κ])%I RES q1 q2 → - Frame p q1.[κ] q2.[κ] RES | 5. + Global Declare Instance frame_lft_tok p κ q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p q1.[κ] q2.[κ] q.[κ] | 5. Global Declare Instance idx_bor_own_fractional i : Fractional (λ q, idx_bor_own q i)%I. Global Declare Instance idx_bor_own_as_fractional i q : AsFractional (idx_bor_own q i) (λ q, idx_bor_own q i)%I q. - Global Declare Instance frame_idx_bor_own p i q1 q2 RES : - FrameFractionalHyps p (idx_bor_own q1 i) (λ q, idx_bor_own q i)%I RES q1 q2 → - Frame p (idx_bor_own q1 i) (idx_bor_own q2 i) RES | 5. + Global Declare Instance frame_idx_bor_own p i q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p (idx_bor_own q1 i) (idx_bor_own q2 i) (idx_bor_own q i) | 5. (** Laws *) Parameter lft_tok_sep : ∀ q κ1 κ2, q.[κ1] ∗ q.[κ2] ⊣⊢ q.[κ1 ⊓ κ2]. diff --git a/theories/lifetime/model/accessors.v b/theories/lifetime/model/accessors.v index ff73ee32..acab5721 100644 --- a/theories/lifetime/model/accessors.v +++ b/theories/lifetime/model/accessors.v @@ -130,16 +130,17 @@ Proof. iDestruct ("HPP'" with "HP'") as "$". iApply fupd_mask_intro; [solve_ndisj|]. iIntros "Hclose HP'". iDestruct ("HPP'" with "HP'") as "HP". - iMod "Hclose" as "_". iMod (slice_fill _ _ true with "Hs HP Hbox") as "Hbox". - solve_ndisj. by rewrite lookup_insert. iFrame. - iApply "Hclose'". iExists _, _. iFrame. rewrite big_sepS_later. - iApply "Hclose''". iLeft. iFrame "%". iExists _, _. iFrame. - iExists _. iFrame. + iMod "Hclose" as "_". iMod (slice_fill _ _ true with "Hs HP Hbox") as "Hbox"; + first solve_ndisj. + { by rewrite lookup_insert. } + iFrame. + iApply "Hclose'". iFrame. rewrite big_sepS_later. + iApply "Hclose''". iLeft. iFrame "% ∗". rewrite insert_insert -(fmap_insert _ _ _ Bor_in) insert_id //. - - iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]". done. + - iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]"; first done. iMod ("Hclose'" with "[-Hbor]") as "_". - + iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto. - + iMod (lft_incl_dead with "Hle H†"). done. iFrame. + + iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto. + + iMod (lft_incl_dead with "Hle H†"); first done. iFrame. iApply fupd_mask_subseteq. solve_ndisj. Qed. @@ -249,12 +250,13 @@ Proof. iMod ("Hclose'" with "[- H◯]"); last first. { iModIntro. iExists (i.1). iSplit; first by iApply lft_incl_refl. iExists j. iFrame. iExists Q. rewrite -bi.iff_refl. auto. } - iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". - iLeft. iFrame "%". iExists _, _. iFrame. iNext. + iFrame. rewrite big_sepS_later. iApply "Hclose''". + iLeft. iFrame "%". iExists _, _. rewrite -!fmap_delete -fmap_insert -(fmap_insert _ _ _ Bor_in). - iExists _. iFrame "Hbox H●". - rewrite big_sepM_insert. by rewrite big_sepM_delete. - by rewrite -fmap_None -lookup_fmap fmap_delete. + iFrame. + rewrite big_sepM_insert; last first. + { by rewrite -fmap_None -lookup_fmap fmap_delete. } + by rewrite big_sepM_delete. - iRight. iMod (lft_dead_in_tok with "HA") as "[HA #H†]". done. iMod ("Hclose'" with "[-Hbor]") as "_". + iExists _, _. iFrame. rewrite big_sepS_later. iApply "Hclose''". eauto. diff --git a/theories/lifetime/model/borrow.v b/theories/lifetime/model/borrow.v index d12f215b..6b45c08d 100644 --- a/theories/lifetime/model/borrow.v +++ b/theories/lifetime/model/borrow.v @@ -42,8 +42,7 @@ Proof. simpl. iFrame. } iMod ("Hclose" with "[HA HI Hclose']") as "_"; [by iNext; iExists _, _; iFrame|]. iSplitL "HB◯ HsliceB". - + rewrite /bor /raw_bor /idx_bor_own. iModIntro. iExists γB. iFrame. - iExists P. rewrite -bi.iff_refl. auto. + + rewrite /bor /raw_bor /idx_bor_own. iModIntro. iExists γB. iFrame. auto. + clear -HE. iIntros "!> H†". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". iDestruct ("HIlookup" with "HI") as %Hκ. diff --git a/theories/lifetime/model/borrow_sep.v b/theories/lifetime/model/borrow_sep.v index c961c617..1cd04ff3 100644 --- a/theories/lifetime/model/borrow_sep.v +++ b/theories/lifetime/model/borrow_sep.v @@ -50,14 +50,12 @@ Proof. -fmap_None -lookup_fmap fmap_delete //. } rewrite !own_bor_op. iDestruct "Hbor" as "[[H● H◯2] H◯1]". iAssert (&{κ}P)%I with "[H◯1 Hs1]" as "$". - { rewrite /bor /raw_bor /idx_bor_own. iExists κ'. iFrame "#". iExists γ1. - iFrame. iExists P. rewrite -bi.iff_refl. auto. } + { rewrite /bor /raw_bor /idx_bor_own. iFrame "# ∗". auto. } iAssert (&{κ}Q)%I with "[H◯2 Hs2]" as "$". - { rewrite /bor /raw_bor /idx_bor_own. iExists κ'. iFrame "#". iExists γ2. - iFrame. iExists Q. rewrite -bi.iff_refl. auto. } - iApply "Hclose". iExists A, I. iFrame. rewrite big_sepS_later. - iApply "Hclose'". iLeft. iFrame "%". iExists Pb', Pi. iFrame. iExists _. - rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "Hbox H●". + { rewrite /bor /raw_bor /idx_bor_own. iFrame "# ∗". auto. } + iApply "Hclose". iFrame. rewrite big_sepS_later. + iApply "Hclose'". iLeft. + rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "%∗". rewrite !big_sepM_insert /=. + by rewrite left_id -(big_sepM_delete _ _ _ _ EQB). + by rewrite -fmap_None -lookup_fmap fmap_delete. @@ -119,12 +117,11 @@ Proof. rewrite own_bor_op. iDestruct "Hbor" as "[H● H◯]". iAssert (&{κ}(P ∗ Q))%I with "[H◯ Hslice]" as "$". { rewrite /bor /raw_bor /idx_bor_own. iExists (κ1 ⊓ κ2). - iSplit; first by iApply (lft_incl_glb with "Hκ1 Hκ2"). - iExists γ. iFrame. iExists _. iFrame. iNext. iModIntro. + iSplit; first by iApply (lft_incl_glb with "Hκ1 Hκ2"). iFrame. iNext. iModIntro. by iSplit; iIntros "[HP HQ]"; iSplitL "HP"; (iApply "HPP'" || iApply "HQQ'"). } - iApply "Hclose". iExists A, I. iFrame. rewrite big_sepS_later. - iApply "Hclose'". iLeft. iFrame "%". iExists Pb, Pi. iFrame. iExists _. - rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "Hbox H●". + iApply "Hclose". iFrame. rewrite big_sepS_later. + iApply "Hclose'". iLeft. + rewrite /to_borUR -!fmap_delete -!fmap_insert. iFrame "%∗". rewrite !big_sepM_insert /=. + rewrite (big_sepM_delete _ _ _ _ EQB1) /=. iNext. simpl. rewrite [([∗ map] _ ∈ delete _ _, _)%I](big_sepM_delete _ _ j2 Bor_in) /=. diff --git a/theories/lifetime/model/creation.v b/theories/lifetime/model/creation.v index f6e00a1b..16e4e64a 100644 --- a/theories/lifetime/model/creation.v +++ b/theories/lifetime/model/creation.v @@ -37,7 +37,7 @@ Proof. { (* FIXME [solve_ndisj] fails *) set_solver-. } { intros i s. by rewrite lookup_fmap fmap_Some=> -[? [/HB -> ->]]. } rewrite lft_vs_unfold; iDestruct "Hvs" as (n) "[Hcnt Hvs]". - iDestruct (big_sepS_filter_acc (.⊂ κ) _ _ (dom _ I) with "Halive") + iDestruct (big_sepS_filter_acc (.⊂ κ) _ _ (dom I) with "Halive") as "[Halive Halive']". { intros κ'. rewrite elem_of_dom. eauto. } iApply fupd_trans. iApply fupd_mask_mono; last @@ -51,7 +51,7 @@ Proof. rewrite /Iinv. iFrame "Hdead Halive' HI". iModIntro. iMod (lft_inh_kill with "[$Hinh $HQ]"); first set_solver+. iModIntro. rewrite /lft_inv_dead. iExists Q. iFrame. - rewrite /lft_bor_dead. iExists (dom _ B), P. + rewrite /lft_bor_dead. iExists (dom B), P. rewrite !gset_to_gmap_dom -map_fmap_compose. rewrite (map_fmap_ext _ ((1%Qp,.) ∘ to_agree) B); last naive_solver. iFrame. @@ -73,12 +73,12 @@ Proof. { iModIntro. rewrite /Iinv. iFrame. iApply (@big_sepS_impl with "[$HK]"); iModIntro. rewrite /lft_inv. iIntros (κ Hκ) "[[[_ %]|[$ _]] _]". set_solver. } - destruct (minimal_exists_L (⊂) Kalive) + destruct (minimal_exists_elem_of_L (⊂) Kalive) as (κ & [Hκalive HκK]%elem_of_filter & Hκmin); first done. iDestruct (@big_sepS_delete with "HK") as "[[Hκinv Hκ] HK]"; first done. iDestruct (lft_inv_alive_in with "Hκinv") as "Hκalive"; first done. assert (κ ∉ K') as HκK' by set_solver +HκK HKK'. - specialize (IH (K ∖ {[ κ ]})). feed specialize IH; [set_solver +HκK|]. + ospecialize (IH (K ∖ {[ κ ]}) _); [set_solver +HκK|]. iMod (IH ({[ κ ]} ∪ K') with "[HI Halive Hκalive] HK") as "[[HI Halive] Hdead]". { set_solver +HKK'. } { intros κ' κ''. @@ -103,7 +103,7 @@ Proof. Qed. Definition kill_set (I : gmap lft lft_names) (Λ : atomic_lft) : gset lft := - filter (Λ ∈.) (dom (gset lft) I). + filter (Λ ∈.) (dom I). Lemma elem_of_kill_set I Λ κ : κ ∈ kill_set I Λ ↔ Λ ∈ κ ∧ is_Some (I !! κ). Proof. by rewrite /kill_set elem_of_filter elem_of_dom. Qed. @@ -114,7 +114,7 @@ Lemma lft_create E : Proof. iIntros (?) "#LFT". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". - destruct (exist_fresh (dom (gset atomic_lft) A)) as [Λ HΛ%not_elem_of_dom]. + destruct (exist_fresh (dom A)) as [Λ HΛ%not_elem_of_dom]. iMod (own_update with "HA") as "[HA HΛ]". { apply auth_update_alloc, (alloc_singleton_local_update _ Λ (Cinl 1%Qp))=>//. by rewrite lookup_fmap HΛ. } @@ -142,8 +142,8 @@ Proof. (exclusive_local_update _ (Cinr ())). } iDestruct "HΛ" as "#HΛ". iModIntro; iNext. pose (K := kill_set I Λ). - pose (K' := filter (lft_alive_in A) (dom (gset lft) I) ∖ K). - destruct (proj1 (subseteq_disjoint_union_L (K ∪ K') (dom (gset lft) I))) as (K''&HI&HK''). + pose (K' := filter (lft_alive_in A) (dom I) ∖ K). + destruct (proj1 (subseteq_disjoint_union_L (K ∪ K') (dom I))) as (K''&HI&HK''). { set_solver+. } assert (K ## K') by set_solver+. rewrite HI !big_sepS_union //. iDestruct "Hinv" as "[[HinvK HinvD] Hinv]". diff --git a/theories/lifetime/model/definitions.v b/theories/lifetime/model/definitions.v index 63070fd0..775c4600 100644 --- a/theories/lifetime/model/definitions.v +++ b/theories/lifetime/model/definitions.v @@ -39,24 +39,24 @@ Definition borUR := gmapUR slice_name (prodR fracR (agreeR bor_stateO)). Definition inhUR := gset_disjUR slice_name. Class lftGS Σ := LftG { - lft_box :> boxG Σ; - alft_inG :> inG Σ (authR alftUR); + lft_box :: boxG Σ; + alft_inG :: inG Σ (authR alftUR); alft_name : gname; - ilft_inG :> inG Σ (authR ilftUR); + ilft_inG :: inG Σ (authR ilftUR); ilft_name : gname; - lft_bor_inG :> inG Σ (authR borUR); - lft_cnt_inG :> inG Σ (authR natUR); - lft_inh_inG :> inG Σ (authR inhUR); + lft_bor_inG :: inG Σ (authR borUR); + lft_cnt_inG :: inG Σ (authR natUR); + lft_inh_inG :: inG Σ (authR inhUR); }. Definition lftGS' := lftGS. Class lftGpreS Σ := LftGPreS { - lft_preG_box :> boxG Σ; - alft_preG_inG :> inG Σ (authR alftUR); - ilft_preG_inG :> inG Σ (authR ilftUR); - lft_preG_bor_inG :> inG Σ (authR borUR); - lft_preG_cnt_inG :> inG Σ (authR natUR); - lft_preG_inh_inG :> inG Σ (authR inhUR); + lft_preG_box :: boxG Σ; + alft_preG_inG :: inG Σ (authR alftUR); + ilft_preG_inG :: inG Σ (authR ilftUR); + lft_preG_bor_inG :: inG Σ (authR borUR); + lft_preG_cnt_inG :: inG Σ (authR natUR); + lft_preG_inh_inG :: inG Σ (authR inhUR); }. Definition lftGpreS' := lftGpreS. @@ -132,7 +132,7 @@ Section defs. Definition lft_vs_inv_go (κ : lft) (lft_inv_alive : ∀ κ', κ' ⊂ κ → iProp Σ) (I : gmap lft lft_names) : iProp Σ := (own_ilft_auth I ∗ - [∗ set] κ' ∈ dom _ I, ∀ Hκ : κ' ⊂ κ, lft_inv_alive κ' Hκ)%I. + [∗ set] κ' ∈ dom I, ∀ Hκ : κ' ⊂ κ, lft_inv_alive κ' Hκ)%I. Definition lft_vs_go (κ : lft) (lft_inv_alive : ∀ κ', κ' ⊂ κ → iProp Σ) (Pb Pi : iProp Σ) : iProp Σ := @@ -176,7 +176,7 @@ Section defs. (∃ (A : gmap atomic_lft bool) (I : gmap lft lft_names), own_alft_auth A ∗ own_ilft_auth I ∗ - [∗ set] κ ∈ dom _ I, lft_inv A κ)%I. + [∗ set] κ ∈ dom I, lft_inv A κ)%I. Definition lft_ctx : iProp Σ := inv mgmtN lfts_inv. @@ -260,7 +260,7 @@ Qed. Lemma lft_vs_inv_unfold κ (I : gmap lft lft_names) : lft_vs_inv κ I ⊣⊢ own_ilft_auth I ∗ - [∗ set] κ' ∈ dom _ I, ⌜κ' ⊂ κ⌝ → lft_inv_alive κ'. + [∗ set] κ' ∈ dom I, ⌜κ' ⊂ κ⌝ → lft_inv_alive κ'. Proof. rewrite /lft_vs_inv /lft_vs_inv_go. by setoid_rewrite pure_impl_forall. Qed. diff --git a/theories/lifetime/model/faking.v b/theories/lifetime/model/faking.v index 94b0f796..6c190897 100644 --- a/theories/lifetime/model/faking.v +++ b/theories/lifetime/model/faking.v @@ -9,9 +9,9 @@ Context `{!invGS Σ, !lftGS Σ}. Implicit Types κ : lft. Lemma ilft_create A I κ : - own_alft_auth A -∗ own_ilft_auth I -∗ ▷ ([∗ set] κ ∈ dom _ I, lft_inv A κ) + own_alft_auth A -∗ own_ilft_auth I -∗ ▷ ([∗ set] κ ∈ dom I, lft_inv A κ) ==∗ ∃ A' I', ⌜is_Some (I' !! κ)⌝ ∗ - own_alft_auth A' ∗ own_ilft_auth I' ∗ ▷ ([∗ set] κ ∈ dom _ I', lft_inv A' κ). + own_alft_auth A' ∗ own_ilft_auth I' ∗ ▷ ([∗ set] κ ∈ dom I', lft_inv A' κ). Proof. iIntros "HA HI Hinv". destruct (decide (is_Some (I !! κ))) as [?|HIκ%eq_None_not_Some]. diff --git a/theories/lifetime/model/primitive.v b/theories/lifetime/model/primitive.v index 9909792b..5dcb652f 100644 --- a/theories/lifetime/model/primitive.v +++ b/theories/lifetime/model/primitive.v @@ -77,7 +77,7 @@ Qed. Lemma own_bor_valid κ x : own_bor κ x -∗ ✓ x. Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed. Lemma own_bor_valid_2 κ x y : own_bor κ x -∗ own_bor κ y -∗ ✓ (x ⋅ y). -Proof. apply wand_intro_r. rewrite -own_bor_op. apply own_bor_valid. Qed. +Proof. apply entails_wand, wand_intro_r. rewrite -own_bor_op. iApply own_bor_valid. Qed. Lemma own_bor_update κ x y : x ~~> y → own_bor κ x ==∗ own_bor κ y. Proof. iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update. @@ -85,7 +85,7 @@ Qed. Lemma own_bor_update_2 κ x1 x2 y : x1 ⋅ x2 ~~> y → own_bor κ x1 ⊢ own_bor κ x2 ==∗ own_bor κ y. Proof. - intros. apply wand_intro_r. rewrite -own_bor_op. by apply own_bor_update. + intros. apply wand_intro_r. rewrite -own_bor_op. by iApply own_bor_update. Qed. Lemma own_cnt_auth I κ x : own_ilft_auth I -∗ own_cnt κ x -∗ ⌜is_Some (I !! κ)⌝. @@ -111,7 +111,7 @@ Qed. Lemma own_cnt_valid κ x : own_cnt κ x -∗ ✓ x. Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed. Lemma own_cnt_valid_2 κ x y : own_cnt κ x -∗ own_cnt κ y -∗ ✓ (x ⋅ y). -Proof. apply wand_intro_r. rewrite -own_cnt_op. apply own_cnt_valid. Qed. +Proof. apply entails_wand, wand_intro_r. rewrite -own_cnt_op. iApply own_cnt_valid. Qed. Lemma own_cnt_update κ x y : x ~~> y → own_cnt κ x ==∗ own_cnt κ y. Proof. iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update. @@ -119,7 +119,7 @@ Qed. Lemma own_cnt_update_2 κ x1 x2 y : x1 ⋅ x2 ~~> y → own_cnt κ x1 -∗ own_cnt κ x2 ==∗ own_cnt κ y. Proof. - intros. apply wand_intro_r. rewrite -own_cnt_op. by apply own_cnt_update. + intros. apply entails_wand, wand_intro_r. rewrite -own_cnt_op. by iApply own_cnt_update. Qed. Lemma own_inh_auth I κ x : own_ilft_auth I -∗ own_inh κ x -∗ ⌜is_Some (I !! κ)⌝. @@ -145,7 +145,7 @@ Qed. Lemma own_inh_valid κ x : own_inh κ x -∗ ✓ x. Proof. iDestruct 1 as (γs) "[#? Hx]". by iApply own_valid. Qed. Lemma own_inh_valid_2 κ x y : own_inh κ x -∗ own_inh κ y -∗ ✓ (x ⋅ y). -Proof. apply wand_intro_r. rewrite -own_inh_op. apply own_inh_valid. Qed. +Proof. apply entails_wand, wand_intro_r. rewrite -own_inh_op. iApply own_inh_valid. Qed. Lemma own_inh_update κ x y : x ~~> y → own_inh κ x ==∗ own_inh κ y. Proof. iDestruct 1 as (γs) "[#Hκ Hx]"; iExists γs. iFrame "Hκ". by iApply own_update. @@ -153,20 +153,20 @@ Qed. Lemma own_inh_update_2 κ x1 x2 y : x1 ⋅ x2 ~~> y → own_inh κ x1 ⊢ own_inh κ x2 ==∗ own_inh κ y. Proof. - intros. apply wand_intro_r. rewrite -own_inh_op. by apply own_inh_update. + intros. apply wand_intro_r. rewrite -own_inh_op. by iApply own_inh_update. Qed. (** Alive and dead *) Global Instance lft_alive_in_dec A κ : Decision (lft_alive_in A κ). Proof. refine (cast_if (decide (set_Forall (λ Λ, A !! Λ = Some true) - (dom (gset atomic_lft) κ)))); + (dom κ)))); rewrite /lft_alive_in; by setoid_rewrite <-gmultiset_elem_of_dom. Qed. Global Instance lft_dead_in_dec A κ : Decision (lft_dead_in A κ). Proof. refine (cast_if (decide (set_Exists (λ Λ, A !! Λ = Some false) - (dom (gset atomic_lft) κ)))); + (dom κ)))); rewrite /lft_dead_in; by setoid_rewrite <-gmultiset_elem_of_dom. Qed. @@ -174,9 +174,9 @@ Lemma lft_alive_or_dead_in A κ : (∃ Λ, Λ ∈ κ ∧ A !! Λ = None) ∨ (lft_alive_in A κ ∨ lft_dead_in A κ). Proof. rewrite /lft_alive_in /lft_dead_in. - destruct (decide (set_Exists (λ Λ, A !! Λ = None) (dom (gset _) κ))) + destruct (decide (set_Exists (λ Λ, A !! Λ = None) (dom κ))) as [(Λ & ?%gmultiset_elem_of_dom & HAΛ)|HA%(not_set_Exists_Forall _)]; first eauto. - destruct (decide (set_Exists (λ Λ, A !! Λ = Some false) (dom (gset _) κ))) + destruct (decide (set_Exists (λ Λ, A !! Λ = Some false) (dom κ))) as [(Λ & HΛ%gmultiset_elem_of_dom & ?)|HA'%(not_set_Exists_Forall _)]; first eauto. right; left. intros Λ HΛ%gmultiset_elem_of_dom. move: (HA _ HΛ) (HA' _ HΛ)=> /=. case: (A !! Λ)=>[[]|]; naive_solver. @@ -300,18 +300,18 @@ Proof. intros p q. rewrite /idx_bor_own -own_bor_op /own_bor. f_equiv=>?. rewrite -auth_frag_op singleton_op -pair_op agree_idemp. done. Qed. -Global Instance frame_lft_tok p κ q1 q2 RES : - FrameFractionalHyps p q1.[κ] (λ q, q.[κ])%I RES q1 q2 → - Frame p q1.[κ] q2.[κ] RES | 5. -Proof. apply: frame_fractional. Qed. +Global Instance frame_lft_tok p κ q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p q1.[κ] q2.[κ] q.[κ] | 5. +Proof. apply: (frame_fractional (λ q, q.[κ])%I). Qed. Global Instance idx_bor_own_as_fractional i q : AsFractional (idx_bor_own q i) (λ q, idx_bor_own q i)%I q. Proof. split. done. apply _. Qed. -Global Instance frame_idx_bor_own p i q1 q2 RES : - FrameFractionalHyps p (idx_bor_own q1 i) (λ q, idx_bor_own q i)%I RES q1 q2 → - Frame p (idx_bor_own q1 i) (idx_bor_own q2 i) RES | 5. -Proof. apply: frame_fractional. Qed. +Global Instance frame_idx_bor_own p i q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p (idx_bor_own q1 i) (idx_bor_own q2 i) (idx_bor_own q i) | 5. +Proof. apply: (frame_fractional (λ q, idx_bor_own q i)%I). Qed. (** Lifetime inclusion *) Lemma lft_incl_acc E κ κ' q : @@ -333,7 +333,7 @@ Qed. Lemma lft_incl_intro κ κ' : □ ((∀ q, q.[κ] ={↑lftN}=∗ ∃ q', q'.[κ'] ∗ (q'.[κ'] ={↑lftN}=∗ q.[κ])) ∗ ([†κ'] ={↑lftN}=∗ [†κ])) -∗ κ ⊑ κ'. -Proof. reflexivity. Qed. +Proof. iIntros "?". done. Qed. Lemma lft_intersect_incl_l κ κ' : ⊢ κ ⊓ κ' ⊑ κ. Proof. @@ -364,7 +364,7 @@ Lemma lft_intersect_acc κ κ' q q' : q.[κ] -∗ q'.[κ'] -∗ ∃ q'', q''.[κ ⊓ κ'] ∗ (q''.[κ ⊓ κ'] -∗ q.[κ] ∗ q'.[κ']). Proof. iIntros "Hκ Hκ'". - destruct (Qp_lower_bound q q') as (qq & q0 & q'0 & -> & ->). + destruct (Qp.lower_bound q q') as (qq & q0 & q'0 & -> & ->). iExists qq. rewrite -lft_tok_sep. iDestruct "Hκ" as "[$ Hκ]". iDestruct "Hκ'" as "[$ Hκ']". iIntros "[Hκ+ Hκ'+]". iSplitL "Hκ Hκ+"; iApply fractional_split; iFrame. @@ -414,8 +414,7 @@ Proof. rewrite /bor /raw_bor /idx_bor /bor_idx. iSplit. - iDestruct 1 as (κ') "[? Hraw]". iDestruct "Hraw" as (s) "[??]". iExists (κ', s). by iFrame. - - iDestruct 1 as ([κ' s]) "[[??]?]". - iExists κ'. iFrame. iExists s. by iFrame. + - iDestruct 1 as ([κ' s]) "[[??]?]". iFrame. Qed. Lemma bor_iff κ P P' : ▷ □ (P ↔ P') -∗ &{κ}P -∗ &{κ}P'. diff --git a/theories/lifetime/model/reborrow.v b/theories/lifetime/model/reborrow.v index d8f6d63e..ce623d9e 100644 --- a/theories/lifetime/model/reborrow.v +++ b/theories/lifetime/model/reborrow.v @@ -89,7 +89,7 @@ Proof. iMod ("Hvs" $! I with "[HI Hinv Hvs' Hinh HB● Hbox HB] [$HPb $Hi] Hκ†") as "($ & $ & Hcnt')". { rewrite lft_vs_inv_unfold. iFrame "HI". - iApply (big_sepS_delete _ (dom (gset lft) I) with "[- $Hinv]"); first done. + iApply (big_sepS_delete _ (dom I) with "[- $Hinv]"); first done. iIntros (_). rewrite lft_inv_alive_unfold. iExists Pb', Pi'. iFrame "Hvs' Hinh". rewrite /lft_bor_alive. iExists (<[i:=Bor_in]>B'). rewrite /to_borUR !fmap_insert. iFrame. @@ -111,7 +111,7 @@ Proof. { rewrite /lft_inv_dead. iDestruct "Hinvκ" as (Pi) "[Hbordead H]". iMod (raw_bor_fake with "Hbordead") as "[Hbordead $]"; first solve_ndisj. iApply "Hclose". iExists _, _. iFrame. - rewrite (big_opS_delete _ (dom _ _) κ') ?elem_of_dom // /lft_inv /lft_inv_dead. + rewrite (big_opS_delete _ (dom _) κ') ?elem_of_dom // /lft_inv /lft_inv_dead. auto 10 with iFrame. } rewrite {1}/raw_bor. iDestruct "Hraw" as (i') "[Hbor H]". iDestruct "H" as (P') "[#HP' #Hs']". @@ -140,7 +140,7 @@ Proof. iRewrite "HeqPb'". iFrame. by iApply "HP'". } iDestruct "HH" as "([HI Hinvκ] & $ & Halive & Hvs)". iApply "Hclose". iExists _, _. iFrame. - rewrite (big_opS_delete _ (dom _ _) κ') ?elem_of_dom //. + rewrite (big_opS_delete _ (dom _) κ') ?elem_of_dom //. iDestruct ("Hclose'" with "Hinvκ") as "$". rewrite /lft_inv lft_inv_alive_unfold. auto 10 with iFrame. Qed. diff --git a/theories/prophecy/prophecy.v b/theories/prophecy/prophecy.v index 6c4ac900..c8d18946 100644 --- a/theories/prophecy/prophecy.v +++ b/theories/prophecy/prophecy.v @@ -187,8 +187,8 @@ Local Definition add_line ξ it (S: proph_smryUR) : proph_smryUR := .<[ξ.(pv_ty) := <[ξ.(pv_id) := it]> (S ξ.(pv_ty))]> S. Definition prophΣ: gFunctors := #[GFunctor prophUR]. -Class prophPreG Σ := ProphPreG { proph_preG_inG :> inG Σ prophUR }. -Class prophG Σ := ProphG { proph_inG :> prophPreG Σ; proph_name: gname }. +Class prophPreG Σ := ProphPreG { proph_preG_inG :: inG Σ prophUR }. +Class prophG Σ := ProphG { proph_inG :: prophPreG Σ; proph_name: gname }. Global Instance subG_prophPreG Σ : subG prophΣ Σ → prophPreG Σ. Proof. solve_inG. Qed. @@ -245,17 +245,17 @@ Proof. Qed. Global Instance proph_tok_as_fractional q ξ : AsFractional q:[ξ] (λ q, q:[ξ]%I) q. Proof. split; by [|apply _]. Qed. -Global Instance frame_proph_tok p ξ q1 q2 RES : - FrameFractionalHyps p q1:[ξ] (λ q, q:[ξ])%I RES q1 q2 → - Frame p q1:[ξ] q2:[ξ] RES | 5. +Global Instance frame_proph_tok p ξ q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p q1:[ξ] q2:[ξ] q:[ξ] | 5. Proof. apply: frame_fractional. Qed. Global Instance proph_toks_as_fractional q ξl : AsFractional q:+[ξl] (λ q, q:+[ξl]%I) q. Proof. split; by [|apply _]. Qed. -Global Instance frame_proph_toks p ξl q1 q2 RES : - FrameFractionalHyps p q1:+[ξl] (λ q, q:+[ξl])%I RES q1 q2 → - Frame p q1:+[ξl] q2:+[ξl] RES | 5. -Proof. apply: frame_fractional. Qed. +Global Instance frame_proph_toks p ξl q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p q1:+[ξl] q2:+[ξl] q:+[ξl] | 5. +Proof. intro. apply: (frame_fractional (λ q, q:+[ξl]%I)). Qed. Global Instance proph_obs_persistent φπ : Persistent .⟨φπ⟩ := _. Global Instance proph_obs_timeless φπ : Timeless .⟨φπ⟩ := _. @@ -280,7 +280,7 @@ Lemma proph_tok_combine ξl ζl q q' : q:+[ξl] -∗ q':+[ζl] -∗ ∃q'', q'':+[ξl ++ ζl] ∗ (q'':+[ξl ++ ζl] -∗ q:+[ξl] ∗ q':+[ζl]). Proof. - case (Qp_lower_bound q q')=> [q''[?[?[->->]]]]. iIntros "[ξl ξl'][ζl ζl']". + case (Qp.lower_bound q q')=> [q''[?[?[->->]]]]. iIntros "[ξl ξl'][ζl ζl']". iExists q''. iFrame "ξl ζl". iIntros "[ξl ζl]". iSplitL "ξl ξl'"; iApply fractional_split; iFrame. Qed. @@ -303,7 +303,7 @@ Lemma proph_intro 𝔄i (I: gset positive) E : ↑prophN ⊆ E → proph_ctx ={E}=∗ ∃i, ⌜i ∉ I⌝ ∗ 1:[PrVar 𝔄i i]. Proof. iIntros (?) "?". iInv prophN as (S) ">[(%L &%& %Sim) ●S]". - case (exist_fresh (I ∪ dom _ (S 𝔄i))) + case (exist_fresh (I ∪ dom (S 𝔄i))) => [i /not_elem_of_union [? /not_elem_of_dom EqNone]]. set ξ := PrVar 𝔄i i. set S' := add_line ξ (fitem 1) S. iMod (own_update _ _ (● S' ⋅ ◯ line ξ (fitem 1)) with "●S") as "[●S' ?]". @@ -311,7 +311,7 @@ Proof. discrete_fun_insert_local_update, alloc_singleton_local_update. } iModIntro. iSplitL "●S'"; last first. { by iModIntro; iExists i; iFrame. } iModIntro. iExists S'. iFrame "●S'". iPureIntro. exists L. - split; [done|]. case=> [𝔅i j]?. rewrite /S' /add_line /discrete_fun_insert -Sim. + split; [done|]. case=> [𝔅i j]?. rewrite /S' /add_line /discrete_fun_insert -Sim /=. case (decide (𝔄i = 𝔅i))=> [?|?]; [|done]. subst=>/=. case (decide (i = j))=> [<-|?]; [|by rewrite lookup_insert_ne]. rewrite lookup_insert EqNone. split=> Eqv; [apply (inj Some) in Eqv|]; inversion Eqv. @@ -372,7 +372,7 @@ Proof. - move=> [Eq'|/InLNe ?]; [|done]. by dependent destruction Eq'. } have Eqv : S' 𝔅i !! j ≡ S 𝔅i !! j. { rewrite /S' /add_line /discrete_fun_insert. - case (decide (𝔄i = 𝔅i))=> [?|?]; [|done]. simpl_eq. + case (decide (pv_ty ξ = 𝔅i))=> [?|?]; [|done]. simpl_eq. case (decide (i = j))=> [?|?]; [|by rewrite lookup_insert_ne]. by subst. } rewrite Eqv Sim. split; [by right|]. case; [|done]=> Eq. by dependent destruction Eq. Qed. @@ -384,10 +384,10 @@ Implicit Type φπ ψπ: proph Prop. Lemma proph_obs_true φπ : (∀π, φπ π) → ⊢ ⟨π, φπ π⟩. Proof. move=> ?. iExists []. by iSplit. Qed. -Lemma proph_obs_impl φπ ψπ : (∀π, φπ π → ψπ π) → .⟨φπ⟩ -∗ .⟨ψπ⟩. +Lemma proph_obs_impl φπ ψπ : (∀π, φπ π → ψπ π) → .⟨φπ⟩ ⊢ .⟨ψπ⟩. Proof. move=> Imp. do 2 f_equiv. move=> ?. by apply Imp. Qed. -Lemma proph_obs_eq φπ ψπ : (∀π, φπ π = ψπ π) → .⟨φπ⟩ -∗ .⟨ψπ⟩. +Lemma proph_obs_eq φπ ψπ : (∀π, φπ π = ψπ π) → .⟨φπ⟩ ⊢ .⟨ψπ⟩. Proof. move=> Eq. apply proph_obs_impl=> ?. by rewrite Eq. Qed. Lemma proph_obs_and φπ ψπ : .⟨φπ⟩ -∗ .⟨ψπ⟩ -∗ ⟨π, φπ π ∧ ψπ π⟩. @@ -398,6 +398,8 @@ Qed. Global Instance proph_obs_from_sep φπ ψπ : FromSep ⟨π, φπ π ∧ ψπ π⟩ .⟨φπ⟩ .⟨ψπ⟩. Proof. rewrite /FromSep. iIntros "#[??]". by iApply proph_obs_and. Qed. +Global Instance proph_obs_combine_sep_as φπ ψπ : CombineSepAs .⟨φπ⟩ .⟨ψπ⟩ ⟨π, φπ π ∧ ψπ π⟩. +Proof. rewrite /CombineSepAs. iIntros "#[??]". by iApply proph_obs_and. Qed. Lemma proph_obs_sat E φπ : ↑prophN ⊆ E → proph_ctx -∗ .⟨φπ⟩ ={E}=∗ ⌜∃π₀, φπ π₀⌝. diff --git a/theories/prophecy/syn_type.v b/theories/prophecy/syn_type.v index bad469b8..006c4c51 100644 --- a/theories/prophecy/syn_type.v +++ b/theories/prophecy/syn_type.v @@ -53,14 +53,15 @@ Fixpoint syn_type_beq 𝔄 𝔅 : bool := Lemma syn_type_eq_correct 𝔄 𝔅 : syn_type_beq 𝔄 𝔅 ↔ 𝔄 = 𝔅. Proof. - move: 𝔄 𝔅. fix FIX 1. - have FIXl: ∀𝔄l 𝔅l, forall2b syn_type_beq 𝔄l 𝔅l ↔ 𝔄l = 𝔅l. - { elim=> [|?? IH][|??]//. rewrite andb_True FIX IH. - split; by [move=> [->->]|move=> [=]]. } - move=> 𝔄 𝔅. case 𝔄; case 𝔅=>//= *; - rewrite ?andb_True ?FIX ?FIXl ?bool_decide_spec; - try (by split; [move=> ->|move=> [=]]); - by split; [move=> [->->]|move=> [=]]. + move: 𝔄 𝔅. fix FIX 1=> 𝔄 𝔅. + destruct 𝔄 as [| | | | | | | | | |𝔄l|𝔄l], 𝔅 as [| | | | | | | | | |𝔅l|𝔅l]=>//=; + try by rewrite ?andb_True !FIX ?bool_decide_spec; intuition congruence. + - assert (forall2b syn_type_beq 𝔄l 𝔅l ↔ 𝔄l = 𝔅l) as ->. + { revert 𝔅l; induction 𝔄l as [|𝔄 𝔄l IH]; move=>[|𝔅 𝔅l] //=. rewrite ?andb_True FIX IH. intuition congruence. } + intuition congruence. + - assert (forall2b syn_type_beq 𝔄l 𝔅l ↔ 𝔄l = 𝔅l) as ->. + { revert 𝔅l; induction 𝔄l as [|𝔄 𝔄l IH]; move=>[|𝔅 𝔅l] //=. rewrite ?andb_True FIX IH. intuition congruence. } + intuition congruence. Qed. Global Instance syn_type_beq_dec: EqDecision syn_type. Proof. diff --git a/theories/typing/array.v b/theories/typing/array.v index 9fbe383c..10ac5e0a 100644 --- a/theories/typing/array.v +++ b/theories/typing/array.v @@ -82,8 +82,8 @@ Section typing. setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "[%] tys Na κ₊") as (q'' ?) "(Na & ↦' & (%&>->& #tys) & Toκ₊)". { apply subseteq_difference_r. { symmetry. apply shr_locsE_disj. } - move: HF. rewrite -plus_assoc shr_locsE_shift. set_solver. } - case (Qp_lower_bound q' q'')=> [q'''[?[?[->->]]]]. iExists q''', (_ ++ _). + move: HF. rewrite -assoc shr_locsE_shift. set_solver. } + case (Qp.lower_bound q' q'')=> [q'''[?[?[->->]]]]. iExists q''', (_ ++ _). rewrite heap_mapsto_vec_app. iDestruct (ty_size_eq with "ty") as ">->". iDestruct "↦" as "[$ ↦r]". iDestruct "↦'" as "[$ ↦r']". iDestruct (na_own_acc with "Na") as "[$ ToNa]". @@ -174,7 +174,7 @@ Section typing. rewrite vzip_with_app !vec_to_list_app concat_app. iExists _, _. iSplit; [done|]. iDestruct "tys" as "[tys tys']". iSplitL "tys"; iExists _; by iFrame. - iApply bi.equiv_iff. - rewrite vec_to_list_app big_sepL_app vec_to_list_length. do 5 f_equiv. + rewrite vec_to_list_app big_sepL_app length_vec_to_list. do 5 f_equiv. by rewrite shift_loc_assoc_nat -Nat.mul_add_distr_r. Qed. diff --git a/theories/typing/array_idx.v b/theories/typing/array_idx.v index 6ed51926..93f0318e 100644 --- a/theories/typing/array_idx.v +++ b/theories/typing/array_idx.v @@ -271,4 +271,4 @@ End array_idx. Global Hint Resolve tctx_extract_split_own_array tctx_extract_idx_shr_array tctx_extract_split_uniq_array | 5 : lrust_typing. -Global Hint Resolve tctx_extract_merge_own_array | 40 : lrust_typing_merge. +Global Hint Resolve tctx_extract_merge_own_array | 60 : lrust_typing_merge. diff --git a/theories/typing/array_util.v b/theories/typing/array_util.v index 093beb94..deab2948 100644 --- a/theories/typing/array_util.v +++ b/theories/typing/array_util.v @@ -43,7 +43,7 @@ Section array_util. Proof. induction aπl as [|??? IH]; inv_vec wll; [by iIntros|]. iIntros (??) "/=[ty tys]". iDestruct (ty_size_eq with "ty") as %?. - iDestruct (IH with "tys") as %?. iPureIntro. rewrite app_length. lia. + iDestruct (IH with "tys") as %?. iPureIntro. rewrite length_app. lia. Qed. Lemma ty_share_big_sepL {𝔄} (ty: type 𝔄) E aπl d κ l tid q : @@ -111,8 +111,9 @@ Section array_util. iDestruct "κ'" as "[κ' κ'₊]". iDestruct "tys" as "[ty tys]". iMod (ty_shr_proph with "LFT In In' ty κ'") as "Upd"; [done|]. setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "tys κ'₊") as "Upd'". - iIntros "!>!>". iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand with "Upd"). - iIntros "[>(%&%&%& ξl & Toκ') >(%&%&%& ζl & Toκ'₊)] !>". + iIntros "!>!>". iMod "Upd". iMod "Upd'". iCombine "Upd Upd'" as "Upd". + iApply (step_fupdN_wand with "Upd"). + iIntros "!> [>(%&%&%& ξl & Toκ') >(%&%&%& ζl & Toκ'₊)] !>". iDestruct (proph_tok_combine with "ξl ζl") as (?) "[ξζl Toξζl]". iExists _, _. iFrame "ξζl". iSplit; [iPureIntro; by apply proph_dep_vec_S|]. iIntros "ξζl". iDestruct ("Toξζl" with "ξζl") as "[ξl ζl]". @@ -165,8 +166,8 @@ Section array_util. { iApply step_fupdN_full_intro. iFrame "L". iPureIntro. by exists [#]. } iDestruct "tys" as "[ty tys]". iDestruct "L" as "[L L₊]". setoid_rewrite <-shift_loc_assoc_nat. iMod (Rl with "LFT E L ty") as "Upd"; [done|]. - iMod ("IH" with "L₊ tys") as "Upd'". iCombine "Upd Upd'" as "Upd". iIntros "!>!>". - iApply (step_fupdN_wand with "Upd"). iIntros "[>(%Eq &$&$) >(%Eq' &$&$)] !%". + iMod ("IH" with "L₊ tys") as "Upd'". iIntros "!>!>". iMod "Upd". iMod "Upd'". iCombine "Upd Upd'" as "Upd". + iApply (step_fupdN_wand with "Upd"). iIntros "!> [>(%Eq &$&$) >(%Eq' &$&$)] !%". move: Eq=> [b Eq]. move: Eq'=> [bl Eq']. exists (b ::: bl). fun_ext=>/= π. by move: (equal_f Eq π) (equal_f Eq' π)=>/= <-<-. Qed. diff --git a/theories/typing/borrow.v b/theories/typing/borrow.v index e2e31b0e..e03aed58 100644 --- a/theories/typing/borrow.v +++ b/theories/typing/borrow.v @@ -71,8 +71,7 @@ Section borrow. iDestruct (uniq_agree with "Hvo Hpc") as "%Hag"; inversion Hag; subst; clear Hag. iMod (uniq_resolve with "PROPH Hvo Hpc Hdt") as "(Hobs & Hpc & Hdt)"; [done|done| ]. iMod ("Hclose''" with "Hdt") as "[Hown Htok]". - iMod ("Hclose'" with "[H↦ Hown Hpc]") as "[Huniq Htok2]". - { iFrame "#∗". iExists _. iFrame. } + iMod ("Hclose'" with "[H↦ Hown Hpc]") as "[Huniq Htok2]"; first by iFrame "#∗". do 2 (iMod (bor_sep with "LFT Huniq") as "[_ Huniq]"; [done|]). iDestruct (ty.(ty_share) with "LFT [$] Huniq Htok") as "Hshr"; [done|]. iModIntro. wp_seq. diff --git a/theories/typing/examples/nonlexical.v b/theories/typing/examples/nonlexical.v deleted file mode 100644 index a543759a..00000000 --- a/theories/typing/examples/nonlexical.v +++ /dev/null @@ -1,132 +0,0 @@ -From iris.proofmode Require Import proofmode. -From lrust.typing Require Import typing lib.option. - -(* Typing "problem case #3" from: - http://smallcultfollowing.com/babysteps/blog/2016/04/27/non-lexical-lifetimes-introduction/ - - Differences with the original implementation: - - We ignore dropping. - We have to add a Copy bound on the key type. - We do not support panic, hence we do not support option::unwrap. We use - unwrap_or as a replacement. -*) - -Section non_lexical. - Context `{!typeG Σ} (HashMap K V : type) `{!Copy K} - (defaultv get_mut insert: val). - - Hypothesis defaultv_type : - typed_val defaultv (fn(∅) → V). - - Hypothesis get_mut_type : - typed_val get_mut (fn(∀ '(α, β), ∅; &uniq{α} HashMap, &shr{β} K) → - option (&uniq{α} V)). - - Hypothesis insert_type : - typed_val insert (fn(∀ α, ∅; &uniq{α} HashMap, K, V) → option V). - - Definition get_default : val := - fn: ["map"; "key"] := - let: "get_mut" := get_mut in - let: "map'" := !"map" in - - Newlft;; - - Newlft;; - letalloc: "map0" <- "map'" in - Share;; - letalloc: "key0" <- "key" in - letcall: "o" := "get_mut" ["map0"; "key0"]%E in - Endlft;; - withcont: "k": - case: !"o" of - [ (* None *) - Endlft;; - - let: "insert" := insert in - Newlft;; - letalloc: "map0" <- "map'" in - letalloc: "key0" <-{K.(ty_size)} !"key" in - let: "defaultv" := defaultv in - letcall: "default0" := "defaultv" [] in - letcall: "old" := "insert" ["map0"; "key0"; "default0"]%E in - Endlft;; - delete [ #(option V).(ty_size); "old"];; (* We should drop here. *) - - Newlft;; - letalloc: "map0" <- "map'" in - Share;; - letalloc: "key0" <- "key" in - letcall: "r'" := "get_mut" ["map0"; "key0"]%E in - Endlft;; - let: "unwrap" := option_unwrap (&uniq{static}V) in - letcall: "r" := "unwrap" ["r'"]%E in - "k" ["r"]; - - (* Some *) - letalloc: "r" <-{1} !("o" +ₗ #1) in - "k" ["r"] - ] - cont: "k" ["r"] := - delete [ #(option (&uniq{static}V)).(ty_size); "o"];; - delete [ #1; "map"];; delete [ #K.(ty_size); "key"];; (* We should drop key here *) - return: ["r"]. - - Lemma get_default_type : - typed_val get_default (fn(∀ m, ∅; &uniq{m} HashMap, K) → &uniq{m} V). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (m ϝ ret p). inv_vec p=>map key. simpl_subst. - iApply type_let; [iApply get_mut_type|solve_typing|]. - iIntros (get_mut'). simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (map'). simpl_subst. - iApply (type_newlft [m]). iIntros (κ). - iApply (type_newlft [ϝ]). iIntros (α). - iApply (type_letalloc_1 (&uniq{κ}_)); [solve_typing..|]. iIntros (map0). simpl_subst. - iApply (type_share key _ α); [solve_typing..|]. - iApply (type_letalloc_1 (&shr{α}K)); [solve_typing..|]. iIntros (key0). simpl_subst. - iApply (type_letcall (κ, α)); [solve_typing..|]. iIntros (o). simpl_subst. - iApply type_endlft. - iApply (type_cont [_] [ϝ ⊑ₗ []] - (λ r, [o ◁ box (Π[uninit 1;uninit 1]); map ◁ box (uninit 1); - key ◁ box K; (r!!!0%fin:val) ◁ box (&uniq{m} V)])). - { iIntros (k). simpl_subst. - iApply type_case_own; - [solve_typing| constructor; [|constructor; [|constructor]]; left]. - - iApply type_endlft. - iApply type_let; [iApply insert_type|solve_typing|]. - iIntros (insert'). simpl_subst. - iApply (type_newlft [ϝ]). iIntros (β). clear map0 key0. - iApply (type_letalloc_1 (&uniq{β}_)); [solve_typing..|]. - iIntros (map0). simpl_subst. - iApply (type_letalloc_n _ (box K)); [solve_typing..|]. iIntros (key0). simpl_subst. - iApply type_let; [iApply defaultv_type|solve_typing|]. - iIntros (defaultv'). simpl_subst. - iApply (type_letcall ()); [solve_typing..|]. iIntros (default0). simpl_subst. - iApply (type_letcall β); [solve_typing..|]. iIntros (old). simpl_subst. - iApply type_endlft. - iApply type_delete; [solve_typing..|]. - iApply (type_newlft [ϝ]). iIntros (γ). clear map0 key0. - iApply (type_letalloc_1 (&uniq{m}_)); [solve_typing..|]. - iIntros (map0). simpl_subst. - iApply (type_share key _ γ); [solve_typing..|]. - iApply (type_letalloc_1 (&shr{γ}K)); [solve_typing..|]. - iIntros (key0). simpl_subst. - iApply (type_letcall (m, γ)); [solve_typing..|]. iIntros (r'). simpl_subst. - iApply type_endlft. - iApply type_let; [iApply (option_unwrap_type (&uniq{m}V))|solve_typing|]. - iIntros (unwrap'). simpl_subst. - iApply (type_letcall ()); [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_jump; solve_typing. - - iApply type_equivalize_lft. - 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. - iApply type_delete; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. -End non_lexical. diff --git a/theories/typing/fixpoint.v b/theories/typing/fixpoint.v index db14e95c..9e7eb51a 100644 --- a/theories/typing/fixpoint.v +++ b/theories/typing/fixpoint.v @@ -48,11 +48,23 @@ Section S. (Tn (2 + i)).(ty_shr) vπ d κ tid l ≡{n}≡ (Tn (2 + n)).(ty_shr) vπ d κ tid l). Proof using HT. move: i. elim: n=> /=[|n IH]=> i ?. - - split; [done|]. apply HT=>//; [apply type_contractive_ty_size| - apply (Tn_ty_lft_const (S i) 1)|apply (Tn_ty_E_const i 0)]. - - case i as [|]; [lia|]. case (IH i) as [??]; [lia|]. - split; (apply HT=>//; [apply type_contractive_ty_size| - apply (Tn_ty_lft_const (2 + i) (2 + n))|apply (Tn_ty_E_const (S i) (S n))]). + - split; [intros; by apply dist_later_0|]. apply HT. + + apply type_contractive_ty_size. + + apply (Tn_ty_lft_const (S i) 1). + + apply (Tn_ty_E_const i 0). + + done. + + intros. apply dist_later_0. + - case i as [|]; [lia|]. case (IH i) as [??]; [lia|]. split. + + intros. apply dist_later_S. apply HT; try done. + * apply type_contractive_ty_size. + * apply (Tn_ty_lft_const (2 + i) (2 + n)). + * apply (Tn_ty_E_const (S i) (S n)). + + intros. apply HT. + * apply type_contractive_ty_size. + * apply (Tn_ty_lft_const (2 + i) (2 + n)). + * apply (Tn_ty_E_const (S i) (S n)). + * destruct n; [done|]. intros. apply IH; lia. + * intros. apply dist_later_S. done. Qed. Program Definition own_shr_chain := {| chain_car n := ((Tn (3 + n)).(ty_own), (Tn (3 + n)).(ty_shr)) : @@ -60,7 +72,7 @@ Section S. (proph 𝔄 -d> nat -d> lft -d> thread_id -d> loc -d> iPropO Σ) |}. Next Obligation. move=> n i Hni. split=>/=. - - move=> >. apply (Tn_cauchy (S _)). lia. + - move=> >. apply (Tn_cauchy (S _)); lia. - move=> >. apply dist_S, Tn_cauchy. lia. Qed. @@ -109,20 +121,20 @@ Section S. Qed. Next Obligation. move=> *. apply @limit_preserving; [|move=> ?; by apply ty_shr_lft_mono]. - apply limit_preserving_entails; [done|]=> ??? Eq. f_equiv; apply Eq. + apply limit_preserving_entails; [done|]=> ??? Eq. do 2 f_equiv; apply Eq. Qed. Next Obligation. move=> *. apply @limit_preserving; [|move=> ?; by apply (ty_share (Tn' _))]. - apply limit_preserving_entails; [done|]=> ??? Eq. do 6 f_equiv; [|f_equiv]; apply Eq. + apply limit_preserving_entails; [done|]=> ??? Eq. do 7 f_equiv; [|f_equiv]; apply Eq. Qed. Next Obligation. move=> *. apply @limit_preserving; [|move=> ?; by apply (ty_own_proph (Tn' _))]. apply limit_preserving_entails; [done|]=> ??? Eq. - do 2 f_equiv; [|do 13 f_equiv]; apply Eq. + do 3 f_equiv; [|do 13 f_equiv]; apply Eq. Qed. Next Obligation. move=> *. apply @limit_preserving; [|move=> ?; by apply (ty_shr_proph (Tn' _))]. - apply limit_preserving_entails; [done|]=> ??? Eq. do 3 f_equiv. apply Eq. + apply limit_preserving_entails; [done|]=> ??? Eq. do 4 f_equiv. apply Eq. Qed. Lemma fix_ty_Tn'_dist n : fix_ty ≡{n}≡ Tn' (3 + n). @@ -152,7 +164,7 @@ Section fix_ty. (T (fix_ty T)).(ty_own) vπ d tid vl. { move=> *. apply equiv_dist=> n. etrans; [apply dist_S, conv_compl|]. rewrite/= (EqOwn n). symmetry. apply HT=>// *; [apply lft_equiv_refl| |]. - - move: n=> [|n]; [done|]. + - destruct n; [by apply dist_later_0|]. apply dist_later_S. case (fix_ty_Tn'_dist T (S n))=> [_ _ _ Eq _]. apply dist_S, Eq. - case (fix_ty_Tn'_dist T n)=> [_ _ _ _ Eq]. apply Eq. } have EqShr': ∀vπ d κ tid l, (fix_ty T).(ty_shr) vπ d κ tid l ≡ @@ -161,7 +173,7 @@ Section fix_ty. rewrite/= (EqShr n). symmetry. apply HT=>// *; [apply lft_equiv_refl| |]. - move: n=> [|[|n]]; [done|done|]. case (fix_ty_Tn'_dist T (S n))=> [_ _ _ Eq _]. apply dist_S, Eq. - - move: n=> [|n]; [done|]. + - destruct n; [by apply dist_later_0|]. apply dist_later_S. case (fix_ty_Tn'_dist T n)=> [_ _ _ _ Eq]. apply Eq. } apply eqtype_id_unfold. iIntros "*_!>_". iSplit; [iPureIntro; by apply HT|]. iSplit; [|iSplit; iIntros "!> *"]. @@ -193,7 +205,7 @@ Section fix_ty. move=> Eq. have Eq': compl (own_shr_chain T) ≡{n}≡ compl (own_shr_chain T'). { have Eq'': Tn T (3 + n) ≡{n}≡ Tn T' (3 + n). - { rewrite /Tn. elim (S (3 + n)); [done|]=> ? IH. by rewrite !Nat_iter_S IH Eq. } + { rewrite /Tn. elim (S (3 + n)); [done|]=> ? IH. by rewrite !Nat.iter_succ IH Eq. } etrans; [apply conv_compl|]. etrans; [|symmetry; apply conv_compl]. split; repeat move=> ? /=; apply Eq''. } split=>/=; try apply Eq; try apply Eq'. by rewrite /Tn /= (Eq base) Eq. @@ -259,7 +271,7 @@ Section fix_ty. - move=> > ?. eapply @limit_preserving. { apply limit_preserving_forall=> ?. apply limit_preserving_entails; [done|]=> ??? Eq. - f_equiv; [|do 11 f_equiv]; apply Eq. } + do 2 f_equiv; [|do 11 f_equiv]; apply Eq. } move=> n. have ->: (Tn T 0).(ty_size) = (Tn T (3 + n)).(ty_size). { rewrite /Tn /=. apply type_contractive_ty_size. } by apply copy_shr_acc. @@ -288,7 +300,7 @@ Section fix_ty. { elim=> [|? H]; apply Loop; [apply base_resolve|apply H]. } rewrite /fix_ty=> > /=. eapply @limit_preserving; [|move=> ?; apply Rslv]. apply limit_preserving_forall=> ?. - apply limit_preserving_entails; [done|]=> ??? Eq. do 4 f_equiv. apply Eq. + apply limit_preserving_entails; [done|]=> ??? Eq. do 5 f_equiv. apply Eq. Qed. Lemma fix_real {𝔅} E L (f: _ → 𝔅) : @@ -300,7 +312,7 @@ Section fix_ty. eapply @limit_preserving; [|move=> ?; apply Rl]; apply limit_preserving_forall=> ?; apply limit_preserving_entails; [done|]=> ??? Eq; - do 3 f_equiv; [apply Eq|]; do 5 f_equiv); [|do 2 f_equiv]; apply Eq. + do 4 f_equiv; [apply Eq|]; do 5 f_equiv); [|do 2 f_equiv]; apply Eq. Qed. End fix_ty. @@ -319,12 +331,12 @@ Section fix_subtyping. subtype E L (fix_ty T) (fix_ty T') f. Proof. move=> Loop qL. - have Incl: llctx_interp L qL -∗ □ (elctx_interp E -∗ + have Incl: llctx_interp L qL ⊢ □ (elctx_interp E -∗ ∀n, type_incl (Tn T n) (Tn T' n) f). { rewrite intuitionistically_into_persistently -wand_forall persistently_forall. apply forall_intro=> n. rewrite -intuitionistically_into_persistently. - move: qL. apply Loop. elim n=> [|??]; [solve_typing|by apply Loop]. } - rewrite Incl /type_incl -!persistent_and_sep /=. do 2 f_equiv. + apply wand_entails, Loop. elim n=> [|??]; [solve_typing|by apply Loop]. } + rewrite Incl /type_incl -!persistent_and_sep /=. apply entails_wand. do 2 f_equiv. (* FIXME : change the definition of limit_preserving so that it applies even if the limti is not computed with compl. *) apply and_intro; [|apply and_intro; [|apply and_intro]]. @@ -343,12 +355,12 @@ Section fix_subtyping. subtype E L (fix_ty T) (fix_ty T') f. Proof. move=> Loop qL. - have Incl: llctx_interp L qL -∗ □ (elctx_interp E -∗ + have Incl: llctx_interp L qL ⊢ □ (elctx_interp E -∗ ∀n, type_incl (Tn T n) (Tn T' n) f). { rewrite intuitionistically_into_persistently -wand_forall persistently_forall. apply forall_intro=> n. rewrite -intuitionistically_into_persistently. - move: qL. apply Loop. elim n=> [|??]; [solve_typing|by apply Loop]. } - rewrite Incl /type_incl -!persistent_and_sep /=. do 2 f_equiv. + apply wand_entails, Loop. elim n=> [|??]; [solve_typing|by apply Loop]. } + rewrite Incl /type_incl -!persistent_and_sep /=. apply entails_wand. do 2 f_equiv. apply and_intro; [|apply and_intro; [|apply and_intro]]. - iIntros "H". iDestruct ("H" $! 0%nat) as "($&_)". - iIntros "H". iDestruct ("H" $! 0%nat) as "(_&$&_)". diff --git a/theories/typing/function.v b/theories/typing/function.v index 85fdee56..8d2ba1d1 100644 --- a/theories/typing/function.v +++ b/theories/typing/function.v @@ -130,7 +130,10 @@ Section typing. - apply (type_lft_morphism_const _ static [])=>//= ?. apply lft_equiv_refl. - move=> *. by apply Eq. - move=>/= n *. apply bi.exist_ne=> ?. apply bi.sep_ne; [done|]. - apply uPred_primitive.later_contractive. destruct n=>/=; [done|by apply Eq]. } + apply uPred_primitive.later_contractive. destruct n; [by apply dist_later_0|]. + apply dist_later_S, Eq; try done. + * intros. destruct n; [by apply dist_later_0|]. by apply dist_later_S. + * intros. eapply dist_later_dist_lt; [|done]. lia. } move=>/= n ty ty' *. apply bi.exist_ne=> ?. apply bi.sep_ne; [done|]. do 5 apply bi.exist_ne=> ?. f_equiv. f_contractive. (do 2 f_equiv)=> x. (do 5 f_equiv)=> wl. rewrite /typed_body. (do 3 f_equiv)=> aπl. do 2 f_equiv. @@ -140,8 +143,9 @@ Section typing. - by apply Ne. - by iApply type_lft_morphism_lft_equiv_proper. - apply type_lft_morphism_elctx_interp_proper=>//. apply _. - - apply dist_dist_later. by apply Ne. - - apply dist_S. by apply Ne. } + - destruct n; [apply dist_later_0|apply dist_later_S]. apply Ne; try done. + intros. eapply dist_later_dist_lt; [|done]; lia. + - by apply Ne. } move: (NeIT x)=> [?[->NeITl]]. do 5 f_equiv; [|do 3 f_equiv; [|f_equiv]]. - apply equiv_dist. rewrite /fp_E /= !elctx_interp_app. do 2 f_equiv; [|f_equiv; [|f_equiv]]. @@ -154,9 +158,9 @@ Section typing. + rewrite !elctx_interp_ty_outlives_E. apply lft_incl_equiv_proper_r. by iApply type_lft_morphism_lft_equiv_proper. - rewrite !cctx_interp_singleton /cctx_elt_interp. do 3 f_equiv. case=>/= ??. - do 4 f_equiv. rewrite /tctx_elt_interp. do 6 f_equiv. by apply EqBox. - - clear -NeITl EqBox. induction NeITl, wl, aπl; [done|]=>/=. - f_equiv; [|done]. rewrite /tctx_elt_interp. do 6 f_equiv. by apply EqBox. + do 4 f_equiv. rewrite /tctx_elt_interp. do 6 f_equiv. eapply dist_le; [by apply EqBox|]. lia. + - clear -NeITl EqBox Hlt. induction NeITl, wl, aπl; [done|]=>/=. + f_equiv; [|done]. rewrite /tctx_elt_interp. do 6 f_equiv. eapply dist_le; [by apply EqBox|]. lia. Qed. Global Instance fn_send {A 𝔄l 𝔅} (fp: A → fn_params 𝔄l 𝔅) : Send (fn fp). @@ -206,7 +210,7 @@ Section typing. rewrite subst_plv_renew. set wl := plistc_renew _ wl'. iDestruct ("Big" with "[$E $Efp']") as "(Efp & InIl & InO)". iApply ("fn" $! _ _ _ _ _ - (plist_map_with (λ _ _, (∘)) fl aπl') (λ π b, postπ' π (g b)) + (plist_map_with (λ _ _ f g x, f (g x)) fl aπl') (λ π b, postπ' π (g b)) with "LFT TIME PROPH UNIQ Efp Na L [C] [T] [Obs]"). - rewrite !cctx_interp_singleton. iRevert "InO C". iClear "#". iIntros "#(_&_& InO &_) C". iIntros (?[??]) "Na L /=[(%&%&%& ⧖ & oty) _] Obs". @@ -270,7 +274,7 @@ Section typing. iMod (lft_create with "LFT") as (ϝ) "[ϝ #To†ϝ]"; [done|]. iMod (bor_create _ ϝ with "LFT κl") as "[Borκl Toκl]"; [done|]. iDestruct (frac_bor_lft_incl with "LFT [>Borκl]") as "#?". - { iApply (bor_fracture with "LFT"); [done|]. by rewrite Qp_mul_1_r. } + { iApply (bor_fracture with "LFT"); [done|]. by rewrite Qp.mul_1_r. } iDestruct (ToEfp ϝ with "L [$E]") as "#Efp". { clear ToEfp. iInduction κl as [|κ κl] "IH"; [done|]=>/=. iSplit. { iApply lft_incl_trans; [done|]. iApply lft_intersect_incl_l. } diff --git a/theories/typing/lft_contexts.v b/theories/typing/lft_contexts.v index 685830cd..d29bda08 100644 --- a/theories/typing/lft_contexts.v +++ b/theories/typing/lft_contexts.v @@ -61,16 +61,16 @@ Section lft_contexts. Definition llctx_interp (L : llctx) (q : Qp) : iProp Σ := ([∗ list] x ∈ L, llctx_elt_interp x q)%I. - Global Arguments llctx_interp _ _%Qp. + Global Arguments llctx_interp _ _%_Qp. Global Instance llctx_interp_fractional L : Fractional (llctx_interp 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 frame_llctx_interp p L q1 q2 RES : - FrameFractionalHyps p (llctx_interp L q1) (llctx_interp L) RES q1 q2 → - Frame p (llctx_interp L q1) (llctx_interp L q2) RES | 5. + Global Instance frame_llctx_interp p L q1 q2 q : + FrameFractionalQp q1 q2 q → + Frame p (llctx_interp L q1) (llctx_interp L q2) (llctx_interp L q) | 5. Proof. apply: frame_fractional. Qed. Global Instance llctx_interp_permut : @@ -199,7 +199,7 @@ Section lft_contexts. Proof. iIntros (? Hal) "#HE [HL1 HL2]". iMod (Hal with "HE HL1") as (q') "[Htok Hclose]"; first done. - destruct (Qp_lower_bound (q/2) q') as (qq & q0 & q'0 & Hq & ->). rewrite Hq. + destruct (Qp.lower_bound (q/2) q') as (qq & q0 & q'0 & Hq & ->). rewrite Hq. iExists qq. iDestruct "HL2" as "[$ HL]". iDestruct "Htok" as "[$ Htok]". iIntros "!> Htok' HL'". iCombine "HL'" "HL" as "HL". rewrite -Hq. iFrame. iApply "Hclose". iFrame. @@ -217,7 +217,7 @@ Section lft_contexts. inversion_clear Hκs. iIntros "HL". iMod (lctx_lft_alive_tok κ 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 & -> & ->). + destruct (Qp.lower_bound q' q'') as (qq & q0 & q'0 & -> & ->). iExists qq. iDestruct "HL" as "[$ HL2]". iDestruct "Hκ" as "[Hκ1 Hκ2]". iDestruct "Hκs" as "[Hκs1 Hκs2]". iModIntro. simpl. rewrite -lft_tok_sep. iSplitL "Hκ1 Hκs1". { by iFrame. } @@ -256,12 +256,12 @@ Section lft_contexts. - iDestruct "HL1" as "[HL1 HL2]". iMod (Hκ with "HE HL1") as (q') "[Htok' Hclose]". done. iMod ("IH" with "HL2") as (q'') "[Htok'' Hclose']". - destruct (Qp_lower_bound q' q'') as (q0 & q'2 & q''2 & -> & ->). + 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 & -> & ->). + 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]". iDestruct "Htok'" as "[$ Htok']". iIntros "!>[Hfold Hκ0]". iMod ("Hclose'" with "[$Hfold $Htok']") as "$". diff --git a/theories/typing/lib/arc.v b/theories/typing/lib/arc.v deleted file mode 100644 index 0d0bfc40..00000000 --- a/theories/typing/lib/arc.v +++ /dev/null @@ -1,1202 +0,0 @@ -From iris.proofmode Require Import tactics. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lang.lib Require Import memcpy arc. -From lrust.lifetime Require Import at_borrow. -From lrust.typing Require Export type. -From lrust.typing Require Import typing option. -Set Default Proof Using "Type". - -Definition arcN := lrustN .@ "arc". -Definition arc_invN := arcN .@ "inv". -Definition arc_shrN := arcN .@ "shr". -Definition arc_endN := arcN .@ "end". - -Section arc. - Context `{!typeG Σ, !arcG Σ}. - - (* Preliminary definitions. *) - Let P1 ν q := (q.[ν])%I. - Global Instance P1_fractional ν : Fractional (P1 ν). - Proof. unfold P1. apply _. Qed. - Let P2 l n := († l…(2 + n) ∗ (l +ₗ 2) ↦∗: λ vl, ⌜length vl = n⌝)%I. - Definition arc_persist tid ν (γ : gname) (l : loc) (ty : type) : iProp Σ := - (is_arc (P1 ν) (P2 l ty.(ty_size)) arc_invN γ l ∗ - (* We use this disjunction, and not simply [ty_shr] here, *) - (* because [weak_new] cannot prove ty_shr, even for a dead *) - (* lifetime. *) - (ty.(ty_shr) (ν ⊓ ty_lft ty) tid (l +ₗ 2) ∨ [†ν]) ∗ - □ (1.[ν] ={↑lftN ∪ ↑lft_userN ∪ ↑arc_endN}[↑lft_userN]▷=∗ - [†ν] ∗ ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid ∗ † l…(2 + ty.(ty_size))))%I. - - Lemma arc_persist_type_incl tid ν γ l ty1 ty2: - type_incl ty1 ty2 -∗ arc_persist tid ν γ l ty1 -∗ arc_persist tid ν γ l ty2. - Proof. - iIntros "#(Heqsz & Hl & Hinclo & Hincls) #(?& Hs & Hvs)". - iDestruct "Heqsz" as %->. iFrame "#". iSplit. - - iDestruct "Hs" as "[?|?]"; last auto. iLeft. iApply "Hincls". - iApply (ty_shr_mono with "[] [//]"). - iApply lft_intersect_mono; [iApply lft_incl_refl|done]. - - iIntros "!> Hν". iMod ("Hvs" with "Hν") as "H". iModIntro. iNext. - iMod "H" as "($ & H & $)". iDestruct "H" as (vl) "[??]". iExists _. - iFrame. by iApply "Hinclo". - Qed. - - Lemma arc_persist_send tid tid' ν γ l ty `{!Send ty,!Sync ty} : - arc_persist tid ν γ l ty -∗ arc_persist tid' ν γ l ty. - Proof. - iIntros "#($ & ? & Hvs)". rewrite sync_change_tid. iFrame "#". - iIntros "!> Hν". iMod ("Hvs" with "Hν") as "H". iModIntro. iNext. - iMod "H" as "($ & H & $)". iDestruct "H" as (vl) "?". iExists _. - by rewrite send_change_tid. - Qed. - - (* Model of arc *) - (* The ty_own part of the arc type cointains either the - shared protocol or the full ownership of the - content. The reason is that get_mut does not have the - masks to rebuild the invariants. *) - Definition full_arc_own l ty tid : iProp Σ:= - (l ↦ #1 ∗ (l +ₗ 1) ↦ #1 ∗ † l…(2 + ty.(ty_size)) ∗ - ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid)%I. - Definition shared_arc_own l ty tid : iProp Σ:= - (∃ γ ν q, arc_persist tid ν γ l ty ∗ arc_tok γ q ∗ q.[ν])%I. - Lemma arc_own_share E l ty tid q : - ↑lftN ⊆ E → - lft_ctx -∗ (q).[ty_lft ty] -∗ full_arc_own l ty tid - ={E}=∗ (q).[ty_lft ty] ∗ shared_arc_own l ty tid. - Proof. - iIntros (?) "#LFT Htok (Hl1 & Hl2 & H† & Hown)". - iMod (lft_create with "LFT") as (ν) "[Hν #Hν†]"=>//. - (* TODO: We should consider changing the statement of - bor_create to dis-entangle the two masks such that this is no - longer necessary. *) - iApply fupd_trans. iApply (fupd_mask_mono (↑lftN))=>//. - iMod (bor_create _ ν with "LFT Hown") as "[HP HPend]"=>//. iModIntro. - iDestruct (lft_intersect_acc with "Hν Htok") as (q') "[Htok Hclose]". - iMod (ty_share with "LFT [] [HP] Htok") as "[#? Htok]"; first solve_ndisj. - { iApply lft_intersect_incl_r. } - { iApply (bor_shorten with "[] HP"). iApply lft_intersect_incl_l. } - iDestruct ("Hclose" with "Htok") as "[Hν $]". - iMod (create_arc (P1 ν) (P2 l ty.(ty_size)) arc_invN with "Hl1 Hl2 Hν") - as (γ q'') "(#? & ? & ?)". - iExists _, _, _. iFrame "∗#". iCombine "H†" "HPend" as "H". - iMod (inv_alloc arc_endN _ ([†ν] ∨ _)%I with "[H]") as "#INV"; - first by iRight; iApply "H". iIntros "!> !> Hν". - iMod (inv_acc with "INV") as "[[H†|[$ Hvs]] Hclose]"; - [set_solver|iDestruct (lft_tok_dead with "Hν H†") as ">[]"|]. - rewrite difference_union_distr_l_L difference_diag_L right_id_L - difference_disjoint_L; last first. - { apply disjoint_union_l. split; solve_ndisj. } - iMod ("Hν†" with "Hν") as "H". iModIntro. iNext. iApply fupd_trans. - iMod "H" as "#Hν". - iMod fupd_intro_mask' as "Hclose2"; last iMod ("Hvs" with "Hν") as "$". - { set_solver-. } - iIntros "{$Hν} !>". - iMod "Hclose2" as "_". iApply "Hclose". auto. - Qed. - - Program Definition arc (ty : type) := - {| ty_size := 1; ty_lfts := ty.(ty_lfts); ty_E := ty.(ty_E); - ty_own tid vl := - match vl return _ with - | [ #(LitLoc l) ] => full_arc_own l ty tid ∨ shared_arc_own l ty tid - | _ => False end; - ty_shr κ tid l := - ∃ (l' : loc), &frac{κ} (λ q, l ↦{q} #l') ∗ - □ ∀ F q, ⌜↑shrN ∪ ↑lftN ⊆ F⌝ -∗ q.[κ] - ={F}[F∖↑shrN]▷=∗ q.[κ] ∗ ∃ γ ν q', κ ⊑ ν ∗ - arc_persist tid ν γ l' ty ∗ - &at{κ, arc_shrN}(arc_tok γ q') - |}%I. - Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". Qed. - Next Obligation. - iIntros (ty E κ l tid q ?) "#LFT #Hincl Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done. - (* Ideally, we'd change ty_shr to say "l ↦{q} #l" in the fractured borrow, - but that would be additional work here... *) - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done. - destruct vl as [|[[|l'|]|][|]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - setoid_rewrite heap_mapsto_vec_singleton. - iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx. - iDestruct "Hb" as (i) "(#Hpb&Hpbown)". - set (C := (∃ _ _ _, _ ∗ _ ∗ &at{_,_} _)%I). - iMod (inv_alloc shrN _ (idx_bor_own 1 i ∨ C)%I - with "[Hpbown]") as "#Hinv"; first by iLeft. - iIntros "!> !>" (F q' ?) "Htok". - iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. - iDestruct "INV" as "[>Hbtok|#Hshr]". - - iAssert (&{κ} _)%I with "[Hbtok]" as "Hb". - { rewrite bor_unfold_idx. iExists _. by iFrame. } - iClear "H↦ Hinv Hpb". iDestruct "Htok" as "[Htok1 Htok2]". - iMod (bor_acc_cons with "LFT Hb Htok1") as "[HP Hclose2]"; first solve_ndisj. - iModIntro. iNext. - iAssert (shared_arc_own l' ty tid ∗ (q' / 2).[κ])%I with "[>HP Htok2]" as "[HX $]". - { iDestruct "HP" as "[?|$]"; last done. - iMod (lft_incl_acc with "Hincl Htok2") as (q'') "[Htok Hclose]"; first solve_ndisj. - iMod (arc_own_share with "LFT Htok [$]") as "[Htok $]"; first solve_ndisj. - by iApply "Hclose". } - iDestruct "HX" as (γ ν q'') "[#Hpersist H]". - iMod ("Hclose2" with "[] H") as "[HX $]"; first by unfold shared_arc_own; auto 10. - iAssert C with "[>HX]" as "#$". - { iExists _, _, _. iFrame "Hpersist". - iMod (bor_sep with "LFT HX") as "[Hrc Hlft]"; first solve_ndisj. - iDestruct (frac_bor_lft_incl with "LFT [> Hlft]") as "$". - { iApply (bor_fracture with "LFT"); first solve_ndisj. by rewrite Qp_mul_1_r. } - iApply (bor_share with "Hrc"); solve_ndisj. } - iApply ("Hclose1" with "[]"). by auto. - - iMod ("Hclose1" with "[]") as "_"; first by auto. - iApply step_fupd_intro; first solve_ndisj. by iFrame. - Qed. - Next Obligation. - iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]". - iExists _. iSplit; first by iApply frac_bor_shorten. - iModIntro. iIntros (F q) "% Htok". - iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj. - iMod ("Hshr" with "[] Htok") as "Hshr2"; first done. - iModIntro. iNext. iMod "Hshr2" as "[Htok HX]". - iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν ?) "(? & ? & ?)". - iExists _, _, _. iModIntro. iFrame. iSplit. - - by iApply lft_incl_trans. - - by iApply at_bor_shorten. - Qed. - - Global Instance arc_type_contractive : TypeContractive arc. - Proof. - split. - - apply (type_lft_morphism_add _ static [] [])=>?. - + rewrite left_id. iApply lft_equiv_refl. - + by rewrite /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid vl. destruct vl as [|[[|l|]|] [|]]=>//=. - rewrite /full_arc_own /shared_arc_own /arc_persist Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ α, ty1.(ty_shr) (α ⊓ ty_lft ty1) tid (l +ₗ 2) ≡{n}≡ - ty2.(ty_shr) (α ⊓ ty_lft ty2) tid (l +ₗ 2)) as Hs'. - { intros. rewrite Hs. apply equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - repeat (apply Ho || apply dist_S, Ho || apply Hs' || f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. rewrite /= /arc_persist Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ l α, dist_later n (ty1.(ty_shr) (α ⊓ ty_lft ty1) tid (l +ₗ 2)) - (ty2.(ty_shr) (α ⊓ ty_lft ty2) tid (l +ₗ 2))) as Hs'. - { intros. rewrite Hs. apply dist_dist_later, equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - repeat (apply dist_S, Ho || apply Hs' || f_contractive || f_equiv). - Qed. - - Global Instance arc_ne : NonExpansive arc. - Proof. - unfold arc, full_arc_own, shared_arc_own, arc_persist. - intros n x y Hxy. constructor; [done|by apply Hxy..| |]. - - intros. rewrite ![X in X {| ty_own := _ |}]/ty_own. - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - Qed. - - Lemma arc_subtype ty1 ty2 : - type_incl ty1 ty2 -∗ type_incl (arc ty1) (arc ty2). - Proof. - iIntros "#Hincl". iPoseProof "Hincl" as "(#Hsz & Hl & #Hoincl & #Hsincl)". - iSplit; [done|]. iSplit; [done|]. iSplit; iModIntro. - - iIntros "* Hvl". destruct vl as [|[[|vl|]|] [|]]; try done. - iDestruct "Hvl" as "[(Hl1 & Hl2 & H† & Hc) | Hvl]". - { iLeft. iFrame. iDestruct "Hsz" as %->. - iFrame. iApply (heap_mapsto_pred_wand with "Hc"). iApply "Hoincl". } - iDestruct "Hvl" as (γ ν 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". - iModIntro. iNext. iMod "H" as "[$ H]". - iDestruct "H" as (γ ν q') "(Hlft & Hpersist & Hna)". - iExists _, _, _. iFrame. by iApply arc_persist_type_incl. - Qed. - - Global Instance arc_mono E L : - Proper (subtype E L ==> subtype E L) arc. - Proof. - iIntros (ty1 ty2 Hsub ?) "HL". iDestruct (Hsub with "HL") as "#Hsub". - iIntros "!> #HE". iApply arc_subtype. by iApply "Hsub". - Qed. - Global Instance arc_proper E L : - Proper (eqtype E L ==> eqtype E L) arc. - Proof. intros ??[]. by split; apply arc_mono. Qed. - - Global Instance arc_send ty : - Send ty → Sync ty → Send (arc ty). - Proof. - intros Hse Hsy tid tid' vl. destruct vl as [|[[|l|]|] []]=>//=. - unfold full_arc_own, shared_arc_own. - repeat (apply send_change_tid || apply bi.exist_mono || - (apply arc_persist_send; apply _) || f_equiv || intros ?). - Qed. - - Global Instance arc_sync ty : - Send ty → Sync ty → Sync (arc ty). - Proof. - intros Hse Hsy ν tid tid' l. - repeat (apply send_change_tid || apply bi.exist_mono || - (apply arc_persist_send; apply _) || f_equiv || intros ?). - Qed. - - (* Model of weak *) - Program Definition weak (ty : type) := - {| ty_size := 1; ty_lfts := ty.(ty_lfts); ty_E := ty.(ty_E); - ty_own tid vl := - match vl return _ with - | [ #(LitLoc l) ] => ∃ γ ν, arc_persist tid ν γ l ty ∗ weak_tok γ - | _ => False end; - ty_shr κ tid l := - ∃ (l' : loc), &frac{κ} (λ q, l ↦{q} #l') ∗ - □ ∀ F q, ⌜↑shrN ∪ ↑lftN ⊆ F⌝ -∗ q.[κ] - ={F}[F∖↑shrN]▷=∗ q.[κ] ∗ ∃ γ ν, - arc_persist tid ν γ l' ty ∗ &at{κ, arc_shrN}(weak_tok γ) - |}%I. - Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". Qed. - Next Obligation. - iIntros (ty E κ l tid q ?) "#LFT Hincl Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done. - (* Ideally, we'd change ty_shr to say "l ↦{q} #l" in the fractured borrow, - but that would be additional work here... *) - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done. - destruct vl as [|[[|l'|]|][|]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - setoid_rewrite heap_mapsto_vec_singleton. - iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx. - iDestruct "Hb" as (i) "(#Hpb&Hpbown)". - set (C := (∃ _ _, _ ∗ &at{_,_} _)%I). - iMod (inv_alloc shrN _ (idx_bor_own 1 i ∨ C)%I - with "[Hpbown]") as "#Hinv"; first by iLeft. - iIntros "!> !> * % Htok". - iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. - iDestruct "INV" as "[>Hbtok|#Hshr]". - - iAssert (&{κ} _)%I with "[Hbtok]" as "Hb". - { rewrite bor_unfold_idx. iExists _. by iFrame. } - iClear "H↦ Hinv Hpb". - iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose2]"; first solve_ndisj. - iDestruct "HP" as (γ ν) "[#Hpersist H]". - iMod ("Hclose2" with "[] H") as "[HX $]"; first by auto 10. - iModIntro. iNext. iAssert C with "[>HX]" as "#$". - { iExists _, _. iFrame "Hpersist". iApply bor_share; solve_ndisj. } - iApply ("Hclose1" with "[]"). by auto. - - iMod ("Hclose1" with "[]") as "_"; first by auto. - iApply step_fupd_intro; first solve_ndisj. by iFrame. - Qed. - Next Obligation. - iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]". - iExists _. iSplit; first by iApply frac_bor_shorten. - iModIntro. iIntros (F q) "% Htok". - iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj. - iMod ("Hshr" with "[] Htok") as "Hshr2"; first done. - iModIntro. iNext. iMod "Hshr2" as "[Htok HX]". - iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν) "[??]". - iExists _, _. iModIntro. iFrame. by iApply at_bor_shorten. - Qed. - - Global Instance weak_type_contractive : TypeContractive weak. - Proof. - split. - - apply (type_lft_morphism_add _ static [] [])=>?. - + rewrite left_id. iApply lft_equiv_refl. - + by rewrite /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid vl. destruct vl as [|[[|l|]|] [|]]=>//=. - rewrite /arc_persist Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ α, ty1.(ty_shr) (α ⊓ ty_lft ty1) tid (l +ₗ 2) ≡{n}≡ - ty2.(ty_shr) (α ⊓ ty_lft ty2) tid (l +ₗ 2)) as Hs'. - { intros. rewrite Hs. apply equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - repeat (apply Ho || apply dist_S, Ho || apply Hs' || f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. rewrite /= /arc_persist Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ l α, dist_later n (ty1.(ty_shr) (α ⊓ ty_lft ty1) tid (l +ₗ 2)) - (ty2.(ty_shr) (α ⊓ ty_lft ty2) tid (l +ₗ 2))) as Hs'. - { intros. rewrite Hs. apply dist_dist_later, equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - repeat (apply dist_S, Ho || apply Hs' || f_contractive || f_equiv). - Qed. - - Global Instance weak_ne : NonExpansive weak. - Proof. - unfold weak, arc_persist. intros n x y Hxy. constructor; [done|by apply Hxy..| |]. - - intros. rewrite ![X in X {| ty_own := _ |}]/ty_own. - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - Qed. - - Lemma weak_subtype ty1 ty2 : - type_incl ty1 ty2 -∗ type_incl (weak ty1) (weak ty2). - Proof. - iIntros "#Hincl". iPoseProof "Hincl" as "(#Hsz & #Hoincl & #Hsincl)". - iSplit; [done|]. iSplit; [done|]. iSplit; iModIntro. - - iIntros "* Hvl". destruct vl as [|[[|vl|]|] [|]]; try done. - 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". - iModIntro. iNext. iMod "H" as "[$ H]". iDestruct "H" as (γ ν) "[Hpersist Hna]". - iExists _, _. iFrame. by iApply arc_persist_type_incl. - Qed. - - Global Instance weak_mono E L : - Proper (subtype E L ==> subtype E L) weak. - Proof. - iIntros (ty1 ty2 Hsub ?) "HL". iDestruct (Hsub with "HL") as "#Hsub". - iIntros "!> #HE". iApply weak_subtype. by iApply "Hsub". - Qed. - Global Instance weak_proper E L : - Proper (eqtype E L ==> eqtype E L) weak. - Proof. intros ??[]. by split; apply weak_mono. Qed. - - Global Instance weak_send ty : - Send ty → Sync ty → Send (weak ty). - Proof. - intros Hse Hsy tid tid' vl. destruct vl as [|[[|l|]|] []]=>//=. - repeat (apply send_change_tid || apply bi.exist_mono || - (apply arc_persist_send; apply _) || f_equiv || intros ?). - Qed. - - Global Instance weak_sync ty : - Send ty → Sync ty → Sync (weak ty). - Proof. - intros Hse Hsy ν tid tid' l. - repeat (apply send_change_tid || apply bi.exist_mono || - (apply arc_persist_send; apply _) || f_equiv || intros ?). - Qed. - - (* Code : constructors *) - Definition arc_new ty : val := - fn: ["x"] := - let: "arcbox" := new [ #(2 + ty.(ty_size))%nat ] in - let: "arc" := new [ #1 ] in - "arcbox" +ₗ #0 <- #1;; - "arcbox" +ₗ #1 <- #1;; - "arcbox" +ₗ #2 <-{ty.(ty_size)} !"x";; - "arc" <- "arcbox";; - delete [ #ty.(ty_size); "x"];; return: ["arc"]. - - Lemma arc_new_type ty : - typed_val (arc_new ty) (fn(∅; ty) → arc ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (rc); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]". - rewrite !tctx_hasty_val. - iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x. - iDestruct (ownptr_uninit_own with "Hrcbox") - as (lrcbox vlrcbox) "(% & Hrcbox↦ & Hrcbox†)". subst rcbox. inv_vec vlrcbox=>???. - iDestruct (heap_mapsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]". - iDestruct (heap_mapsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]". - rewrite shift_loc_assoc. - iDestruct (ownptr_uninit_own with "Hrc") as (lrc vlrc) "(% & Hrc↦ & Hrc†)". subst rc. - inv_vec vlrc=>?. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write. - wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz. - wp_apply (wp_memcpy with "[$Hrcbox↦2 $Hx↦]"); [by auto using vec_to_list_length..|]. - iIntros "[Hrcbox↦2 Hx↦]". wp_seq. wp_write. - (* Finish up the proof. *) - iApply (type_type _ _ _ [ #lx ◁ box (uninit (ty_size ty)); #lrc ◁ box (arc ty)] - with "[] LFT HE Hna HL Hk [>-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=. iFrame. - iSplitL "Hx↦"; first by iExists _; rewrite ->uninit_own; auto. - iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. iLeft. - rewrite freeable_sz_full_S. iFrame. iExists _. by iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition weak_new ty : val := - fn: [] := - let: "arcbox" := new [ #(2 + ty.(ty_size))%nat ] in - let: "w" := new [ #1 ] in - "arcbox" +ₗ #0 <- #0;; - "arcbox" +ₗ #1 <- #1;; - "w" <- "arcbox";; - return: ["w"]. - - Lemma weak_new_type ty : - typed_val (weak_new ty) (fn(∅) → weak ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg. simpl_subst. - iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (rc); simpl_subst. - iIntros (tid) "#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]". - iDestruct (heap_mapsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]". - rewrite shift_loc_assoc. - iDestruct (ownptr_uninit_own with "Hrc") as (lrc vlrc) "(% & Hrc↦ & Hrc†)". subst rc. - inv_vec vlrc=>?. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - iMod (lft_create with "LFT") as (ν) "[Hν Hν†]"; first done. - wp_bind (_ +ₗ _)%E. iSpecialize ("Hν†" with "Hν"). - iApply wp_mask_mono; last iApply (wp_step_fupd with "Hν†"); [set_solver-..|]. - wp_op. iIntros "#H† !>". rewrite shift_loc_0. wp_write. wp_op. wp_write. wp_write. - iApply (type_type _ _ _ [ #lrc ◁ box (weak ty)] - with "[] LFT HE Hna HL Hk [>-]"); last first. - { rewrite tctx_interp_singleton tctx_hasty_val' //=. iFrame. - iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iMod (create_weak (P1 ν) (P2 lrcbox ty.(ty_size)) - with "Hrcbox↦0 Hrcbox↦1 [-]") as (γ) "[Ha Htok]". - { rewrite freeable_sz_full_S. iFrame. iExists _. iFrame. - by rewrite vec_to_list_length. } - iExists γ, ν. iFrame. iSplitL; first by auto. iIntros "!>!>!> Hν". - iDestruct (lft_tok_dead with "Hν H†") as "[]". } - iApply type_jump; solve_typing. - Qed. - - (* Code : deref *) - Definition arc_deref : val := - fn: ["arc"] := - let: "x" := new [ #1 ] in - let: "arc'" := !"arc" in - "x" <- (!"arc'" +ₗ #2);; - delete [ #1; "arc" ];; return: ["x"]. - - Lemma arc_deref_type ty : - typed_val arc_deref (fn(∀ α, ∅; &shr{α}(arc ty)) → &shr{α}ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (x); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#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. - destruct rc' as [[|rc'|]|]; try done. simpl of_val. - iDestruct "Hrc'" as (l') "[Hrc' #Hdelay]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hdelay" with "[] Hα1"); last iApply (wp_step_fupd with "Hdelay"); [done..|]. - iMod (frac_bor_acc with "LFT Hrc' Hα2") as (q') "[Hrc'↦ Hclose2]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose2" with "[$Hrc'↦]") as "Hα2". iIntros "!> [Hα1 #Hproto] !>". - iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & _)". - iDestruct "Hpersist" as "(_ & [Hshr|Hν†] & _)"; last first. - { iMod (lft_incl_dead with "Hαν Hν†") as "Hα†". done. - iDestruct (lft_tok_dead with "Hα1 Hα†") as "[]". } - wp_op. wp_write. iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ rcx ◁ box (&shr{α}(arc ty)); #lrc2 ◁ box (&shr{α}ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame "Hrcx". iFrame "Hx†". iExists [_]. rewrite heap_mapsto_vec_singleton. - iFrame "Hx". iApply (ty_shr_mono with "[] Hshr"). - iApply lft_incl_glb; [done|]. iApply elctx_interp_ty_outlives_E. - rewrite !elctx_interp_app /=. iDestruct "HE" as "(_ & [[_ $] _] & _)". } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Code : getters *) - Definition arc_strong_count : val := - fn: ["arc"] := - let: "r" := new [ #1 ] in - let: "arc'" := !"arc" in - let: "arc''" := !"arc'" in - "r" <- strong_count ["arc''"];; - delete [ #1; "arc" ];; return: ["r"]. - - Lemma arc_strong_count_type ty : - typed_val arc_strong_count (fn(∀ α, ∅; &shr{α}(arc ty)) → int). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let. - wp_apply (strong_count_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I - with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!> Hα". - iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|]. - iExists _. iFrame. iMod fupd_intro_mask' as "Hclose2"; last iModIntro. set_solver. - iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". } - iIntros (c) "[Hα _]". iMod ("Hclose1" with "Hα HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(arc ty)); #c ◁ int; r ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition arc_weak_count : val := - fn: ["arc"] := - let: "r" := new [ #1 ] in - let: "arc'" := !"arc" in - let: "arc''" := !"arc'" in - "r" <- weak_count ["arc''"];; - delete [ #1; "arc" ];; return: ["r"]. - - Lemma arc_weak_count_type ty : - typed_val arc_weak_count (fn(∀ α, ∅; &shr{α}(arc ty)) → int). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let. - wp_apply (weak_count_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I - with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!> Hα". - iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|]. - iExists _. iFrame. iMod fupd_intro_mask' as "Hclose2"; last iModIntro. set_solver. - iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". } - iIntros (c) "[Hα _]". iMod ("Hclose1" with "Hα HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(arc ty)); #c ◁ int; r ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Code : clone, [up/down]grade. *) - Definition arc_clone : val := - fn: ["arc"] := - let: "r" := new [ #1 ] in - let: "arc'" := !"arc" in - let: "arc''" := !"arc'" in - clone_arc ["arc''"];; - "r" <- "arc''";; - delete [ #1; "arc" ];; return: ["r"]. - - Lemma arc_clone_type ty : - typed_val arc_clone (fn(∀ α, ∅; &shr{α}(arc ty)) → arc ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let. - wp_apply (clone_arc_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I - with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!> Hα". - iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|]. - iExists _. iFrame. iMod fupd_intro_mask' as "Hclose2"; last iModIntro. set_solver. - iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". } - iIntros (q'') "[Hα Hown]". wp_seq. iMod ("Hclose1" with "Hα HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(arc ty)); #l' ◁ arc ty; r ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. simpl. unfold shared_arc_own. auto 10 with iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition weak_clone : val := - fn: ["w"] := - let: "r" := new [ #1 ] in - let: "w'" := !"w" in - let: "w''" := !"w'" in - clone_weak ["w''"];; - "r" <- "w''";; - delete [ #1; "w" ];; return: ["r"]. - - Lemma weak_clone_type ty : - typed_val weak_clone (fn(∀ α, ∅; &shr{α}(weak ty)) → weak ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν) "#[Hpersist Hrctokb]". iModIntro. wp_let. - wp_apply (clone_weak_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I - with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!> Hα". - iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>$ Hclose1]"; [solve_ndisj..|]. - iMod fupd_intro_mask' as "Hclose2"; last iModIntro. set_solver. - iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". } - iIntros "[Hα Hown]". wp_seq. iMod ("Hclose1" with "Hα HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(weak ty)); #l' ◁ weak ty; r ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. simpl. eauto 10 with iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition downgrade : val := - fn: ["arc"] := - let: "r" := new [ #1 ] in - let: "arc'" := !"arc" in - let: "arc''" := !"arc'" in - downgrade ["arc''"];; - "r" <- "arc''";; - delete [ #1; "arc" ];; return: ["r"]. - - Lemma downgrade_type ty : - typed_val downgrade (fn(∀ α, ∅; &shr{α}(arc ty)) → weak ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q') "#(Hαν & Hpersist & Hrctokb)". iModIntro. wp_let. - wp_apply (downgrade_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I - with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!> Hα". - iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>Htok Hclose1]"; [solve_ndisj..|]. - iExists _. iFrame. iMod fupd_intro_mask' as "Hclose2"; last iModIntro. set_solver. - iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". } - iIntros "[Hα Hown]". wp_seq. iMod ("Hclose1" with "Hα HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(arc ty)); #l' ◁ weak ty; r ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. simpl. eauto 10 with iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition upgrade : val := - fn: ["w"] := - let: "r" := new [ #2 ] in - let: "w'" := !"w" in - let: "w''" := !"w'" in - if: upgrade ["w''"] then - "r" <-{Σ some} "w''";; - delete [ #1; "w" ];; return: ["r"] - else - "r" <-{Σ none} ();; - delete [ #1; "w" ];; return: ["r"]. - - Lemma upgrade_type ty : - typed_val upgrade (fn(∀ α, ∅; &shr{α}(weak ty)) → option (arc ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". clear q'. iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν) "#[Hpersist Hrctokb]". iModIntro. wp_let. - wp_apply (upgrade_spec (P1 ν) (P2 l' ty.(ty_size)) _ _ _ (q.[α])%I - with "[] [] [$Hα1 $Hα2]"); first by iDestruct "Hpersist" as "[$ _]". - { iIntros "!> Hα". - iMod (at_bor_acc_tok with "LFT Hrctokb Hα") as "[>$ Hclose1]"; [solve_ndisj..|]. - iMod fupd_intro_mask' as "Hclose2"; last iModIntro. set_solver. - iIntros "Htok". iMod "Hclose2" as "_". by iApply "Hclose1". } - iIntros ([] q') "[Hα Hown]"; wp_if; iMod ("Hclose1" with "Hα HL") as "HL". - - (* Finish up the proof (sucess). *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(weak ty)); #l' ◁ arc ty; - r ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. simpl. unfold shared_arc_own. eauto 10 with iFrame. } - iApply (type_sum_assign (option (arc ty))); [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - - (* Finish up the proof (fail). *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(weak ty)); r ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - unlock. iFrame. } - iApply (type_sum_unit (option (arc ty))); [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Code : drop *) - Definition arc_drop ty : val := - fn: ["arc"] := - let: "r" := new [ #(option ty).(ty_size) ] in - let: "arc'" := !"arc" in - "r" <-{Σ none} ();; - (if: drop_arc ["arc'"] then - "r" <-{ty.(ty_size),Σ some} !("arc'" +ₗ #2);; - if: drop_weak ["arc'"] then - delete [ #(2 + ty.(ty_size)); "arc'" ] - else #☠ - else #☠);; - delete [ #1; "arc"];; return: ["r"]. - - Lemma arc_drop_type ty : - typed_val (arc_drop ty) (fn(∅; arc ty) → option ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new (option ty).(ty_size)); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iApply (type_sum_unit (option ty)); [solve_typing..|]. - iIntros (tid) "#LFT #HE Hna HL Hk [Hr [Hrcx [Hrc' _]]]". - rewrite !tctx_hasty_val. destruct rc' as [[|rc'|]|]=>//=. - iAssert (shared_arc_own rc' ty tid ∗ llctx_interp [ϝ ⊑ₗ []] 1)%I - with "[>Hrc' HL]" as "[Hrc' HL]". - { iDestruct "Hrc'" as "[?|$]"; last done. - iMod (lctx_lft_alive_tok ty_lft ty with "HE HL") - as (?) "(Htok & HL & Hclose)"; [done| |]. - { change (ty_outlives_E (arc ty)) with (ty_outlives_E ty). - eapply (lctx_lft_alive_incl _ _ ϝ); first solve_typing. - iIntros (?) "_ !>". rewrite !elctx_interp_app /=. - iIntros "(_ & _ & [? _] & _ & _)". by iApply elctx_interp_ty_outlives_E. } - iMod (arc_own_share with "LFT Htok [$]") as "[Htok $]"; first solve_ndisj. - by iApply ("Hclose" with "Htok"). } - iDestruct "Hrc'" as (γ ν q) "(#Hpersist & Htok & Hν)". - wp_bind (drop_arc _). iApply (drop_arc_spec with "[] [$Htok Hν]"); - [by iDestruct "Hpersist" as "[$?]"|done|]. - iNext. iIntros (b) "Hdrop". wp_bind (if: _ then _ else _)%E. - iApply (wp_wand _ _ _ (λ _, ty_own (box (option ty)) tid [r])%I with "[Hdrop Hr]"). - { destruct b; wp_if=>//. destruct r as [[]|]; try done. - (* FIXME: 'simpl' reveals a match here. Didn't we forbid that for ty_own? *) - rewrite {3}[option]lock. simpl. rewrite -{2 3}lock. (* FIXME: Tried using contextual pattern, did not work. *) - iDestruct "Hr" as "[Hr ?]". iDestruct "Hr" as ([|d vl]) "[H↦ Hown]"; - first by iDestruct "Hown" as (???) "[>% ?]". - rewrite heap_mapsto_vec_cons. iDestruct "H↦" as "[H↦0 H↦1]". - iDestruct "Hpersist" as "(? & ? & H†)". wp_bind (_ <- _)%E. - iDestruct "Hdrop" as "[Hν Hdrop]". iSpecialize ("H†" with "Hν"). - iApply wp_mask_mono; last iApply (wp_step_fupd with "H†"); first done. - { set_solver-. (* FIXME [solve_ndisj] fails *) } - iDestruct "Hown" as (???) "(_ & Hlen & _)". wp_write. iIntros "(#Hν & Hown & H†)!>". - wp_seq. wp_op. wp_op. iDestruct "Hown" as (?) "[H↦ Hown]". - iDestruct (ty_size_eq with "Hown") as %?. rewrite Max.max_0_r. - iDestruct "Hlen" as %[= ?]. wp_apply (wp_memcpy with "[$H↦1 $H↦]"); [congruence..|]. - iIntros "[? Hrc']". wp_seq. iMod ("Hdrop" with "[Hrc' H†]") as "Htok". - { unfold P2. auto with iFrame. } - wp_apply (drop_weak_spec with "[//] Htok"). unlock. iIntros ([]); last first. - { iIntros "_". wp_if. unlock. iFrame. iExists (_::_). rewrite heap_mapsto_vec_cons. - iFrame. iExists 1%nat, _, []. rewrite /= !right_id_L Max.max_0_r. - auto 10 with iFrame. } - iIntros "([H† H1] & H2 & H3)". iDestruct "H1" as (vl1) "[H1 Heq]". - iDestruct "Heq" as %<-. wp_if. - wp_apply (wp_delete _ _ _ (_::_::vl1) with "[H1 H2 H3 H†]"). - { simpl. lia. } - { rewrite 2!heap_mapsto_vec_cons shift_loc_assoc. auto with iFrame. } - iFrame. iIntros "_". iExists (_::_). rewrite heap_mapsto_vec_cons. iFrame. - iExists 1%nat, _, []. rewrite !right_id_L. iFrame. iSplit; [by auto|simpl]. - auto with lia. } - iIntros (?) "Hr". wp_seq. - (* Finish up the proof. *) - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); r ◁ box (option ty) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. unlock. iFrame. - by rewrite tctx_hasty_val. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition weak_drop ty : val := - fn: ["arc"] := - let: "r" := new [ #0] in - let: "arc'" := !"arc" in - (if: drop_weak ["arc'"] then delete [ #(2 + ty.(ty_size)); "arc'" ] - else #☠);; - delete [ #1; "arc"];; return: ["r"]. - - Lemma weak_drop_type ty : - typed_val (weak_drop ty) (fn(∅; weak ty) → unit). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 0); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. destruct rc' as [[|rc'|]|]=>//=. - iDestruct "Hrc'" as (γ ν) "[#Hpersist Htok]". wp_bind (drop_weak _). - iApply (drop_weak_spec with "[] [Htok]"); [by iDestruct "Hpersist" as "[$?]"|by auto|]. - iIntros "!> * Hdrop". wp_bind (if: _ then _ else _)%E. - iApply (wp_wand _ _ _ (λ _, True)%I with "[Hdrop]"). - { destruct b; wp_if=>//. iDestruct "Hdrop" as "((? & H↦) & ? & ?)". - iDestruct "H↦" as (vl) "[? Heq]". iDestruct "Heq" as %<-. - wp_apply (wp_delete _ _ _ (_::_::vl) with "[-]"); [simpl;lia| |done]. - rewrite !heap_mapsto_vec_cons shift_loc_assoc. iFrame. } - iIntros (?) "_". wp_seq. - (* Finish up the proof. *) - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); r ◁ box (uninit 0) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. unlock. iFrame. - by rewrite tctx_hasty_val. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Code : other primitives *) - Definition arc_try_unwrap ty : val := - fn: ["arc"] := - let: "r" := new [ #(Σ[ ty; arc ty ]).(ty_size) ] in - let: "arc'" := !"arc" in - if: try_unwrap ["arc'"] then - (* Return content *) - "r" <-{ty.(ty_size),Σ 0} !("arc'" +ₗ #2);; - (* Decrement the "strong weak reference" *) - (if: drop_weak ["arc'"] then delete [ #(2 + ty.(ty_size)); "arc'" ] - else #☠);; - delete [ #1; "arc"];; return: ["r"] - else - "r" <-{Σ 1} "arc'";; - delete [ #1; "arc"];; return: ["r"]. - - Lemma arc_try_unwrap_type ty : - typed_val (arc_try_unwrap ty) (fn(∅; arc ty) → Σ[ ty; arc ty ]). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new (Σ[ ty; arc ty ]).(ty_size)); - [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#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 ∗ llctx_interp [ϝ ⊑ₗ []] 1)%I - with "[>Hrc' HL]" as "[Hrc' HL]". - { iDestruct "Hrc'" as "[?|$]"; last done. - iMod (lctx_lft_alive_tok ty_lft ty with "HE HL") - as (?) "(Htok & HL & Hclose)"; [done| |]. - { change (ty_outlives_E (arc ty)) with (ty_outlives_E ty). - eapply (lctx_lft_alive_incl _ _ ϝ); first solve_typing. - iIntros (?) "_ !>". rewrite !elctx_interp_app /=. - iIntros "(_ & _ & [?_] & _ & _)". by iApply elctx_interp_ty_outlives_E. } - iMod (arc_own_share with "LFT Htok [$]") as "[Htok $]"; first solve_ndisj. - by iApply ("Hclose" with "Htok"). } - iDestruct "Hrc'" as (γ ν q) "(#Hpersist & Htok & Hν)". - wp_apply (try_unwrap_spec with "[] [Hν Htok]"); - [by iDestruct "Hpersist" as "[$?]"|iFrame|]. - iIntros ([]) "H"; wp_if. - - iDestruct "H" as "[Hν Hend]". rewrite -(lock [r]). destruct r as [[|r|]|]=>//=. - iDestruct "Hr" as "[Hr >Hfree]". iDestruct "Hr" as (vl0) "[>Hr Hown]". - iDestruct "Hown" as ">Hlen". - destruct vl0 as [|? vl0]=>//; iDestruct "Hlen" as %[=Hlen0]. - rewrite heap_mapsto_vec_cons. iDestruct "Hr" as "[Hr0 Hr1]". - iDestruct "Hpersist" as "(Ha & _ & H†)". wp_bind (_ <- _)%E. - iSpecialize ("H†" with "Hν"). - iApply wp_mask_mono; last iApply (wp_step_fupd with "H†"); first done. - { (* FIXME [solve_ndisj] fails *) set_solver-. } - wp_write. iIntros "(#Hν & Hown & H†) !>". wp_seq. wp_op. wp_op. - rewrite -(firstn_skipn ty.(ty_size) vl0) heap_mapsto_vec_app. - iDestruct "Hr1" as "[Hr1 Hr2]". iDestruct "Hown" as (vl) "[Hrc' Hown]". - iDestruct (ty_size_eq with "Hown") as %Hlen. - wp_apply (wp_memcpy with "[$Hr1 $Hrc']"); rewrite /= ?firstn_length_le; try lia. - iIntros "[Hr1 Hrc']". wp_seq. wp_bind (drop_weak _). - iMod ("Hend" with "[$H† Hrc']") as "Htok"; first by eauto. - iApply (drop_weak_spec with "Ha Htok"). - iIntros "!> * Hdrop". wp_bind (if: _ then _ else _)%E. - iApply (wp_wand _ _ _ (λ _, True)%I with "[Hdrop]"). - { destruct b; wp_if=>//. iDestruct "Hdrop" as "((? & H↦) & ? & ?)". - iDestruct "H↦" as (vl') "[? Heq]". iDestruct "Heq" as %<-. - wp_apply (wp_delete _ _ _ (_::_::vl') with "[-]"); [simpl; lia| |done]. - rewrite !heap_mapsto_vec_cons shift_loc_assoc. iFrame. } - iIntros (?) "_". wp_seq. - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); #r ◁ box (Σ[ ty; arc ty ]) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val. unlock. - iFrame. iCombine "Hr1" "Hr2" as "Hr1". iCombine "Hr0" "Hr1" as "Hr". - rewrite -[in _ +ₗ ty.(ty_size)]Hlen -heap_mapsto_vec_app - -heap_mapsto_vec_cons tctx_hasty_val' //. iFrame. iExists _. iFrame. - iExists O, _, _. iSplit; first by auto. iFrame. iIntros "!> !% /=". - rewrite app_length drop_length. lia. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); #rc' ◁ arc ty; - r ◁ box (uninit (Σ[ ty; arc ty ]).(ty_size)) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. iRight. iExists _, _, _. auto with iFrame. } - iApply (type_sum_assign Σ[ ty; arc ty ]); [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition arc_get_mut : val := - fn: ["arc"] := - let: "r" := new [ #2 ] in - let: "arc'" := !"arc" in - let: "arc''" := !"arc'" in - if: is_unique ["arc''"] then - Skip;; - (* Return content *) - "r" <-{Σ some} "arc''" +ₗ #2;; - delete [ #1; "arc"];; return: ["r"] - else - "r" <-{Σ none} ();; - delete [ #1; "arc"];; return: ["r"]. - - Lemma arc_get_mut_type ty : - typed_val arc_get_mut (fn(∀ α, ∅; &uniq{α}(arc ty)) → option (&uniq{α}ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. - iDestruct "Hrc'" as "[#? Hrc']". destruct rc' as [[|rc'|]|]=>//=. - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & HL & Hclose1)"; - [solve_typing..|]. - iMod (bor_exists with "LFT Hrc'") as (rcvl) "Hrc'"=>//. - iMod (bor_sep with "LFT Hrc'") as "[Hrc'↦ Hrc]"=>//. - destruct rcvl as [|[[|rc|]|][|]]; try by - iMod (bor_persistent with "LFT Hrc Hα") as "[>[] ?]". - rewrite heap_mapsto_vec_singleton. - iMod (bor_acc with "LFT Hrc'↦ Hα") as "[Hrc'↦ Hclose2]"=>//. wp_read. - iMod ("Hclose2" with "Hrc'↦") as "[_ [Hα1 Hα2]]". - iMod (bor_acc_cons with "LFT Hrc Hα1") as "[Hrc Hclose2]"=>//. wp_let. - iAssert (shared_arc_own rc ty tid ∗ (q / 2).[α])%I with "[>Hrc Hα2]" as "[Hrc Hα2]". - { iDestruct "Hrc" as "[Hrc|$]"=>//. - iMod (lft_incl_acc with "[//] Hα2") as (q') "[Htok Hclose]"; [solve_ndisj|]. - iMod (arc_own_share with "LFT Htok Hrc") as "[Htok $]"; [solve_ndisj|]. - by iApply "Hclose". } - iDestruct "Hrc" as (γ ν q') "[#(Hi & Hs & #Hc) Htoks]". - wp_apply (is_unique_spec with "Hi Htoks"). iIntros ([]) "H"; wp_if. - - iDestruct "H" as "(Hrc & Hrc1 & Hν)". iSpecialize ("Hc" with "Hν"). wp_bind Skip. - iApply wp_mask_mono; last iApply (wp_step_fupd with "Hc"); first done. - { (* FIXME [solve_ndisj] fails *) set_solver-. } - wp_seq. iIntros "(#Hν & Hown & H†) !>". wp_seq. - iMod ("Hclose2" with "[Hrc Hrc1 H†] Hown") as "[Hb Hα1]". - { iIntros "!> Hown !>". iLeft. iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); #rc +ₗ #2 ◁ &uniq{α}ty; - r ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. unlock. by iFrame. } - iApply (type_sum_assign (option (&uniq{α}ty))); [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - - iMod ("Hclose2" with "[] [H]") as "[_ Hα1]"; - [by iIntros "!> H"; iRight; iApply "H"|iExists _, _, _; iFrame "∗#"|]. - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); r ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. unlock. iFrame. } - iApply (type_sum_unit (option (&uniq{α}ty))); [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition arc_make_mut (ty : type) (clone : val) : val := - fn: ["arc"] := - let: "r" := new [ #1 ] in - let: "arc'" := !"arc" in - let: "arc''" := !"arc'" in - (case: try_unwrap_full ["arc''"] of - [ "r" <- "arc''" +ₗ #2;; - delete [ #1; "arc"];; return: ["r"]; - - let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in - "rcbox" +ₗ #0 <- #1;; - "rcbox" +ₗ #1 <- #1;; - "rcbox" +ₗ #2 <-{ty.(ty_size)} !"arc''" +ₗ #2;; - "arc'" <- "rcbox";; - "r" <- "rcbox" +ₗ #2;; - delete [ #1; "arc"];; return: ["r"]; - - letalloc: "x" <- "arc''" +ₗ #2 in - letcall: "c" := clone ["x"]%E in (* FIXME : why do I need %E here ? *) - let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in - "rcbox" +ₗ #0 <- #1;; - "rcbox" +ₗ #1 <- #1;; - "rcbox" +ₗ #2 <-{ty.(ty_size)} !"c";; - delete [ #ty.(ty_size); "c"];; - "arc'" <- "rcbox";; - letalloc: "rcold" <- "arc''" in - (* FIXME : here, we are dropping the old arc pointer. In the - case another strong reference has been dropped while - cloning, it is possible that we are actually dropping the - last reference here. This means that we may have to drop an - instance of [ty] here. Instead, we simply forget it. *) - let: "drop" := arc_drop ty in - letcall: "content" := "drop" ["rcold"]%E in - delete [ #(option ty).(ty_size); "content"];; - "r" <- "rcbox" +ₗ #2;; - delete [ #1; "arc"];; return: ["r"] - ]). - - Lemma arc_make_mut_type ty clone : - typed_val clone (fn(∀ α, ∅; &shr{α}ty) → ty) → - typed_val (arc_make_mut ty clone) (fn(∀ α, ∅; &uniq{α}(arc ty)) → &uniq{α} ty). - Proof. - intros Hclone E L. iApply type_fn; [solve_typing..|]. rewrite [(2 + ty_size ty)%nat]lock. - iIntros "/= !>". iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. - iDestruct "Hrc'" as "[#? Hrc']". destruct rc' as [[|rc'|]|]=>//=. - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. - iMod (bor_acc_cons with "LFT Hrc' Hα1") as "[Hrc' Hclose2]"=>//. - iDestruct "Hrc'" as (rcvl) "[Hrc'↦ Hrc]". - destruct rcvl as [|[[|rc|]|][|]]; try by iDestruct "Hrc" as ">[]". - rewrite heap_mapsto_vec_singleton. wp_read. - iAssert (shared_arc_own rc ty tid ∗ (q / 2).[α])%I with "[>Hrc Hα2]" as "[Hrc Hα2]". - { iDestruct "Hrc" as "[Hrc|$]"=>//. - iMod (lft_incl_acc with "[//] Hα2") as (q') "[Htok Hclose]"; [solve_ndisj|]. - iMod (arc_own_share with "LFT Htok Hrc") as "[Htok $]"; [solve_ndisj|]. - by iApply "Hclose". } - iDestruct "Hrc" as (γ ν q') "[#(Hi & Hs & #Hc) Htoks]". wp_let. - wp_apply (try_unwrap_full_spec with "Hi Htoks"). iIntros (x). - pose proof (fin_to_nat_lt x). destruct (fin_to_nat x) as [|[|[]]]; last lia. - - iIntros "(Hrc0 & Hrc1 & HP1)". wp_case. wp_bind (_ +ₗ _)%E. - iSpecialize ("Hc" with "HP1"). - iApply wp_mask_mono; last iApply (wp_step_fupd with "Hc"); first done. - { (* FIXME [solve_ndisj] fails *) set_solver-. } - wp_op. iIntros "(#Hν & Hrc2 & H†)". iModIntro. - iMod ("Hclose2" with "[Hrc'↦ Hrc0 Hrc1 H†] Hrc2") as "[Hb Hα1]". - { iIntros "!> Hrc2". iExists [_]. rewrite heap_mapsto_vec_singleton. - iFrame. iLeft. by iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); #(rc +ₗ 2) ◁ &uniq{α}ty; - r ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. unlock. by iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - - iIntros "[Hν Hweak]". wp_case. iSpecialize ("Hc" with "Hν"). wp_bind (new _). - iApply wp_mask_mono; last iApply (wp_step_fupd with "Hc"); first done. - { (* FIXME [solve_ndisj] fails *) set_solver-. } - wp_apply wp_new=>//. lia. iIntros (l) "(H† & Hlvl) (#Hν & Hown & H†') !>". - rewrite -!lock Nat2Z.id. - wp_let. wp_op. rewrite !heap_mapsto_vec_cons shift_loc_assoc shift_loc_0. - iDestruct "Hlvl" as "(Hrc0 & Hrc1 & Hrc2)". wp_write. wp_op. wp_write. - wp_op. wp_op. iDestruct "Hown" as (vl) "[H↦ Hown]". - iDestruct (ty_size_eq with "Hown") as %Hsz. - wp_apply (wp_memcpy with "[$Hrc2 $H↦]"). - { by rewrite repeat_length. } { by rewrite Hsz. } - iIntros "[H↦ H↦']". wp_seq. wp_write. - iMod ("Hclose2" $! ((l +ₗ 2) ↦∗: ty_own ty tid)%I - with "[Hrc'↦ Hrc0 Hrc1 H†] [H↦ Hown]") as "[Hb Hα1]"; [|by auto with iFrame|]. - { iIntros "!> H↦". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iLeft. iFrame. iDestruct "H†" as "[?|%]"=>//. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ rcx ◁ box (uninit 1); (#l +ₗ #2) ◁ &uniq{α}ty; - r ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. unlock. by iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - - iIntros "[Htok Hν]". wp_case. wp_apply wp_new=>//. - iIntros (l) "/= (H† & Hl)". wp_let. wp_op. - rewrite heap_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 //. } - 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ν Hα2") as (q'') "[Hαν Hclose3]". - rewrite -[ν ⊓ α](right_id_L). - iApply (type_call_iris _ [ν ⊓ α] (ν ⊓ α) [_] with - "LFT HE Hna Hαν Hclone [Hl H†]"); [solve_typing| |]. - { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. - iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iApply ty_shr_mono; last done. iApply lft_intersect_mono; [|done]. - iApply lft_incl_refl. } - iIntros ([[|cl|]|]) "Hna Hαν Hcl //". wp_rec. - iDestruct "Hcl" as "[Hcl Hcl†]". iDestruct "Hcl" as (vl) "[Hcl↦ Hown]". - iDestruct (ty_size_eq with "Hown") as %Hsz. - iDestruct ("Hclose3" with "Hαν") as "[Hν Hα2]". - wp_apply wp_new=>//. lia. iIntros (l') "(Hl'† & Hl')". wp_let. wp_op. - rewrite shift_loc_0. rewrite -!lock Nat2Z.id. - rewrite !heap_mapsto_vec_cons shift_loc_assoc. - iDestruct "Hl'" as "(Hl' & Hl'1 & Hl'2)". wp_write. wp_op. wp_write. wp_op. - wp_apply (wp_memcpy with "[$Hl'2 $Hcl↦]"). - { by rewrite repeat_length. } { by rewrite Hsz. } - iIntros "[Hl'2 Hcl↦]". wp_seq. rewrite freeable_sz_full. - wp_apply (wp_delete with "[$Hcl↦ Hcl†]"); - [lia|by replace (length vl) with (ty.(ty_size))|]. - iIntros "_". wp_seq. wp_write. - iMod ("Hclose2" $! ((l' +ₗ 2) ↦∗: ty_own ty tid)%I with - "[Hrc'↦ Hl' Hl'1 Hl'†] [Hl'2 Hown]") as "[Hl' Hα1]". - { iIntros "!> H". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iLeft. iFrame. iDestruct "Hl'†" as "[?|%]"=>//. } - { iExists _. iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ #rc ◁ arc ty; #l' +ₗ #2 ◁ &uniq{α}ty; - r ◁ box (uninit 1); rcx ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 3!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - !tctx_hasty_val' //. unlock. iFrame "∗#". iRight. - iExists _, _, _. iFrame "∗#". } - iApply type_letalloc_1; [solve_typing..|]. iIntros (rcold). simpl_subst. - iApply type_let. apply arc_drop_type. solve_typing. iIntros (drop). simpl_subst. - iApply (type_letcall ()); [solve_typing..|]. iIntros (content). simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. -End arc. diff --git a/theories/typing/lib/cell.v b/theories/typing/lib/cell.v index 7893d0eb..2e2f36b1 100644 --- a/theories/typing/lib/cell.v +++ b/theories/typing/lib/cell.v @@ -52,7 +52,7 @@ Section cell. Proof. split; [by apply type_lft_morphism_id_like|done| |]. - move=> */=. by do 9 f_equiv. - - move=> */=. do 13 (f_contractive || f_equiv). by simpl in *. + - move=> */=. do 13 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance cell_copy {𝔄} (ty: type 𝔄) : Copy ty → Copy (cell ty). @@ -265,8 +265,7 @@ Section cell. { iApply proph_obs_impl; [|done]=>/= π. by case (vπ π)=> ? _[_ ?]. } iSplit; [|done]. rewrite tctx_hasty_val. iExists _. iFrame "⧖ †". iNext. rewrite split_mt_uniq_bor. iFrame "In". iExists _, _, _. by iFrame. - - iNext. iExists _, _. iFrame "⧖' Pc'". iExists _. iFrame "↦'". - iExists _, _, _. iSplit; [done|]. iFrame "ty ⧖'". + - iNext. iFrame "⧖' Pc' ↦' ty". iExists _. iSplit; [done|]. iApply proph_obs_impl; [|done]=>/= π. by case (vπ π)=>/= ??[? _]. - iIntros "!> (%&%&_&_&%&?&%&%&%&>->&_& ⧖'' &?)". iExists _, _. iFrame "⧖''". iSplitL "Vo Pc"; last first. { iExists _. by iFrame. } @@ -309,9 +308,8 @@ Section cell. - iNext. iExists _, _. iFrame "⧖ Pc'". iExists _. iFrame "↦' ty". iApply proph_obs_impl; [|done]=>/= π. case (vπ π)=>/= ??[->[Imp ?]]. apply Imp=>//. apply (aπ π). - - iIntros "!> (%&%&(#⧖' & Pc' &%& ↦ & Obs' & ty)) !>!>". iExists _, _. - iFrame "⧖'". iSplitL "ToPc". { iApply "ToPc". by iApply proph_eqz_refl. } - iExists _. iFrame "↦". iExists _, _, _. iSplit; [done|]. by iFrame. + - iIntros "!> (%&%&(#⧖' & Pc' &%& ↦ & Obs' & ty)) !>!>". + iFrame "⧖' ↦ ty Obs'". iExists _. iSplitL; [|done]. iApply "ToPc". iApply proph_eqz_refl. Qed. (** Updating the Invariant *) @@ -348,8 +346,7 @@ Section cell. iMod (uniq_resolve _ [] 1%Qp with "PROPH Vo Pc []") as "(Obs'' & Pc &_)"; [done..|]. iCombine "Obs Obs'" as "Obs". iCombine "Obs'' Obs" as "#?". iMod ("Toα" with "[↦ ty Pc]") as "[_ α]". - { iNext. iExists _, _. iFrame "⧖ Pc". iExists _. iFrame "↦". iExists _, _, _. - iSplit; [done|]. iFrame "⧖ ty". iApply proph_obs_impl; [|done]=>/= π. + { iNext. iFrame "⧖ Pc ↦ ty". iExists _. iSplitR; [done|]. iApply proph_obs_impl; [|done]=>/= π. move: (equal_f Eq π)=>/=. case (vπ π)=>/= ??->[_[[Imp _]?]]. by apply Imp. } iMod ("ToL" with "α L") as "L". rewrite cctx_interp_singleton. iApply ("C" $! [# #_] -[_] with "Na L [↦r †r] []"). diff --git a/theories/typing/lib/diverging_static.v b/theories/typing/lib/diverging_static.v deleted file mode 100644 index ccfd10c9..00000000 --- a/theories/typing/lib/diverging_static.v +++ /dev/null @@ -1,55 +0,0 @@ -From iris.proofmode Require Import proofmode. -From lrust.typing Require Export type. -From lrust.typing Require Import typing. -Set Default Proof Using "Type". - -Section diverging_static. - Context `{!typeG Σ}. - - (* Show that we can convert any live borrow to 'static with an infinite - loop. *) - Definition diverging_static_loop (call_once : val) : val := - fn: ["x"; "f"] := - let: "call_once" := call_once in - letcall: "ret" := "call_once" ["f"; "x"]%E in - withcont: "loop": - "loop" [] - cont: "loop" [] := - "loop" []. - - Lemma diverging_static_loop_type T F call_once : - (* F : FnOnce(&'static T), as witnessed by the impl call_once *) - typed_val call_once (fn(∅; F, &shr{static} T) → unit) → - typed_val (diverging_static_loop call_once) - (fn(∀ α, λ ϝ, ty_outlives_E T static; &shr{α} T, F) → ∅). - Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x f. simpl_subst. - iApply type_let; [apply Hf|solve_typing|]. iIntros (call); simpl_subst. - (* Drop to Iris *) - iIntros (tid) "#LFT #TIME #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α & _ & _)"; - [solve_typing..|]. - iMod (lft_eternalize with "Hα") as "#Hα". - iAssert (type_incl (box (&shr{α} T)) (box (&shr{static} T))) as "#[_ [_ [Hincl _]]]". - { iApply box_type_incl. iApply shr_type_incl; first done. - iApply type_incl_refl. } - wp_let. rewrite (tctx_hasty_val _ call). iDestruct "Hcall" as (?) "[_ Hcall]". - iApply (type_call_iris _ [ϝ] () [_; _] with "LFT TIME HE Hna [Hϝ] Hcall [Hx Hf]"). - - solve_typing. - - by rewrite /= (right_id static). - - simpl. iFrame. iSplit; last done. rewrite !tctx_hasty_val. - iDestruct "Hx" as (?) "[??]". iExists _. iFrame. iApply "Hincl". done. - - iClear "Hα Hincl". iIntros (r depth') "Htl Hϝ1 Hdepth Hret". wp_rec. - (* Go back to the type system. *) - iApply (type_type _ [] _ [] with "[] LFT TIME HE Htl [] Hk [-]"); last first. - { rewrite /tctx_interp /=. done. } - { rewrite /llctx_interp /=. done. } - iApply (type_cont [] [] (λ _, [])). - + iIntros (?). simpl_subst. iApply type_jump; solve_typing. - + iIntros "!>" (? args). inv_vec args. simpl_subst. iApply type_jump; solve_typing. - Qed. -End diverging_static. diff --git a/theories/typing/lib/fake_shared.v b/theories/typing/lib/fake_shared.v deleted file mode 100644 index 55c03115..00000000 --- a/theories/typing/lib/fake_shared.v +++ /dev/null @@ -1,66 +0,0 @@ -From iris.proofmode Require Import proofmode. -From lrust.typing Require Import typing. -Set Default Proof Using "Type". - -Section fake_shared. - Context `{!typeG Σ}. - - Definition fake_shared_box : val := - fn: ["x"] := Skip ;; return: ["x"]. - - Lemma fake_shared_box_type ty : - typed_val fake_shared_box - (fn(∀ '(α, β), ∅; &shr{α}(&shr{β} ty)) → &shr{α}(box ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #TIME #HE Hna HL Hk HT". - rewrite tctx_interp_singleton tctx_hasty_val. - iDestruct "HT" as (depth) "[#Hdepth HT]". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. - iAssert (▷ (own_ptr 1 (&shr{α}(box ty))).(ty_own) depth tid [x])%I with "[HT]" as "HT". - { destruct x as [[|l|]|], depth as [|depth]=>//=. iDestruct "HT" as "[H $]". - iNext. iDestruct "H" as ([|[[]|][]]) "[H↦ H]"; try done. - iExists _. iFrame. iDestruct "H" as (vl) "[#Hf H]". - iNext. destruct vl as [|[[|l'|]|][]]; try done. iExists l'. iSplit. - { iApply frac_bor_iff; last done. iIntros "!>!> *". - rewrite heap_mapsto_vec_singleton. iSplit; auto. } - by iApply ty_shr_mono. } - do 2 wp_seq. - iApply (type_type [] _ _ [ x ◁ box (&shr{α}(box ty)) ] - with "[] LFT TIME [] Hna HL Hk [HT]"); last first. - { rewrite tctx_interp_singleton tctx_hasty_val. auto. } - { by rewrite /elctx_interp. } - iApply type_jump; simpl; solve_typing. - Qed. - - Definition fake_shared_uniq : val := - fn: ["x"] := Skip ;; return: ["x"]. - - Lemma fake_shared_uniq_type ty : - 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 #TIME #HE Hna HL Hk HT". - rewrite tctx_interp_singleton tctx_hasty_val. - iDestruct "HT" as (depth) "[#Hdepth HT]". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|]. - (* FIXME: WTF, using &uniq{β} here does not work. *) - iAssert (▷ (own_ptr 1 (&shr{α} (uniq_bor β ty))).(ty_own) depth tid [x])%I with "[HT]" as "HT". - { destruct x as [[|l|]|], depth as [|depth]=>//=. iDestruct "HT" as "[H $]". - iNext. iDestruct "H" as ([|[[]|][]]) "[H↦ H]"; try done. - iExists _. iFrame. iDestruct "H" as (vl) "[#Hf H]". - iNext. destruct vl as [|[[|l'|]|][]]; try done. iExists l'. iSplit. - { iApply frac_bor_iff; last done. iIntros "!>!> *". - rewrite heap_mapsto_vec_singleton. iSplit; auto. } - by iApply ty_shr_mono. } - do 2 wp_seq. - iApply (type_type [] _ _ [ x ◁ box (&shr{α}(&uniq{β} ty)) ] - with "[] LFT TIME [] Hna HL Hk [HT]"); last first. - { rewrite tctx_interp_singleton tctx_hasty_val. auto. } - { by rewrite /elctx_interp. } - iApply type_jump; simpl; solve_typing. - Qed. -End fake_shared. diff --git a/theories/typing/lib/join.v b/theories/typing/lib/join.v deleted file mode 100644 index 321c75a9..00000000 --- a/theories/typing/lib/join.v +++ /dev/null @@ -1,96 +0,0 @@ -From iris.proofmode Require Import proofmode. -From lrust.lang Require Import spawn. -From lrust.typing Require Export type. -From lrust.typing Require Import typing. -Set Default Proof Using "Type". - -Definition joinN := lrustN .@ "join". - -Section join. - Context `{!typeG Σ, !spawnG Σ}. - - (* This model is very far from rayon::join, which uses a persistent - work-stealing thread-pool. Still, the important bits are right: - One of the closures is executed in another thread, and the - closures can refer to on-stack data (no 'static' bound). *) - Definition join (call_once_A call_once_B : val) (R_A R_B : type) : val := - fn: ["fA"; "fB"] := - let: "call_once_A" := call_once_A in - let: "call_once_B" := call_once_B in - let: "join" := spawn [λ: ["c"], - letcall: "r" := "call_once_A" ["fA"]%E in - finish ["c"; "r"]] in - letcall: "retB" := "call_once_B" ["fB"]%E in - let: "retA" := spawn.join ["join"] in - (* Put the results in a pair. *) - let: "ret" := new [ #(R_A.(ty_size) + R_B.(ty_size)) ] in - "ret" +ₗ #0 <-{R_A.(ty_size)} !"retA";; - "ret" +ₗ #R_A.(ty_size) <-{R_B.(ty_size)} !"retB";; - delete [ #R_A.(ty_size); "retA"] ;; delete [ #R_B.(ty_size); "retB"] ;; - return: ["ret"]. - - (* TODO: Find a better spot for this. *) - Lemma Z_nat_add (n1 n2 : nat) : Z.to_nat (n1 + n2) = (n1 + n2)%nat. - Proof. rewrite Z2Nat.inj_add; [|lia..]. rewrite !Nat2Z.id //. Qed. - - Lemma join_type A B R_A R_B call_once_A call_once_B - `(!Send A, !Send B, !Send R_A, !Send R_B) : - (* A : FnOnce() -> R_A, as witnessed by the impl call_once_A *) - typed_val call_once_A (fn(∅; A) → R_A) → - (* B : FnOnce() -> R_B, as witnessed by the impl call_once_B *) - typed_val call_once_B (fn(∅; B) → R_B) → - typed_val (join call_once_A call_once_B R_A R_B) (fn(∅; A, B) → Π[R_A; R_B]). - Proof using Type*. - intros HfA HfB E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>envA envB. simpl_subst. - iApply type_let; [apply HfA|solve_typing|]. iIntros (fA); simpl_subst. - iApply type_let; [apply HfB|solve_typing|]. iIntros (fB); simpl_subst. - (* Drop to Iris. *) - iIntros (tid) "#LFT #TIME #HE Hna HL Hk (HfB & HfA & HenvA & HenvB & _)". - iMod (lctx_lft_alive_tok ϝ with "HE HL") as (qϝ1) "(Hϝ1 & HL & Hclose1)"; - [solve_typing..|]. - (* FIXME: using wp_apply here breaks calling solve_to_val. *) - wp_bind (spawn _). - iApply ((spawn_spec joinN (λ v, - tctx_elt_interp tid (v ◁ box R_A) ∗ (qϝ1).[ϝ])%I) - with "[HfA HenvA Hϝ1]"). - { (* The new thread. *) - simpl_subst. iIntros (c) "Hfin". iMod na_alloc as (tid') "Htl". wp_let. - wp_let. unlock. rewrite !tctx_hasty_val. iDestruct "HfA" as (?) "[_ HfA]". - iApply (type_call_iris _ [ϝ] () [_] with "LFT TIME HE Htl [Hϝ1] HfA [HenvA]"). - - rewrite /fp_E outlives_product. solve_typing. - - by rewrite /= (right_id static). - - iDestruct "HenvA" as (?) "HenvA". - rewrite big_sepL_singleton tctx_hasty_val send_change_tid. auto. - - iIntros (r depth') "Htl Hϝ1 #Hdepth' Hret". - wp_rec. iApply (finish_spec with "[$Hfin Hret Hϝ1]"); last auto. - rewrite (right_id static). iFrame. rewrite tctx_hasty_val. - iExists _. iFrame "Hdepth'". by iApply @send_change_tid. } - iNext. iIntros (c) "Hjoin". wp_let. wp_let. - iMod (lctx_lft_alive_tok ϝ with "HE HL") as (qϝ2) "(Hϝ2 & HL & Hclose2)"; - [solve_typing..|]. - rewrite !tctx_hasty_val. iDestruct "HfB" as (?) "[_ HfB]". - iApply (type_call_iris _ [ϝ] () [_] with "LFT TIME HE Hna [Hϝ2] HfB [HenvB]"). - { rewrite /fp_E outlives_product. solve_typing. } - { by rewrite /= (right_id static). } - { by rewrite big_sepL_singleton tctx_hasty_val. } - (* The return continuation after calling fB in the main thread. *) - iIntros (retB depth') "Hna Hϝ2 Hdepth' HretB". rewrite /= (right_id static). - iMod ("Hclose2" with "Hϝ2 HL") as "HL". wp_rec. - wp_apply (join_spec with "Hjoin"). iIntros (retA) "[HretA Hϝ1]". - iMod ("Hclose1" with "Hϝ1 HL") as "HL". wp_let. - (* Switch back to type system mode. *) - iApply (type_type _ _ _ [ retA ◁ box R_A; retB ◁ box R_B ] - with "[] LFT TIME HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. auto with iFrame. } - iApply (type_new_subtype (Π[uninit R_A.(ty_size); uninit R_B.(ty_size)])); - (* FIXME: solve_typing should handle this without any aid. *) - rewrite ?Z_nat_add; [solve_typing..|]. - iIntros (r); simpl_subst. - iApply (type_memcpy R_A); [solve_typing..|]. - iApply (type_memcpy R_B); [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. -End join. diff --git a/theories/typing/lib/mutex/mutex.v b/theories/typing/lib/mutex/mutex.v index bbfc0d5b..edf309e3 100644 --- a/theories/typing/lib/mutex/mutex.v +++ b/theories/typing/lib/mutex/mutex.v @@ -114,7 +114,7 @@ Section mutex. split; [by apply type_lft_morphism_id_like|by move=>/= ??->|..]. - move=>/= *. by do 13 f_equiv. - move=>/= *. do 7 f_equiv. { by apply equiv_dist, lft_incl_equiv_proper_r. } - rewrite /mutex_body. do 12 (f_contractive || f_equiv). simpl in *. by apply dist_S. + rewrite /mutex_body. do 12 (f_contractive || f_equiv). eapply dist_later_dist_lt; [|done]. lia. Qed. Global Instance mutex_send {𝔄} (ty: type 𝔄) : Send ty → Send (mutex ty). @@ -180,7 +180,7 @@ Section mutex. iIntros "_". do 3 wp_seq. rewrite cctx_interp_singleton. iApply ("C" $! [# #_] -[const Φ] with "Na L [-] []"); last first. { by iApply proph_obs_impl; [|done]=>/= ?[_ ?]. } - rewrite/= right_id (tctx_hasty_val #_). iExists _. iFrame "⧖". + rewrite/= right_id (tctx_hasty_val #_). iExists _. iSplitR; [by iApply "⧖"|]. rewrite/= freeable_sz_full. iFrame "†m". iNext. rewrite split_mt_mutex. iExists _, _, _, _. iSplit; [done|]. iFrame "↦b ⧖". iSplit; last first. { iExists _. iFrame. } @@ -255,11 +255,9 @@ Section mutex. rewrite/= freeable_sz_full. iFrame "†m". iNext. rewrite split_mt_uniq_bor. iFrame "In". iExists _, _, _. by iFrame. - iNext. iExists _, _. by iFrame. - - iIntros "!> big !>!>". iDestruct "big" as (??) "(⧖' &_& ↦ty)". - iExists _, _. iFrame "⧖". - iSplitL "ToPc". { iApply "ToPc". iApply proph_eqz_refl. } - iExists _, _, _, _. iSplit; [done|]. iFrame "↦b ⧖' ↦ty". - by iApply proph_obs_true. + - iIntros "!> big !>!>". iDestruct "big" as (??) "(⧖' &_& $)". iFrame "⧖ ⧖' ↦b". + iExists _. iSplitL. { iApply "ToPc". iApply proph_eqz_refl. } + iExists _. iSplit; [done|]. by iApply proph_obs_true. Qed. End mutex. diff --git a/theories/typing/lib/mutex/mutexguard.v b/theories/typing/lib/mutex/mutexguard.v index 4903e427..c7a763f4 100644 --- a/theories/typing/lib/mutex/mutexguard.v +++ b/theories/typing/lib/mutex/mutexguard.v @@ -116,9 +116,9 @@ Section mutexguard. split; [by eapply type_lft_morphism_add_one|done| |]. - move=>/= *. do 10 f_equiv. { by apply equiv_dist, lft_incl_equiv_proper_r. } rewrite /mutex_body. - f_equiv; [do 2 f_equiv|]; f_contractive; do 9 f_equiv; by simpl in *. + f_equiv; [do 2 f_equiv|]; f_contractive; do 9 f_equiv; by eapply dist_later_dist_lt. - move=>/= *. do 13 f_equiv. { by apply equiv_dist, lft_incl_equiv_proper_r. } - do 17 (f_contractive || f_equiv). by simpl in *. + do 17 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance mutexguard_sync {𝔄} κ (ty: type 𝔄) : @@ -330,7 +330,7 @@ Section mutexguard. Definition mutexguard_drop: val := fn: ["g"] := - release [!"g"];; delete [ #1; "g"];; + release [(!"g")];; delete [ #1; "g"];; return: [new [ #0]]. (* Rust's MutexGuard::drop *) diff --git a/theories/typing/lib/rc/rc.v b/theories/typing/lib/rc/rc.v deleted file mode 100644 index 15b85b12..00000000 --- a/theories/typing/lib/rc/rc.v +++ /dev/null @@ -1,1162 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree excl numbers. -From lrust.lang.lib Require Import memcpy. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Export type. -From lrust.typing Require Import typing option. -Set Default Proof Using "Type". - -Definition rc_stR := - prodUR (optionUR (csumR (prodR fracR positiveR) (exclR unitO))) natUR. -Class rcG Σ := - RcG :> inG Σ (authR rc_stR). -Definition rcΣ : gFunctors := #[GFunctor (authR rc_stR)]. -Global Instance subG_rcΣ {Σ} : subG rcΣ Σ → rcG Σ. -Proof. solve_inG. Qed. - -Definition rc_tok q : authR rc_stR := - ◯ (Some $ Cinl (q, 1%positive), 0%nat). -Definition rc_dropping_tok : authR rc_stR := - ◯ (Some $ Cinr $ Excl (), 0%nat). -Definition weak_tok : authR rc_stR := ◯ (None, 1%nat). - -Definition rcN := lrustN .@ "rc". -Definition rc_invN := rcN .@ "inv". -Definition rc_shrN := rcN .@ "shr". - -Section rc. - Context `{!typeG Σ, !rcG Σ}. - - (* The RC can be in four different states : - - The living state, meaning that some strong reference exists. The - authoritative state is something like (Some (Cinl (q, strong)), weak), - where q is the total fraction owned by strong references, strong is - the number of strong references and weak is the number of weak - references. - - The "dropping" state, meaning that the last strong reference has been - dropped, and that the content in being dropped. The authoritative - state is something like (Some (Cinr (Excl ())), weak), where weak is - the number of weak references. The client owning the Excl also owns - the content of the box. - In our case, this state is not really necesary, because we do not - properly support dropping the content, but just copy it out of the RC - box. However, including it is more realistic, and this state is - still necessary for Arc anyway. - - The weak state, meaning that there only exists weak references. The - authoritative state is something like (None, weak), where weak is the - number of weak references. - - The dead state, meaning that no reference exist anymore. The - authoritative state is something like (None, 0). - - Note that when we are in the living or dropping states, the weak reference - count stored in the heap is actually one plus the actual number of weak - references. This hack (which also exists in Rust's implementation) makes the - implementation of weak_drop easier, because it does not have to check the - strong count. *) - - Definition rc_inv tid ν (γ : gname) (l : loc) (ty : type) : iProp Σ := - (∃ st : rc_stR, own γ (● st) ∗ - match st with - | (Some (Cinl (q, strong)), weak) => ∃ q', - l ↦ #(Zpos strong) ∗ (l +ₗ 1) ↦ #(S weak) ∗ † l…(2 + ty.(ty_size)) ∗ - ⌜(q + q')%Qp = 1%Qp⌝ ∗ q'.[ν] ∗ - (* We keep this view shift decomposed because we store the persistent - part in ty_own, to make it available with one step less. *) - ([†ν] ={↑lftN}=∗ ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid) - | (Some (Cinr _), weak) => - l ↦ #0 ∗ (l +ₗ 1) ↦ #(S weak) - | (None, (S _) as weak) => - l ↦ #0 ∗ (l +ₗ 1) ↦ #weak ∗ † l…(2 + ty.(ty_size)) ∗ - (l +ₗ 2) ↦∗: λ vl, ⌜length vl = ty.(ty_size)⌝ - | _ => True - end)%I. - - Definition rc_persist tid ν (γ : gname) (l : loc) (ty : type) : iProp Σ := - tc_opaque (∃ ty', ▷ type_incl ty' ty ∗ - na_inv tid rc_invN (rc_inv tid ν γ l ty') ∗ - (* We use this disjunction, and not simply [ty_shr] here, - because [weak_new] cannot prove ty_shr, even for a dead - lifetime. *) - (ty.(ty_shr) (ν ⊓ ty_lft ty) tid (l +ₗ 2) ∨ [†ν]) ∗ - □ (1.[ν] ={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=∗ [†ν]))%I. - - Global Instance rc_persist_persistent tid ν γ l ty : Persistent (rc_persist tid ν γ l ty). - Proof. unfold rc_persist, tc_opaque. apply _. Qed. - - Lemma rc_persist_type_incl tid ν γ l ty1 ty2: - type_incl ty1 ty2 -∗ rc_persist tid ν γ l ty1 -∗ rc_persist tid ν γ l ty2. - Proof. - iIntros "#Hincl H". iDestruct "H" as (ty) "#(?&?& Hs &?)". iExists ty. - iFrame "#". iSplit. - - iNext. by iApply type_incl_trans. - - iDestruct "Hs" as "[?|?]"; last auto. - iLeft. iDestruct "Hincl" as "(_&Hlft&_&Hincls)". iApply "Hincls". - iApply (ty_shr_mono with "[] [//]"). iApply lft_intersect_mono; [|done]. - iApply lft_incl_refl. - Qed. - - Program Definition rc (ty : type) := - {| ty_size := 1; - ty_lfts := ty.(ty_lfts); ty_E := ty.(ty_E); - ty_own tid vl := - match vl return _ with - | [ #(LitLoc l) ] => - (* The ty_own part of the rc type cointains either the - shared protocol or the full ownership of the - content. The reason is that get_mut does not have the - masks to rebuild the invariants. *) - (l ↦ #1 ∗ (l +ₗ 1) ↦ #1 ∗ † l…(2 + ty.(ty_size)) ∗ - ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid) ∨ - (∃ γ ν q, rc_persist tid ν γ l ty ∗ own γ (rc_tok q) ∗ q.[ν]) - | _ => False end; - ty_shr κ tid l := - ∃ (l' : loc), &frac{κ} (λ q, l ↦{q} #l') ∗ - □ ∀ F q, ⌜↑shrN ∪ ↑lftN ⊆ F⌝ -∗ q.[κ] - ={F}[F∖↑shrN]▷=∗ q.[κ] ∗ ∃ γ ν q', κ ⊑ ν ∗ - rc_persist tid ν γ l' ty ∗ - &na{κ, tid, rc_shrN}(own γ (rc_tok q')) - |}%I. - Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". Qed. - Next Obligation. - iIntros (ty E κ l tid q ?) "#LFT #Hout Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done. - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done. - destruct vl as [|[[|l'|]|][|]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - setoid_rewrite heap_mapsto_vec_singleton. - iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx. - iDestruct "Hb" as (i) "(#Hpb&Hpbown)". - set (C := (∃ _ _ _, _ ∗ _ ∗ &na{_,_,_} _)%I). - iMod (inv_alloc shrN _ (idx_bor_own 1 i ∨ C)%I - with "[Hpbown]") as "#Hinv"; first by iLeft. - iIntros "!> !>" (F q') "% [Htok1 Htok2]". - iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. - iDestruct "INV" as "[>Hbtok|#Hshr]". - - iAssert (&{κ} _)%I with "[Hbtok]" as "Hb". - { rewrite bor_unfold_idx. iExists _. by iFrame. } - iClear "H↦ Hinv Hpb". - iMod (bor_acc_cons with "LFT Hb Htok1") as "[HP Hclose2]"; first solve_ndisj. - set (X := (∃ _ _ _, _)%I). iModIntro. iNext. - iAssert (X ∗ (q'/2).[κ])%I with "[>HP Htok2]" as "[HX $]". - { iDestruct "HP" as "[(Hl'1 & Hl'2 & H† & Hown)|$]"; last done. - iMod (lft_create with "LFT") as (ν) "[[Hν1 Hν2] #Hν†]"; first solve_ndisj. - (* TODO: We should consider changing the statement of - bor_create to dis-entangle the two masks such that this is no - longer necessary. *) - iApply (fupd_mask_mono (↑lftN)); first solve_ndisj. - iMod (bor_create with "LFT Hown") as "[HP HPend]"; first solve_ndisj. - iMod (own_alloc (● (Some $ Cinl ((1/2)%Qp, 1%positive), 0%nat) ⋅ - ◯ (Some $ Cinl ((1/2)%Qp, 1%positive), 0%nat))) as (γ) "[Ha Hf]". - { by apply auth_both_valid_discrete. } - iMod (na_inv_alloc tid _ _ (rc_inv tid ν γ l' ty) - with "[Ha Hν2 Hl'1 Hl'2 H† HPend]") as "#?". - { rewrite /rc_inv. iExists (Some $ Cinl (_, _), _). iFrame. iExists _. - iFrame "#∗". rewrite Qp_div_2; auto. } - iMod (lft_incl_acc with "Hout Htok2") as (q'') "[Htok Hclose]"; [done|]. - iDestruct (lft_intersect_acc with "Hν1 Htok") as (q''') "[Htok Hclose']". - iMod (ty_share with "LFT [] [HP] Htok") as "[? Htok]"; first solve_ndisj. - { iApply lft_intersect_incl_r. } - { iApply (bor_shorten with "[] HP" ). iApply lft_intersect_incl_l. } - iDestruct ("Hclose'" with "Htok") as "[? Htok]". - iMod ("Hclose" with "Htok") as "$". - iExists _, _, _. iFrame. iExists ty. iFrame "#". iSplitR; last by auto. - by iApply type_incl_refl. } - iDestruct "HX" as (γ ν q'') "(#Hpersist & Hrctok)". - iMod ("Hclose2" with "[] Hrctok") as "[HX $]". - { unfold X. iIntros "!> [??] !>". iNext. iRight. do 3 iExists _. - iFrame "#∗". } - iAssert C with "[>HX]" as "#$". - { iExists _, _, _. iFrame "Hpersist". - iMod (bor_sep with "LFT HX") as "[Hrc Hlft]"; first solve_ndisj. - iDestruct (frac_bor_lft_incl with "LFT [> Hlft]") as "$". - { iApply (bor_fracture with "LFT"); first solve_ndisj. by rewrite Qp_mul_1_r. } - iApply (bor_na with "Hrc"); solve_ndisj. } - iApply ("Hclose1" with "[]"). by auto. - - iMod ("Hclose1" with "[]") as "_"; first by auto. - iApply step_fupd_intro; first solve_ndisj. by iFrame. - Qed. - Next Obligation. - iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]". - iExists _. iSplit; first by iApply frac_bor_shorten. - iModIntro. iIntros (F q) "% Htok". - iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj. - iMod ("Hshr" with "[] Htok") as "Hshr2"; first done. - iModIntro. iNext. iMod "Hshr2" as "[Htok HX]". - iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν ?) "(? & ? & ?)". - iExists _, _, _. iModIntro. iFrame. iSplit. - - by iApply lft_incl_trans. - - by iApply na_bor_shorten. - Qed. - - Global Instance rc_type_contractive : TypeContractive rc. - split. - - apply (type_lft_morphism_add _ static [] []). - + intros. rewrite left_id. iApply lft_equiv_refl. - + intros. by rewrite /= /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid vl. destruct vl as [|[[|l|]|] [|]]=>//=. - rewrite /rc_persist /type_incl Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ α l, ty1.(ty_shr) (α ⊓ ty_lft ty1) tid l ≡{n}≡ - ty2.(ty_shr) (α ⊓ ty_lft ty2) tid l) as Hs'. - { intros. rewrite Hs. apply equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - simpl. repeat (apply Ho || apply dist_S, Hs || apply Hs' || - apply equiv_dist, lft_incl_equiv_proper_l, Hl || - f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. rewrite /= /rc /rc_persist /type_incl Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ l α, dist_later n (ty1.(ty_shr) (α ⊓ ty_lft ty1) tid l) - (ty2.(ty_shr) (α ⊓ ty_lft ty2) tid l)) as Hs'. - { intros. rewrite Hs. apply dist_dist_later, equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - simpl. repeat (apply Ho || apply dist_S, Hs || apply Hs' || - apply equiv_dist, lft_incl_equiv_proper_l, Hl || - f_contractive || f_equiv). - Qed. - - Global Instance rc_ne : NonExpansive rc. - unfold rc, rc_persist, type_incl. - intros n x y Hxy. constructor; [done|by apply Hxy..| |]. - - intros. rewrite ![X in X {| ty_own := _ |}]/ty_own /=. - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - Qed. - - Lemma rc_subtype ty1 ty2 : - type_incl ty1 ty2 -∗ type_incl (rc ty1) (rc ty2). - Proof. - iIntros "#Hincl". iPoseProof "Hincl" as "(#Hsz & #Hl & #Hoincl & #Hsincl)". - iSplit; [done|]. iSplit; [done|]. iSplit; iModIntro. - - iIntros "* Hvl". destruct vl as [|[[|vl|]|] [|]]; try done. - iDestruct "Hvl" as "[(Hl1 & Hl2 & H† & Hc) | Hvl]". - { iLeft. iFrame. iDestruct "Hsz" as %->. - iFrame. iApply (heap_mapsto_pred_wand with "Hc"). iApply "Hoincl". } - 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". - iModIntro. iNext. iMod "H" as "[$ H]". - iDestruct "H" as (γ ν q') "(Hlft & Hpersist & Hna)". - iExists _, _, _. iFrame. by iApply rc_persist_type_incl. - Qed. - - Global Instance rc_mono E L : - Proper (subtype E L ==> subtype E L) rc. - Proof. - iIntros (ty1 ty2 Hsub ?) "HL". iDestruct (Hsub with "HL") as "#Hsub". - iIntros "!> #HE". iApply rc_subtype. by iApply "Hsub". - Qed. - Global Instance rc_proper E L : - Proper (eqtype E L ==> eqtype E L) rc. - Proof. intros ??[]. by split; apply rc_mono. Qed. -End rc. - -Section code. - Context `{!typeG Σ, !rcG Σ}. - - Lemma rc_check_unique ty F tid (l : loc) : - ↑rc_invN ⊆ F → - {{{ na_own tid F ∗ ty_own (rc ty) tid [ #l ] }}} - !#l - {{{ strong, RET #(Zpos strong); l ↦ #(Zpos strong) ∗ - (((⌜strong = 1%positive⌝ ∗ - (∃ weak : Z, (l +ₗ 1) ↦ #weak ∗ - ((⌜weak = 1⌝ ∗ - |={⊤}[↑lft_userN]▷=> † l…(2 + ty.(ty_size)) ∗ - ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid ∗ na_own tid F) ∨ - (⌜weak > 1⌝ ∗ - ((l ↦ #1 -∗ (l +ₗ 1) ↦ #weak - ={⊤}=∗ na_own tid F ∗ ty_own (rc ty) tid [ #l ]) ∧ - (l ↦ #0 -∗ (l +ₗ 1) ↦ #(weak - 1) - ={⊤}[↑lft_userN]▷=∗ ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid ∗ - ((l +ₗ 2) ↦∗: (λ vl, ⌜length vl = ty.(ty_size)⌝) - ={⊤}=∗ na_own tid F)))))) ∧ - (l ↦ #0 ={⊤}[↑lft_userN]▷=∗ - ▷ (l +ₗ 2) ↦∗: ty.(ty_own) tid ∗ † l…(2 + ty.(ty_size)) ∗ na_own tid F ∗ - (na_own tid F ={⊤}=∗ ∃ weak : Z, - (l +ₗ 1) ↦ #weak ∗ - ((⌜weak = 1⌝ ∗ l ↦ #0 ∗ na_own tid F) ∨ - (⌜weak > 1⌝ ∗ - († l…(2 + ty.(ty_size)) -∗ (l +ₗ 1) ↦ #(weak-1) -∗ - (l +ₗ 2) ↦∗: (λ vl, ⌜length vl = ty.(ty_size)⌝) - ={⊤}=∗ na_own tid F)))))) - ) ∨ - (⌜(1 < strong)%positive⌝ ∗ - ((l ↦ #(Zpos strong) ={⊤}=∗ na_own tid F ∗ ty_own (rc ty) tid [ #l ]) ∧ - (l ↦ #(Zpos strong - 1) ={⊤}=∗ na_own tid F)))) - }}}. - Proof. - iIntros (? Φ) "[Hna [(Hl1 & Hl2 & H† & Hown)|Hown]] HΦ". - - wp_read. iApply "HΦ". iFrame "Hl1". iLeft. iSplit. done. iSplit. - + iExists _. iFrame "Hl2". iLeft. iFrame. iSplit; first done. - iApply step_fupd_intro; auto. - + iIntros "Hl1". iFrame. iApply step_fupd_intro; first done. - auto 10 with iFrame. - - iDestruct "Hown" as (γ ν q) "(#Hpersist & Htok & Hν1)". - iPoseProof "Hpersist" as (ty') "(Hincl & Hinv & _ & #Hνend)". - iMod (na_inv_acc with "Hinv Hna") as "(Hproto & Hna & Hclose)"; [solve_ndisj..|]. - iDestruct "Hproto" as ([st weak]) "[>Hst Hproto]". - iDestruct (own_valid_2 with "Hst Htok") as %[[[[=]|(?&st'&[=<-]&EQst'&Hincl)] - %option_included _]%prod_included [Hval _]]%auth_both_valid_discrete. - simpl in EQst'. subst st. destruct Hincl as [Hincl|Hincl]. - + destruct st' as [[]| |]; try by inversion Hincl. apply (inj Cinl) in Hincl. - apply (inj2 pair) in Hincl. destruct Hincl as [<-%leibniz_equiv <-%leibniz_equiv]. - iDestruct "Hproto" as (q') "(Hl1 & Hl2 & Hl† & >Hqq' & Hν & Hν†)". - iDestruct "Hqq'" as %Hqq'. iPoseProof "Hincl" as "(>Hincls & _ & Hinclo & _)". - iDestruct "Hincls" as %Hincls. - wp_read. iApply "HΦ". iFrame "Hl1". iLeft. iSplit; first done. iSplit. - * iExists _. iFrame "Hl2". destruct weak. - -- iLeft. iSplit; first done. rewrite Hincls. iFrame "Hl†". - iMod (own_update_2 with "Hst Htok") as "Hst". - { apply auth_update_dealloc. rewrite -{1}(right_id (None, 0%nat) _ (_, _)). - apply cancel_local_update_unit, _. } - rewrite -[in (1).[ν]%I]Hqq' -[(|={↑lft_userN,⊤}=>_)%I]fupd_trans. - iApply step_fupd_mask_mono; - last iMod ("Hνend" with "[$Hν $Hν1]") as "H†"; try done. - iModIntro. iNext. iMod "H†". - iMod fupd_intro_mask' as "Hclose2"; last iMod ("Hν†" with "H†") as "Hty". - { set_solver-. } - iMod "Hclose2" as "_". iModIntro. - iMod ("Hclose" with "[Hst $Hna]") as "$"; first by iExists _; iFrame. - iModIntro. iNext. iDestruct "Hty" as (vl) "[??]". iExists _. iFrame. - by iApply "Hinclo". - -- iRight. iSplit; first by auto with lia. iSplit. - ++ iIntros "Hl1 Hl2". - iMod ("Hclose" with "[-Htok Hν1]") as "$"; last by auto 10 with iFrame. - iFrame. iExists _. auto with iFrame. - ++ iIntros "Hl1 Hl2". - rewrite -[in (1).[ν]%I]Hqq' -[(|={↑lft_userN,⊤}=>_)%I]fupd_trans. - iApply step_fupd_mask_mono; - last iMod ("Hνend" with "[$Hν $Hν1]") as "H†"; try done. - iModIntro. iNext. iMod "H†". - iMod fupd_intro_mask' as "Hclose2"; last iMod ("Hν†" with "H†") as "Hty". - { set_solver-. } - iMod "Hclose2" as "_". iModIntro. - iMod (own_update_2 with "Hst Htok") as "Hst". - { apply auth_update_dealloc, prod_local_update_1, - (cancel_local_update_unit (Some _) None). } - iSplitL "Hty". - { iDestruct "Hty" as (vy) "[H Hty]". iExists vy. iFrame. - by iApply "Hinclo". } - iIntros "!> H". iApply ("Hclose" with "[>-]"). iFrame. iExists _. - iFrame. rewrite Hincls /= !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. - by iFrame. - * iIntros "Hl1". - iMod (own_update_2 with "Hst Htok") as "[Hst Htok]". - { apply auth_update. etrans. - - rewrite [(Some _, weak)]pair_split. apply cancel_local_update_unit, _. - - apply (op_local_update _ _ (Some (Cinr (Excl tt)), 0%nat)). auto. } - rewrite -[(|={↑lft_userN,⊤}=>_)%I]fupd_trans -[in (1).[ν]%I]Hqq'. - iApply step_fupd_mask_mono; - last iMod ("Hνend" with "[$Hν $Hν1]") as "H†"; try done. - iModIntro. iNext. iMod "H†". - iMod fupd_intro_mask' as "Hclose2"; last iMod ("Hν†" with "H†") as "Hty". - { set_solver-. } - iMod "Hclose2" as "_". iModIntro. - iMod ("Hclose" with "[Hst $Hna Hl1 Hl2]") as "$"; - first by iExists _; iFrame; iFrame. - rewrite Hincls. iFrame. iSplitL "Hty". - { iDestruct "Hty" as (vl) "[??]". iExists _. iFrame. by iApply "Hinclo". } - iIntros "!> Hna". iClear "Hνend". clear q' Hqq' weak Hval. - iMod (na_inv_acc with "Hinv Hna") as "(Hproto & Hna & Hclose)"; [solve_ndisj..|]. - iDestruct "Hproto" as ([st weak]) "[>Hst Hproto]". - iDestruct (own_valid_2 with "Hst Htok") as %[[[[=]|(?&st'&[=<-]&EQst'&Hincl)] - %option_included _]%prod_included [Hval _]]%auth_both_valid_discrete. - simpl in EQst'. subst st. destruct Hincl as [Hincl|Hincl]; first last. - { apply csum_included in Hincl. destruct Hincl as - [->|[(?&?&[=]&?)|(?&?&[=<-]&->&?%exclusive_included)]]; try done. apply _. } - setoid_subst. iDestruct "Hproto" as ">(Hl1 & Hl2)". iExists _. iFrame. - rewrite right_id. destruct weak as [|weak]. - -- iLeft. iFrame. iSplitR; first done. - iApply ("Hclose" with "[>- $Hna]"). iExists (None, 0%nat). - iMod (own_update_2 with "Hst Htok") as "$"; last done. - apply auth_update_dealloc. rewrite -{1}(right_id (None, 0%nat) _ (_, _)). - apply cancel_local_update_unit, _. - -- iRight. iSplitR; first by auto with lia. iIntros "!> Hl† Hl2 Hvl". - iApply ("Hclose" with "[>- $Hna]"). iExists (None, S weak). - rewrite Hincls. iFrame. iSplitR "Hl2"; last first. - { by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. } - iMod (own_update_2 with "Hst Htok") as "$"; last done. - apply auth_update_dealloc, prod_local_update', reflexivity. - rewrite -{1}(right_id None _ (Some _)). apply: cancel_local_update_unit. - + apply csum_included in Hincl. - destruct Hincl as [->|[(?&(?,?)&[=<-]&->&(q',strong)&[]%(inj2 pair))| - (?&?&[=]&?)]]; first done. setoid_subst. - iDestruct "Hproto" as (q'') "(Hl1 & Hl2 & Hl† & >Hqq' & Hν & Hν†)". - iDestruct "Hqq'" as %Hqq'. wp_read. iApply "HΦ". iFrame "Hl1". iRight. - iSplit; first by rewrite !pos_op_plus; auto with lia. iSplit; iIntros "H↦". - * iMod ("Hclose" with "[- Htok Hν1]") as "$"; last by auto 10 with iFrame. - iFrame. iExists _. iFrame. auto with iFrame. - * iMod (own_update_2 with "Hst Htok") as "Hst". - { apply auth_update_dealloc. - rewrite pair_op Cinl_op Some_op -{1}(left_id 0%nat _ weak) pair_op. - apply (cancel_local_update_unit _ (_, _)). } - iApply "Hclose". iFrame. iExists _. iFrame. iExists (q+q'')%Qp. iFrame. - iSplitL; first last. - { rewrite [(_+_)%Qp]assoc [(q'+_)%Qp]comm. auto. } - rewrite pos_op_plus Z.sub_1_r -Pos2Z.inj_pred; last lia. - by rewrite Pos.add_1_l Pos.pred_succ. - Qed. - - Definition rc_strong_count : val := - fn: ["rc"] := - let: "r" := new [ #1 ] in - let: "rc'" := !"rc" in - let: "rc''" := !"rc'" in - let: "strong" := !("rc''" +ₗ #0) in - "r" <- "strong";; - delete [ #1; "rc" ];; return: ["r"]. - - Lemma rc_strong_count_type ty : - typed_val rc_strong_count (fn(∀ α, ∅; &shr{α}(rc ty)) → int). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)". - iModIntro. wp_let. wp_op. rewrite shift_loc_0. - (* Finally, finally... opening the thread-local Rc protocol. *) - iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)". - iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|]. - iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weak]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hrctok") as %[[[[=]|(?&[[q0 s0]| |]&[=<-]&?&Hincl)] - %option_included _]%prod_included [Hval _]]%auth_both_valid_discrete; - setoid_subst; try done; last first. - { exfalso. destruct Hincl as [Hincl|Hincl]. by inversion Hincl. - apply csum_included in Hincl. naive_solver. } - iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >% & Hνtok & Hν†)". - wp_read. wp_let. - (* And closing it again. *) - iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hνtok Hν† $Hna]") as "Hna". - { iExists _. iFrame "Hrc●". iExists _. auto with iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(rc ty)); r ◁ box (uninit 1); - #(Z.pos s0) ◁ int ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. - rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_weak_count : val := - fn: ["rc"] := - let: "r" := new [ #1 ] in - let: "rc'" := !"rc" in - let: "rc''" := !"rc'" in - let: "weak" := !("rc''" +ₗ #1) in - let: "one" := #1 in - let: "weak" := "weak" - "one" in - "r" <- "weak";; - delete [ #1; "rc" ];; return: ["r"]. - - Lemma rc_weak_count_type ty : - typed_val rc_weak_count (fn(∀ α, ∅; &shr{α}(rc ty)) → int). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock [[r]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)". - iModIntro. wp_let. wp_op. - (* Finally, finally... opening the thread-local Rc protocol. *) - iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)". - iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|]. - iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weak]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hrctok") as %[[[[=]|(?&[[q0 weak0]| |]&[=<-]&?&Hincl)] - %option_included _]%prod_included [Hval _]]%auth_both_valid_discrete; - setoid_subst; try done; last first. - { exfalso. destruct Hincl as [Hincl|Hincl]. by inversion Hincl. - apply csum_included in Hincl. naive_solver. } - iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >% & Hνtok & Hν†)". - wp_read. wp_let. - (* And closing it again. *) - iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hνtok Hν† $Hna]") as "Hna". - { iExists _. iFrame "Hrc●". iExists _. auto with iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(rc ty)); r ◁ box (uninit 1); - #(S weak) ◁ int ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. - rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - iFrame. } - iApply type_int. iIntros (?). simpl_subst. - iApply type_minus; [solve_typing..|]. iIntros (?). simpl_subst. - iApply type_assign; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_new ty : val := - fn: ["x"] := - let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in - let: "rc" := new [ #1 ] in - "rcbox" +ₗ #0 <- #1;; - "rcbox" +ₗ #1 <- #1;; - "rcbox" +ₗ #2 <-{ty.(ty_size)} !"x";; - "rc" <- "rcbox";; - delete [ #ty.(ty_size); "x"];; return: ["rc"]. - - Lemma rc_new_type ty : - typed_val (rc_new ty) (fn(∅; ty) → rc ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (rc); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]". - rewrite !tctx_hasty_val. - iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x. - iDestruct (ownptr_uninit_own with "Hrcbox") - as (lrcbox vlrcbox) "(% & Hrcbox↦ & Hrcbox†)". subst rcbox. inv_vec vlrcbox=>???. - iDestruct (heap_mapsto_vec_cons with "Hrcbox↦") as "[Hrcbox↦0 Hrcbox↦1]". - iDestruct (heap_mapsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]". - rewrite shift_loc_assoc. - iDestruct (ownptr_uninit_own with "Hrc") as (lrc vlrc) "(% & Hrc↦ & Hrc†)". subst rc. - inv_vec vlrc=>?. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write. - wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz. - wp_apply (wp_memcpy with "[$Hrcbox↦2 $Hx↦]"); [by auto using vec_to_list_length..|]. - iIntros "[Hrcbox↦2 Hx↦]". wp_seq. wp_write. - iApply (type_type _ _ _ [ #lx ◁ box (uninit (ty_size ty)); #lrc ◁ box (rc ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=. iFrame. - iSplitL "Hx↦". - { iExists _. rewrite uninit_own. auto. } - iNext. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. iLeft. - rewrite freeable_sz_full_S. iFrame. iExists _. iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_clone : val := - fn: ["rc"] := - let: "r" := new [ #1 ] in - let: "rc'" := !"rc" in - let: "rc''" := !"rc'" in - let: "strong" := !("rc''" +ₗ #0) in - "rc''" +ₗ #0 <- "strong" + #1;; - "r" <- "rc''";; - delete [ #1; "rc" ];; return: ["r"]. - - Lemma rc_clone_type ty : - typed_val rc_clone (fn(∀ α, ∅; &shr{α}(rc ty)) → rc ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)". - subst r. inv_vec vlr=>r. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)". - iModIntro. wp_let. wp_op. rewrite shift_loc_0. - (* Finally, finally... opening the thread-local Rc protocol. *) - iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)". - iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|]. - iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weak]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hrctok") as %[[[[=]|(?&[[q0 s0]| |]&[=<-]&?&Hincl)] - %option_included _]%prod_included [Hval _]]%auth_both_valid_discrete; - setoid_subst; try done; last first. - { exfalso. destruct Hincl as [Hincl|Hincl]. by inversion Hincl. - apply csum_included in Hincl. naive_solver. } - iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >% & Hνtok & Hν†)". - wp_read. wp_let. wp_op. rewrite shift_loc_0. wp_op. wp_write. wp_write. - (* And closing it again. *) - iMod (own_update with "Hrc●") as "[Hrc● Hrctok2]". - { apply auth_update_alloc, prod_local_update_1, - (op_local_update_discrete _ _ (Some (Cinl ((qb/2)%Qp, 1%positive))))=>-[/= Hqa _]. - split; simpl; last done. apply frac_valid'. rewrite /= -H comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. } - rewrite right_id -Some_op -Cinl_op -pair_op. iDestruct "Hνtok" as "[Hνtok1 Hνtok2]". - iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hνtok2 Hν† $Hna]") as "Hna". - { iExists _. iFrame "Hrc●". iExists _. rewrite Z.add_comm. iFrame. - rewrite [_ ⋅ _]comm frac_op' -[(_ + _)%Qp]assoc Qp_div_2. auto. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(rc ty)); #lr ◁ box (rc ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame "Hx". iFrame "Hr†". iExists [# #l']. - rewrite heap_mapsto_vec_singleton /=. simpl. eauto 10 with iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_deref : val := - fn: ["rc"] := - let: "x" := new [ #1 ] in - let: "rc'" := !"rc" in - "x" <- (!"rc'" +ₗ #2);; - delete [ #1; "rc" ];; return: ["x"]. - - Lemma rc_deref_type ty : - typed_val rc_deref (fn(∀ α, ∅; &shr{α}(rc ty)) → &shr{α}ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (x); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#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. - destruct rc' as [[|rc'|]|]; try done. simpl of_val. - iDestruct "Hrc'" as (l') "[Hrc' #Hdelay]". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hdelay" with "[] Hα1"); last iApply (wp_step_fupd with "Hdelay"); [done..|]. - iMod (frac_bor_acc with "LFT Hrc' Hα2") as (q') "[Hrc'↦ Hclose2]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose2" with "[$Hrc'↦]") as "Hα2". iIntros "!> [Hα1 #Hproto] !>". - iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & _)". - iDestruct "Hpersist" as (ty') "(_ & _ & [Hshr|Hν†] & _)"; last first. - { iMod (lft_incl_dead with "Hαν Hν†") as "Hα†". done. - iDestruct (lft_tok_dead with "Hα1 Hα†") as "[]". } - wp_op. wp_write. iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ rcx ◁ box (&shr{α}(rc ty)); #lrc2 ◁ box (&shr{α}ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame "Hrcx". iFrame "Hx†". iExists [_]. rewrite heap_mapsto_vec_singleton. - iFrame "Hx". iApply (ty_shr_mono with "[] Hshr"). - iApply lft_incl_glb; [done|]. iApply elctx_interp_ty_outlives_E. - rewrite !elctx_interp_app /=. iDestruct "HE" as "(_ & [[_ $] _] & _)". } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_try_unwrap ty : val := - fn: ["rc"] := - let: "r" := new [ #(Σ[ ty; rc ty ]).(ty_size) ] in - withcont: "k": - let: "rc'" := !"rc" in - let: "strong" := !("rc'" +ₗ #0) in - if: "strong" = #1 then - (* Decrement strong *) - "rc'" +ₗ #0 <- #0;; - Skip;; - (* Return content *) - "r" <-{ty.(ty_size),Σ 0} !("rc'" +ₗ #2);; - (* Decrement weak (removing the one weak ref collectively held by all - strong refs), and deallocate if weak count reached 0. *) - let: "weak" := !("rc'" +ₗ #1) in - if: "weak" = #1 then - delete [ #(2 + ty.(ty_size)); "rc'" ];; - "k" [] - else - "rc'" +ₗ #1 <- "weak" - #1;; - "k" [] - else - "r" <-{Σ 1} "rc'";; - "k" [] - cont: "k" [] := - delete [ #1; "rc"];; return: ["r"]. - - Lemma rc_try_unwrap_type ty : - typed_val (rc_try_unwrap ty) (fn(∅; rc ty) → Σ[ ty; rc ty ]). - Proof. - (* TODO: There is a *lot* of duplication between this proof and the one for drop. *) - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new (Σ[ ty; rc ty ]).(ty_size)); [solve_typing..|]; iIntros (r); simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] - (λ _, [rcx ◁ box (uninit 1); r ◁ box (Σ[ ty; rc ty ])])) ; - [solve_typing..| |]; last first. - { simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iIntros (k). simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. - destruct rc' as [[|rc'|]|]; try done. - rewrite [[ #rc' ]]lock. wp_op. rewrite shift_loc_0 -(lock [ #rc' ]). - wp_apply (rc_check_unique with "[$Hna $Hrc']"); first solve_ndisj. - iIntros (strong) "[Hrc Hc]". wp_let. - iDestruct "Hc" as "[[% [_ Hproto]]|[% [Hproto _]]]". - - subst strong. wp_op. wp_if. wp_op. - rewrite shift_loc_0. wp_write. wp_bind (#☠;;#☠)%E. - iApply (wp_step_fupd with "[Hproto Hrc]"); - [|by iApply ("Hproto" with "Hrc")|]; - [done..|wp_seq; iIntros "(Hty&H†&Hna&Hproto) !>"]. - rewrite <-freeable_sz_full_S, <-(freeable_sz_split _ 2 ty.(ty_size)). - iDestruct "H†" as "[H†1 H†2]". wp_seq. wp_bind (_<-_;;_)%E. - iApply (wp_wand with "[Hna HL Hty Hr H†2]"). - { iApply (type_sum_memcpy_instr 0 [_; (rc ty)] with "LFT HE Hna HL"); first done. - { by eapply (write_own _ (uninit _)). } { apply read_own_move. } - iSplitL "Hr"; first by unlock; rewrite tctx_hasty_val. iSplitL; last done. - rewrite tctx_hasty_val'; last done. iFrame. } - iIntros (?) "(Hna & HL & [Hr [Hrc _]])". wp_seq. - iMod ("Hproto" with "Hna") as (weak) "[H↦weak Hproto]". wp_op. wp_read. wp_let. - iDestruct "Hproto" as "[(% & H↦strong & Hna)|[% Hproto]]". - + subst. wp_op. wp_if. - iApply (type_type _ _ _ - [ r ◁ own_ptr (ty_size Σ[ ty; rc ty ]) (Σ[ ty; rc ty]); - rcx ◁ box (uninit 1); - #rc' +ₗ #2 ◁ own_ptr (2 + ty.(ty_size)) (uninit (ty.(ty_size))); - #rc' +ₗ #0 ◁ own_ptr (2 + ty.(ty_size)) (uninit 2)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. rewrite 3!tctx_interp_cons tctx_interp_singleton. iFrame "Hr Hrc". - rewrite 1!tctx_hasty_val tctx_hasty_val' //. iFrame "Hrcx". - rewrite shift_loc_0 /=. iFrame. iExists [_; _]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - auto with iFrame. } - iApply (type_delete (Π[uninit 2;uninit ty.(ty_size)])); [solve_typing..|]. - iApply type_jump; solve_typing. - + rewrite (tctx_hasty_val' _ (#rc' +ₗ #2)); last done. - iDestruct "Hrc" as "[Hrc H†2]". wp_op; case_bool_decide; first lia. wp_if. wp_op. wp_op. wp_write. - iMod ("Hproto" with "[H†1 H†2] H↦weak Hrc") as "Hna". - { rewrite -freeable_sz_full_S -(freeable_sz_split _ 2 ty.(ty_size)). iFrame. } - iApply (type_type _ _ _ - [ r ◁ own_ptr (ty_size Σ[ ty; rc ty ]) (Σ[ ty; rc ty]); - rcx ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. } - iApply type_jump; solve_typing. - - wp_op; case_bool_decide as Hc; first lia. wp_if. iMod ("Hproto" with "Hrc") as "[Hna Hrc]". - iApply (type_type _ _ _ [ r ◁ own_ptr (ty_size Σ[ ty; rc ty ]) (uninit _); - rcx ◁ box (uninit 1); #rc' ◁ rc ty ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton. - rewrite !tctx_hasty_val tctx_hasty_val' //. unlock. iFrame. } - iApply (type_sum_assign Σ[ ty; rc ty ]); [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_drop ty : val := - fn: ["rc"] := - let: "r" := new [ #(option ty).(ty_size) ] in - withcont: "k": - let: "rc'" := !"rc" in - let: "strong" := !("rc'" +ₗ #0) in - "rc'" +ₗ #0 <- "strong" - #1;; - if: "strong" = #1 then - (* Return content. *) - "r" <-{ty.(ty_size),Σ some} !("rc'" +ₗ #2);; - (* Decrement weak (removing the one weak ref collectively held by all - strong refs), and deallocate if weak count reached 0. *) - let: "weak" := !("rc'" +ₗ #1) in - if: "weak" = #1 then - delete [ #(2 + ty.(ty_size)); "rc'" ];; - "k" [] - else - "rc'" +ₗ #1 <- "weak" - #1;; - "k" [] - else - "r" <-{Σ none} ();; - "k" [] - cont: "k" [] := - delete [ #1; "rc"];; return: ["r"]. - - Lemma rc_drop_type ty : - typed_val (rc_drop ty) (fn(∅; rc ty) → option ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new (option ty).(ty_size)); [solve_typing..|]; iIntros (r); simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] - (λ _, [rcx ◁ box (uninit 1); r ◁ box (option ty)])); - [solve_typing..| |]; last first. - { simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iIntros (k). simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. - destruct rc' as [[|rc'|]|]; try done. rewrite [[ #rc' ]]lock. - wp_op. rewrite shift_loc_0. rewrite -(lock [ #rc' ]). - wp_apply (rc_check_unique with "[$Hna $Hrc']"); first solve_ndisj. - iIntros (strong) "[Hrc Hc]". wp_let. wp_op. rewrite shift_loc_0. - rewrite {3}[Z.pos strong]lock. wp_op. unlock. wp_write. - iDestruct "Hc" as "[[% [_ Hproto]]|[% [_ Hproto]]]". - - subst strong. wp_bind (#1 = #1)%E. - iApply (wp_step_fupd with "[Hproto Hrc]"); - [|by iApply ("Hproto" with "Hrc")|]; - [done..|wp_op; iIntros "(Hty&H†&Hna&Hproto) !>"; wp_if]. - rewrite <-freeable_sz_full_S, <-(freeable_sz_split _ 2 ty.(ty_size)). - iDestruct "H†" as "[H†1 H†2]". wp_bind (_<-_;;_)%E. - iApply (wp_wand with "[Hna HL Hty Hr H†2]"). - { iApply (type_sum_memcpy_instr 1 [unit; _] with "LFT HE Hna HL"); first done. - { by eapply (write_own _ (uninit _)). } { apply read_own_move. } - iSplitL "Hr"; first by unlock; rewrite tctx_hasty_val. iSplitL; last done. - rewrite tctx_hasty_val'; last done. iFrame. } - iIntros (?) "(Hna & HL & [Hr [Hrc _]])". wp_seq. - iMod ("Hproto" with "Hna") as (weak) "[H↦weak Hproto]". wp_op. wp_read. wp_let. - iDestruct "Hproto" as "[(% & H↦strong & Hna)|[% Hproto]]". - + subst. wp_op. wp_if. - iApply (type_type _ _ _ - [ r ◁ own_ptr (ty_size (option ty)) (option ty); - rcx ◁ box (uninit 1); - #rc' +ₗ #2 ◁ own_ptr (2 + ty.(ty_size)) (uninit (ty.(ty_size))); - #rc' +ₗ #0 ◁ own_ptr (2 + ty.(ty_size)) (uninit 2)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. rewrite 3!tctx_interp_cons tctx_interp_singleton. iFrame "Hr Hrc". - rewrite 1!tctx_hasty_val tctx_hasty_val' //. iFrame "Hrcx". - rewrite shift_loc_0 /=. iFrame. iExists [_; _]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - auto with iFrame. } - iApply (type_delete (Π[uninit 2;uninit ty.(ty_size)])); [solve_typing..|]. - iApply type_jump; solve_typing. - + rewrite (tctx_hasty_val' _ (#rc' +ₗ #2)); last done. - iDestruct "Hrc" as "[Hrc H†2]". wp_op. case_bool_decide; first lia. wp_if. wp_op. wp_op. wp_write. - iMod ("Hproto" with "[H†1 H†2] H↦weak Hrc") as "Hna". - { rewrite -freeable_sz_full_S -(freeable_sz_split _ 2 ty.(ty_size)). iFrame. } - iApply (type_type _ _ _ - [ r ◁ own_ptr (ty_size (option ty)) (option ty); rcx ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. } - iApply type_jump; solve_typing. - - wp_op; case_bool_decide as Hc; first lia. wp_if. iMod ("Hproto" with "Hrc") as "Hna". - iApply (type_type _ _ _ [ r ◁ own_ptr (ty_size (option ty)) (uninit _); - rcx ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton. - rewrite !tctx_hasty_val. iFrame. } - iApply (type_sum_unit (option ty)); [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_get_mut : val := - fn: ["rc"] := - let: "r" := new [ #2 ] in - withcont: "k": - let: "rc'" := !"rc" in - let: "rc''" := !"rc'" in - let: "strong" := !("rc''" +ₗ #0) in - if: "strong" = #1 then - let: "weak" := !("rc''" +ₗ #1) in - if: "weak" = #1 then - "r" <-{Σ some} ("rc''" +ₗ #2);; - "k" [] - else - "r" <-{Σ none} ();; - "k" [] - else - "r" <-{Σ none} ();; - "k" [] - cont: "k" [] := - delete [ #1; "rc"];; return: ["r"]. - - Lemma rc_get_mut_type ty : - typed_val rc_get_mut (fn(∀ α, ∅; &uniq{α}(rc ty)) → option (&uniq{α}ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] - (λ _, [rcx ◁ box (uninit 1); r ◁ box (option $ &uniq{α}ty)])); - [solve_typing..| |]; last first. - { simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iIntros (k). simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. - iDestruct "Hrc'" as "[#? Hrc']". destruct rc' as [[|rc'|]|]; try done. - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "(Hα & HL & Hclose1)"; [solve_typing..|]. - iMod (bor_acc_cons with "LFT Hrc' Hα") as "[Hrc' Hclose2]"; first solve_ndisj. - iDestruct (heap_mapsto_ty_own with "Hrc'") as (vl) "[>Hrc' Hrcown]". - inv_vec vl=>l. rewrite heap_mapsto_vec_singleton. - wp_read. destruct l as [[|l|]|]; try done. - wp_let. wp_op. rewrite shift_loc_0. - wp_apply (rc_check_unique with "[$Hna Hrcown]"); first done. - { (* Boy this is silly... *) iDestruct "Hrcown" as "[(?&?&?&?)|?]"; last by iRight. - iLeft. by iFrame. } - iIntros (c) "(Hl1 & Hc)". wp_let. wp_op; case_bool_decide as Hc. - - wp_if. iDestruct "Hc" as "[[% [Hc _]]|[% _]]"; last lia. subst. - iDestruct "Hc" as (weak) "[Hl2 Hweak]". wp_op. wp_read. wp_let. - iDestruct "Hweak" as "[[% Hrc]|[% [Hrc _]]]". - + subst. wp_bind (#1 = #1)%E. iApply (wp_step_fupd with "Hrc"); [done..|]. - wp_op. iIntros "(Hl† & Hty & Hna)!>". wp_if. - iMod ("Hclose2" with "[Hrc' Hl1 Hl2 Hl†] [Hty]") as "[Hty Hα]"; [|iNext; iExact "Hty"|]. - { iIntros "!> Hty". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame "Hrc'". - iLeft. by iFrame. } - iMod ("Hclose1" with "Hα HL") as "HL". - iApply (type_type _ _ _ [ r ◁ box (uninit 2); #l +ₗ #2 ◁ &uniq{α}ty; - rcx ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame "∗#". } - iApply (type_sum_assign (option (&uniq{_}_))); [solve_typing..|]. - iApply type_jump; solve_typing. - + wp_op; case_bool_decide; first lia. wp_if. iMod ("Hrc" with "Hl1 Hl2") as "[Hna Hrc]". - iMod ("Hclose2" with "[] [Hrc' Hrc]") as "[Hrc' Hα]". - { iIntros "!> HX". iModIntro. iExact "HX". } - { iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. auto. } - iMod ("Hclose1" with "Hα HL") as "HL". - iApply (type_type _ _ _ [ r ◁ box (uninit 2); rcx ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - unlock. iFrame. } - iApply (type_sum_unit (option (&uniq{_}_))); [solve_typing..|]. - iApply type_jump; solve_typing. - - wp_if. iDestruct "Hc" as "[[% _]|[% [Hproto _]]]"; first lia. - iMod ("Hproto" with "Hl1") as "[Hna Hrc]". - iMod ("Hclose2" with "[] [Hrc' Hrc]") as "[Hrc' Hα]". - { iIntros "!> HX". iModIntro. iExact "HX". } - { iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. auto. } - iMod ("Hclose1" with "Hα HL") as "HL". - iApply (type_type _ _ _ [ r ◁ box (uninit 2); rcx ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - unlock. iFrame. } - iApply (type_sum_unit (option (&uniq{_}_))); [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* TODO : it is not nice that we have to inline the definition of - rc_new and of rc_deref. *) - Definition rc_make_mut (ty : type) (clone : val) : val := - fn: ["rc"] := - let: "r" := new [ #1 ] in - withcont: "k": - let: "rc'" := !"rc" in - let: "rc''" := !"rc'" in - let: "strong" := !("rc''" +ₗ #0) in - if: "strong" = #1 then - let: "weak" := !("rc''" +ₗ #1) in - if: "weak" = #1 then - (* This is the last strong ref, and there is no weak ref. - We just return a deep ptr into the Rc. *) - "r" <- "rc''" +ₗ #2;; - "k" [] - else - (* This is the last strong ref, but there are weak refs. - We make ourselves a new Rc, move the content, and mark the old one killed - (strong count becomes 0, strong idx removed from weak cnt). - We store the new Rc in our argument (which is a &uniq rc), - and return a deep ptr into it. *) - "rc''" +ₗ #0 <- #0;; - "rc''" +ₗ #1 <- "weak" - #1;; - (* Inlined rc_new("rc''" +ₗ #2) begins. *) - let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in - "rcbox" +ₗ #0 <- #1;; - "rcbox" +ₗ #1 <- #1;; - "rcbox" +ₗ #2 <-{ty.(ty_size)} !"rc''" +ₗ #2;; - "rc'" <- "rcbox";; - (* Inlined rc_new ends. *) - "r" <- "rcbox" +ₗ #2;; - "k" [] - else - (* There are other strong refs, we have to make a copy and clone the content. *) - let: "x" := new [ #1 ] in - "x" <- "rc''" +ₗ #2;; - let: "clone" := clone in - letcall: "c" := "clone" ["x"]%E in (* FIXME : why do I need %E here ? *) - (* Inlined rc_new("c") begins. *) - let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in - "rcbox" +ₗ #0 <- #1;; - "rcbox" +ₗ #1 <- #1;; - "rcbox" +ₗ #2 <-{ty.(ty_size)} !"c";; - delete [ #ty.(ty_size); "c"];; - "rc'" <- "rcbox";; - (* Inlined rc_new ends. *) - letalloc: "rcold" <- "rc''" in - (* FIXME : here, we are dropping the old rc pointer. In the - case another strong reference has been dropped while - cloning, it is possible that we are actually dropping the - last reference here. This means that we may have to drop an - instance of [ty] here. Instead, we simply forget it. *) - let: "drop" := rc_drop ty in - letcall: "content" := "drop" ["rcold"]%E in - delete [ #(option ty).(ty_size); "content"];; - "r" <- "rcbox" +ₗ #2;; - "k" [] - cont: "k" [] := - delete [ #1; "rc"];; return: ["r"]. - - Lemma rc_make_mut_type ty clone : - (* ty : Clone, as witnessed by the impl clone *) - typed_val clone (fn(∀ α, ∅; &shr{α}ty) → ty) → - typed_val (rc_make_mut ty clone) (fn(∀ α, ∅; &uniq{α}(rc ty)) → &uniq{α}ty). - Proof. - intros Hclone E L. iApply type_fn; [solve_typing..|]. - rewrite [(2 + ty_size ty)%nat]lock. - iIntros "/= !>". iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] - (λ _, [rcx ◁ box (uninit 1); r ◁ box (&uniq{α}ty)])); - [solve_typing..| |]; last first. - { simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iIntros (k). simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[rcx]]lock [[r]]lock. - iDestruct "Hrc'" as "[#? Hrc']". destruct rc' as [[|rc'|]|]; try done. - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; - [solve_typing..|]. - iMod (bor_acc_cons with "LFT Hrc' Hα1") as "[Hrc' Hclose2]"; first solve_ndisj. - iDestruct (heap_mapsto_ty_own with "Hrc'") as (vl) "[>Hrc' Hrcown]". - inv_vec vl=>l. rewrite heap_mapsto_vec_singleton. - wp_read. destruct l as [[|l|]|]; try done. wp_let. wp_op. rewrite shift_loc_0. - wp_apply (rc_check_unique with "[$Hna Hrcown]"); first done. - { (* Boy this is silly... *) iDestruct "Hrcown" as "[(?&?&?&?)|?]"; last by iRight. - iLeft. by iFrame. } - iIntros (c) "(Hl1 & Hc)". wp_let. wp_op; case_bool_decide as Hc; wp_if. - - iDestruct "Hc" as "[[% [Hc _]]|[% _]]"; last lia. subst. - iDestruct "Hc" as (weak) "[Hl2 Hweak]". wp_op. wp_read. wp_let. - iDestruct "Hweak" as "[[% Hrc]|[% [_ Hrc]]]". - + subst. wp_bind (#1 = #1)%E. iApply (wp_step_fupd with "Hrc"); [done..|]. - wp_op. iIntros "(Hl† & Hty & Hna)!>". wp_if. - iMod ("Hclose2" with "[Hrc' Hl1 Hl2 Hl†] [Hty]") as "[Hty Hα1]"; [|iNext; iExact "Hty"|]. - { iIntros "!> Hty". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame "Hrc'". - iLeft. by iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ r ◁ box (uninit 1); #l +ₗ #2 ◁ &uniq{α}ty; - rcx ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. unlock. iFrame "∗#". } - iApply type_assign; [solve_typing..|]. - iApply type_jump; solve_typing. - + wp_op; case_bool_decide; first lia. wp_if. wp_op. rewrite shift_loc_0. wp_write. wp_op. - wp_op. wp_write. wp_bind (new _). iSpecialize ("Hrc" with "Hl1 Hl2"). - iApply (wp_step_fupd with "Hrc"); [done..|]. iApply wp_new; first lia. done. - rewrite -!lock Nat2Z.id. - iNext. iIntros (lr) "/= ([H†|%] & H↦lr) [Hty Hproto] !>"; last lia. - rewrite 2!heap_mapsto_vec_cons. iDestruct "H↦lr" as "(Hlr1 & Hlr2 & Hlr3)". - wp_let. wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write. wp_op. wp_op. - iDestruct "Hty" as (vlr) "[H↦vlr Hty]". rewrite shift_loc_assoc. - iDestruct (ty_size_eq with "Hty") as %Hsz. - wp_apply (wp_memcpy with "[$Hlr3 $H↦vlr]"). - { by rewrite repeat_length. } { by rewrite Hsz. } - iIntros "[Hlr3 Hvlr]". wp_seq. wp_write. wp_op. - iMod ("Hproto" with "[Hvlr]") as "Hna"; first by eauto. - iMod ("Hclose2" $! ((lr +ₗ 2) ↦∗: ty_own ty tid)%I - with "[Hrc' Hlr1 Hlr2 H†] [Hlr3 Hty]") as "[Hb Hα1]". - { iIntros "!> H !>". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iLeft. iFrame. } - { iExists _. iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ r ◁ box (uninit 1); #(lr +ₗ 2) ◁ &uniq{α}ty; - rcx ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. unlock. iFrame "∗#". } - iApply type_assign; [solve_typing..|]. - iApply type_jump; solve_typing. - - wp_apply wp_new; [lia|done|]. - iIntros (lr) "/= ([H†|%] & H↦lr)"; last lia. - iDestruct "Hc" as "[[% ?]|[% [Hproto _]]]"; first lia. - iMod ("Hproto" with "Hl1") as "[Hna Hty]". wp_let. wp_op. - rewrite heap_mapsto_vec_singleton. wp_write. - iAssert (∃ γ ν q', rc_persist tid ν γ l ty ∗ own γ (rc_tok q') ∗ q'.[ν] ∗ - (q / 2).[α])%I - with "[>Hty Hα2]" as (γ ν q') "(Hty & Htok & Hν & Hα2)". - { iDestruct "Hty" as "[(Hl1 & Hl2 & Hl† & Hl3)|Hty]"; last by iFrame. - iMod (own_alloc (● (Some $ Cinl ((1/2)%Qp, 1%positive), 0%nat) ⋅ - rc_tok (1/2)%Qp)) as (γ) "[Ha Hf]"; - first by apply auth_both_valid_discrete. - iMod (lft_create with "LFT") as (ν) "[[Hν1 Hν2] #Hν†]"=>//. - iApply (fupd_mask_mono (↑lftN))=>//. iExists γ, ν, (1/2)%Qp. iFrame "Hν2 Hf". - iMod (bor_create _ ν with "LFT Hl3") as "[Hb Hh]"=>//. - iMod (lft_incl_acc with "[//] Hα2") as (q') "[Htok Hclose]"; [done|]. - iDestruct (lft_intersect_acc with "Hν1 Htok") as (q'') "[Htok Hclose']". - iMod (ty_share with "LFT [] [Hb] Htok") as "[Hty Htok]"=>//. - { iApply lft_intersect_incl_r. } - { iApply (bor_shorten with "[] Hb"). iApply lft_intersect_incl_l. } - iDestruct ("Hclose'" with "Htok") as "[? Htok]". - iMod ("Hclose" with "Htok") as "$". - iExists ty. iSplitR; [iApply type_incl_refl|]. - iSplitR "Hty"; last by auto. iApply na_inv_alloc. iExists _. do 2 iFrame. - iExists _. iFrame. by rewrite Qp_div_2. } - iDestruct "Hty" as (ty') "#(Hty'ty & Hinv & Hs & Hν†)". - iDestruct "Hs" as "[Hs|Hν']"; last by iDestruct (lft_tok_dead with "Hν Hν'") as "[]". - wp_bind (of_val clone). iApply (wp_wand with "[Hna]"). - { iApply (Hclone _ [] with "LFT HE Hna"). - by rewrite /llctx_interp. by rewrite /tctx_interp. } - clear clone Hclone. iIntros (clone) "(Hna & _ & Hclone)". - wp_let. wp_let. rewrite tctx_interp_singleton tctx_hasty_val. - iDestruct (lft_intersect_acc with "Hν Hα2") as (q'') "[Hαν Hclose3]". - rewrite -[ν ⊓ α](right_id_L). - iApply (type_call_iris _ [ν ⊓ α] (ν ⊓ α) [_] - with "LFT HE Hna Hαν Hclone [H† H↦lr]"); [solve_typing| |]. - { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full_S. - iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iApply ty_shr_mono; [|done]. iApply lft_intersect_mono; [|done]. - iApply lft_incl_refl. } - iIntros ([[|cl|]|]) "Hna Hαν Hcl //". wp_rec. - iDestruct "Hcl" as "[Hcl Hcl†]". iDestruct "Hcl" as (vl) "[Hcl↦ Hown]". - iDestruct (ty_size_eq with "Hown") as %Hsz. - iDestruct ("Hclose3" with "Hαν") as "[Hν Hα2]". - wp_apply wp_new=>//. lia. iIntros (l') "(Hl'† & Hl')". wp_let. wp_op. - rewrite shift_loc_0. rewrite -!lock Nat2Z.id. - rewrite !heap_mapsto_vec_cons shift_loc_assoc. - iDestruct "Hl'" as "(Hl' & Hl'1 & Hl'2)". wp_write. wp_op. wp_write. wp_op. - wp_apply (wp_memcpy with "[$Hl'2 $Hcl↦]"). - { by rewrite repeat_length. } { by rewrite Hsz. } - iIntros "[Hl'2 Hcl↦]". wp_seq. rewrite freeable_sz_full. - wp_apply (wp_delete with "[$Hcl↦ Hcl†]"); - [lia|by replace (length vl) with (ty.(ty_size))|]. - iIntros "_". wp_seq. wp_write. - iMod ("Hclose2" $! ((l' +ₗ 2) ↦∗: ty_own ty tid)%I with - "[Hrc' Hl' Hl'1 Hl'†] [Hl'2 Hown]") as "[Hl' Hα1]". - { iIntros "!> H". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iLeft. iFrame. iDestruct "Hl'†" as "[?|%]"=>//. } - { iExists _. iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - iApply (type_type _ _ _ [ #l ◁ rc ty; #l' +ₗ #2 ◁ &uniq{α}ty; - r ◁ box (uninit 1); rcx ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 3!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - !tctx_hasty_val' //. unlock. iFrame "∗#". iRight. iExists γ, ν, _. - unfold rc_persist, tc_opaque. iFrame "∗#". eauto. } - iApply type_letalloc_1; [solve_typing..|]. iIntros (rcold). simpl_subst. - iApply type_let. apply rc_drop_type. solve_typing. iIntros (drop). simpl_subst. - iApply (type_letcall ()); [solve_typing..|]. iIntros (content). simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_assign; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. -End code. diff --git a/theories/typing/lib/rc/weak.v b/theories/typing/lib/rc/weak.v deleted file mode 100644 index b800f535..00000000 --- a/theories/typing/lib/rc/weak.v +++ /dev/null @@ -1,504 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree numbers. -From lrust.lang.lib Require Import memcpy. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Export type. -From lrust.typing Require Import typing option. -From lrust.typing.lib Require Export rc. -Set Default Proof Using "Type". - -Section weak. - Context `{!typeG Σ, !rcG Σ}. - - Program Definition weak (ty : type) := - {| ty_size := 1; ty_lfts := ty.(ty_lfts); ty_E := ty.(ty_E); - ty_own tid vl := - match vl return _ with - | [ #(LitLoc l) ] => ∃ γ ν, rc_persist tid ν γ l ty ∗ own γ weak_tok - | _ => False end; - ty_shr κ tid l := - ∃ (l' : loc), &frac{κ} (λ q, l ↦{q} #l') ∗ - □ ∀ F q, ⌜↑shrN ∪ ↑lftN ⊆ F⌝ -∗ q.[κ] - ={F}[F∖↑shrN]▷=∗ q.[κ] ∗ ∃ γ ν, rc_persist tid ν γ l' ty ∗ - &na{κ, tid, rc_shrN}(own γ weak_tok) - |}%I. - Next Obligation. by iIntros (ty tid [|[[]|][]]) "H". Qed. - Next Obligation. - iIntros (ty E κ l tid q ?) "#LFT #? Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb"; first done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]"; first done. - (* Ideally, we'd change ty_shr to say "l ↦{q} #l" in the fractured borrow, - but that would be additional work here... *) - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦"; first done. - destruct vl as [|[[|l'|]|][|]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - setoid_rewrite heap_mapsto_vec_singleton. - iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx. - iDestruct "Hb" as (i) "(#Hpb&Hpbown)". - set (C := (∃ _ _, _ ∗ &na{_,_,_} _)%I). - iMod (inv_alloc shrN _ (idx_bor_own 1 i ∨ C)%I with "[Hpbown]") as "#Hinv"; - first by iLeft. - iIntros "!> !> * % Htok". - iMod (inv_acc with "Hinv") as "[INV Hclose1]"; first solve_ndisj. - iDestruct "INV" as "[>Hbtok|#Hshr]". - - iAssert (&{κ} _)%I with "[Hbtok]" as "Hb". - { rewrite bor_unfold_idx. iExists _. by iFrame. } - iClear "H↦ Hinv Hpb". - iMod (bor_acc_cons with "LFT Hb Htok") as "[HP Hclose2]"; first solve_ndisj. - iDestruct "HP" as (γ ν) "(#Hpersist & Hweak)". - iModIntro. iNext. iMod ("Hclose2" with "[] Hweak") as "[Hb $]"; first by auto 10. - iAssert C with "[>Hb]" as "#HC". - { iExists _, _. iFrame "Hpersist". iApply (bor_na with "Hb"). solve_ndisj. } - iMod ("Hclose1" with "[]") as "_"; by auto. - - iMod ("Hclose1" with "[]") as "_"; first by auto. - iApply step_fupd_intro; first solve_ndisj. auto. - Qed. - Next Obligation. - iIntros (ty κ κ' tid l) "#Hincl H". iDestruct "H" as (l') "[#Hl #Hshr]". - iExists _. iSplit; first by iApply frac_bor_shorten. - iModIntro. iIntros (F q) "% Htok". - iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose]"; first solve_ndisj. - iMod ("Hshr" with "[] Htok") as "Hshr2"; first done. - iModIntro. iNext. iMod "Hshr2" as "[Htok HX]". - iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν) "(? & ?)". - iExists _, _. iFrame. by iApply na_bor_shorten. - Qed. - - Global Instance weak_type_contractive : TypeContractive weak. - split. - - apply (type_lft_morphism_add _ static [] [])=>?. - + rewrite left_id. iApply lft_equiv_refl. - + by rewrite /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid vl. destruct vl as [|[[|l|]|] [|]]=>//=. - rewrite /rc_persist /type_incl Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ α l, ty1.(ty_shr) (α ⊓ ty_lft ty1) tid l ≡{n}≡ - ty2.(ty_shr) (α ⊓ ty_lft ty2) tid l) as Hs'. - { intros. rewrite Hs. apply equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - simpl. repeat (apply Ho || apply dist_S, Hs || apply Hs' || - apply equiv_dist, lft_incl_equiv_proper_l, Hl || - f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. rewrite /= /weak /rc_persist /type_incl Hsz. - assert (∀ α, ⊢ α ⊓ ty_lft ty1 ≡ₗ α ⊓ ty_lft ty2) as Hl'. - { intros α. iApply lft_intersect_equiv_proper; [|done]. iApply lft_equiv_refl. } - assert (∀ l α, dist_later n (ty1.(ty_shr) (α ⊓ ty_lft ty1) tid (l +ₗ 2)) - (ty2.(ty_shr) (α ⊓ ty_lft ty2) tid (l +ₗ 2))) as Hs'. - { intros. rewrite Hs. apply dist_dist_later, equiv_dist. - by iSplit; iApply ty_shr_mono; iDestruct Hl' as "[??]". } - simpl. repeat (apply Ho || apply dist_S, Hs || apply Hs' || - apply equiv_dist, lft_incl_equiv_proper_l, Hl || - f_contractive || f_equiv). - Qed. - - Global Instance weak_ne : NonExpansive weak. - unfold weak, rc_persist, type_incl. - intros n x y Hxy. constructor; [done|by apply Hxy..| |]. - - intros. rewrite ![X in X {| ty_own := _ |}]/ty_own /=. - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - - solve_proper_core ltac:(fun _ => eapply ty_size_ne || eapply ty_lfts_ne || f_equiv). - Qed. - - Lemma weak_subtype ty1 ty2 : - type_incl ty1 ty2 -∗ type_incl (weak ty1) (weak ty2). - Proof. - iIntros "#Hincl". iPoseProof "Hincl" as "(Hsz & #Hl & #Hoincl & #Hsincl)". - iSplit; [done|]. iSplit; [done|]. iSplit; iModIntro. - - iIntros "* Hvl". destruct vl as [|[[|vl|]|] [|]]; try done. - 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". - iModIntro. iNext. iMod "H" as "[$ H]". iDestruct "H" as (γ ν) "(Hpersist & Hshr)". - iExists _, _. iFrame "Hshr". by iApply rc_persist_type_incl. - Qed. - - Global Instance weak_mono E L : - Proper (subtype E L ==> subtype E L) weak. - Proof. - iIntros (ty1 ty2 Hsub ?) "HL". iDestruct (Hsub with "HL") as "#Hsub". - iIntros "!> #HE". iApply weak_subtype. iApply "Hsub"; done. - Qed. - Global Instance weak_proper E L : - Proper (eqtype E L ==> eqtype E L) weak. - Proof. intros ??[]. by split; apply weak_mono. Qed. -End weak. - -Section code. - Context `{!typeG Σ, !rcG Σ}. - - Definition rc_upgrade : val := - fn: ["w"] := - let: "r" := new [ #2 ] in - withcont: "k": - let: "w'" := !"w" in - let: "w''" := !"w'" in - let: "strong" := !("w''" +ₗ #0) in - if: "strong" = #0 then - "r" <-{Σ none} ();; - "k" [] - else - "w''" +ₗ #0 <- "strong" + #1;; - "r" <-{Σ some} "w''";; - "k" [] - cont: "k" [] := - delete [ #1; "w" ];; return: ["r"]. - - Lemma rc_upgrade_type ty : - typed_val rc_upgrade (fn(∀ α, ∅; &shr{α}(weak ty)) → option (rc ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>w. simpl_subst. - iApply (type_new 2); [solve_typing..|]; iIntros (r); simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] - (λ _, [w ◁ box (&shr{α}(weak ty)); r ◁ box (option (rc ty))])) ; - [solve_typing..| |]; last first. - { simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iIntros (k). simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (w'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hw [Hw' [Hr _]]]". - rewrite !tctx_hasty_val [[w]]lock. - destruct w' as [[|lw|]|]; try done. iDestruct "Hw'" as (l') "[#Hlw #Hshr]". - iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)". - subst r. inv_vec vlr=>r0 r1. rewrite !heap_mapsto_vec_cons. - iDestruct "Hr" as "(Hr1 & Hr2 & _)". - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlw Hα2") as (q') "[Hlw↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlw↦]") as "Hα2". iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν) "#(Hpersist & Hwtokb)". - iModIntro. wp_let. wp_op. rewrite shift_loc_0. - (* Finally, finally... opening the thread-local Rc protocol. *) - iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)". - iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|]. - iMod (na_bor_acc with "LFT Hwtokb Hα1 Hna") as "(>Hwtok & Hna & Hclose3)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hwtok") as - %[[_ Hweak%nat_included]%prod_included [Hval _]]%auth_both_valid_discrete. - destruct st as [[[q'' strong]| |]|]; try done. - - (* Success case. *) - iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hl'† & >Hq''q0 & [Hν1 Hν2] & Hν†)". - iDestruct "Hq''q0" as %Hq''q0. - wp_read. wp_let. wp_op. wp_if. wp_op. rewrite shift_loc_0. wp_op. wp_write. - (* Closing the invariant. *) - iMod (own_update with "Hrc●") as "[Hrc● Hrctok2]". - { apply auth_update_alloc, prod_local_update_1, - (op_local_update_discrete _ _ (Some (Cinl ((qb/2)%Qp, 1%positive))))=>-[/= Hqa _]. - split; simpl; last done. apply frac_valid'. - rewrite /= -Hq''q0 comm_L. by apply Qp_add_le_mono_l, Qp_div_le. } - rewrite right_id -Some_op -Cinl_op -pair_op. - iMod ("Hclose3" with "[$Hwtok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hl'† Hν2 Hν† $Hna]") as "Hna". - { iExists _. iFrame "Hrc●". iExists _. rewrite Z.add_comm. iFrame. - rewrite [_ ⋅ _]comm frac_op' -[(_ + _)%Qp]assoc Qp_div_2. auto. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ w ◁ box (&shr{α}(weak ty)); #lr ◁ box (uninit 2); - #l' ◁ rc ty ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite 2!tctx_interp_cons tctx_interp_singleton tctx_hasty_val !tctx_hasty_val' //. - unlock. iFrame "Hr† Hw". iSplitL "Hr1 Hr2". - - iExists [_;_]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. auto with iFrame. - - iRight. auto with iFrame. } - iApply (type_sum_assign (option (rc ty))); [solve_typing..|]. - iApply type_jump; solve_typing. - - (* Failure : dropping *) - (* TODO : The two failure cases are almost identical. *) - iDestruct "Hrcst" as "[Hl'1 Hl'2]". wp_read. wp_let. wp_op. wp_if. - (* Closing the invariant. *) - iMod ("Hclose3" with "[$Hwtok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 $Hna]") as "Hna". - { iExists _. auto with iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ w ◁ box (&shr{α}(weak ty)); #lr ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. iExists [_;_]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. auto with iFrame. } - iApply (type_sum_unit (option (rc ty))); [solve_typing..|]. - iApply type_jump; solve_typing. - - (* Failure : general case *) - destruct weakc as [|weakc]; first by simpl in *; lia. - iDestruct "Hrcst" as "[Hl'1 Hrcst]". wp_read. wp_let. wp_op. wp_if. - (* Closing the invariant. *) - iMod ("Hclose3" with "[$Hwtok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hrcst $Hna]") as "Hna". - { iExists _. auto with iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ w ◁ box (&shr{α}(weak ty)); #lr ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. iExists [_;_]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. auto with iFrame. } - iApply (type_sum_unit (option (rc ty))); [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rc_downgrade : val := - fn: ["rc"] := - let: "r" := new [ #1 ] in - let: "rc'" := !"rc" in - let: "rc''" := !"rc'" in - let: "weak" := !("rc''" +ₗ #1) in - "rc''" +ₗ #1 <- "weak" + #1;; - "r" <- "rc''";; - delete [ #1; "rc" ];; return: ["r"]. - - Lemma rc_downgrade_type ty : - typed_val rc_downgrade (fn(∀ α, ∅; &shr{α}(rc ty)) → weak ty). - Proof. - (* TODO : this is almost identical to rc_clone *) - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)". - subst r. inv_vec vlr=>r. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν q'') "#(Hαν & Hpersist & Hrctokb)". - iModIntro. wp_let. wp_op. - (* Finally, finally... opening the thread-local Rc protocol. *) - iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)". - iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|]. - iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose3)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hrctok") as %[[[[=]|(?&[[q0 weak0]| |]&[=<-]&?&Hincl)] - %option_included _]%prod_included [Hval _]]%auth_both_valid_discrete; - setoid_subst; try done; last first. - { exfalso. destruct Hincl as [Hincl|Hincl]. by inversion Hincl. - apply csum_included in Hincl. naive_solver. } - iDestruct "Hrcst" as (qb) "(Hl'1 & Hl'2 & Hrcst)". - wp_read. wp_let. wp_op. wp_op. wp_write. wp_write. - (* And closing it again. *) - iMod (own_update with "Hrc●") as "[Hrc● Hrctok2]". - { by apply auth_update_alloc, prod_local_update_2, (op_local_update_discrete _ _ 1%nat). } - iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]". - iMod ("Hclose2" with "[Hrc● Hl'1 Hl'2 Hrcst $Hna]") as "Hna". - { iExists _. iFrame "Hrc●". iExists _. - rewrite !Nat2Z.inj_succ Z.add_1_r. iFrame. } - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(rc ty)); #lr ◁ box (weak ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame "Hx". iFrame "Hr†". iExists [# #l']. - rewrite heap_mapsto_vec_singleton /=. eauto 10 with iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Exact same code as downgrade *) - Definition weak_clone : val := - fn: ["w"] := - let: "r" := new [ #1 ] in - let: "w'" := !"w" in - let: "w''" := !"w'" in - let: "weak" := !("w''" +ₗ #1) in - "w''" +ₗ #1 <- "weak" + #1;; - "r" <- "w''";; - delete [ #1; "w" ];; return: ["r"]. - - Lemma weak_clone_type ty : - typed_val weak_clone (fn(∀ α, ∅; &shr{α}(weak ty)) → weak ty). - Proof. - (* TODO : this is almost identical to rc_clone *) - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hr _]]]". - rewrite !tctx_hasty_val [[x]]lock. - destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]". - iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr & Hr†)". - subst r. inv_vec vlr=>r. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - iMod (lctx_lft_alive_tok α with "HE HL") as (q) "([Hα1 Hα2] & HL & Hclose1)"; [solve_typing..|]. - wp_bind (!_)%E. - iSpecialize ("Hshr" with "[] Hα1"); last iApply (wp_step_fupd with "Hshr"); [done..|]. - iMod (frac_bor_acc with "LFT Hlrc Hα2") as (q') "[Hlrc↦ Hclose]"; first solve_ndisj. - iApply wp_fupd. wp_read. - iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]". - iDestruct "Hproto" as (γ ν) "#[Hpersist Hwtokb]". - iModIntro. wp_let. wp_op. - (* Finally, finally... opening the thread-local Rc protocol. *) - iPoseProof "Hpersist" as (ty') "(_ & Hinv & _ & _)". - iAssert (∃ wv : Z, (l' +ₗ 1) ↦ #wv ∗ ((l' +ₗ 1) ↦ #(wv + 1) ={⊤}=∗ - na_own tid ⊤ ∗ (q / 2).[α] ∗ own γ weak_tok))%I - with "[> Hna Hα1]" as (wv) "[Hwv Hclose2]". - { iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose2)"; [solve_ndisj..|]. - iMod (na_bor_acc with "LFT Hwtokb Hα1 Hna") as "(>Hwtok & Hna & Hclose3)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hwtok") as - %[[_ Hweak%nat_included]%prod_included [Hval _]]%auth_both_valid_discrete. - iMod (own_update with "Hrc●") as "[Hrc● $]". - { by apply auth_update_alloc, prod_local_update_2, - (op_local_update_discrete _ _ 1%nat). } - destruct st as [[[q'' strong]| |]|]; try done. - - iExists _. iDestruct "Hrcst" as (q0) "(Hl'1 & >$ & Hrcst)". - iIntros "!> Hl'2". iMod ("Hclose3" with "[$Hwtok] Hna") as "[$ Hna]". - iApply ("Hclose2" with "[- $Hna]"). iExists _. iFrame "Hrc●". - iExists _. rewrite /Z.add /= Pos.add_1_r. iFrame. - - iExists _. iDestruct "Hrcst" as "[Hl'1 >$]". iIntros "!> Hl'2". - iMod ("Hclose3" with "[$Hwtok] Hna") as "[$ Hna]". - iApply ("Hclose2" with "[- $Hna]"). iExists _. iFrame "Hrc●". - rewrite /Z.add /= Pos.add_1_r. iFrame. - - destruct weakc as [|weakc]; first by simpl in Hweak; lia. - iExists _. iDestruct "Hrcst" as "(Hl'1 & >$ & Hrcst)". iIntros "!> Hl'2". - iMod ("Hclose3" with "[$Hwtok] Hna") as "[$ Hna]". - iApply ("Hclose2" with "[- $Hna]"). iExists _. iFrame "Hrc●". - rewrite /Z.add /= Pos.add_1_r. iFrame. } - wp_read. wp_let. wp_op. wp_op. wp_write. - iMod ("Hclose2" with "Hwv") as "(Hna & Hα1 & Hwtok)". - iMod ("Hclose1" with "[$Hα1 $Hα2] HL") as "HL". wp_write. - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ box (&shr{α}(weak ty)); #lr ◁ box (weak ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. iExists [# #l']. rewrite heap_mapsto_vec_singleton /=. - auto 10 with iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition weak_drop ty : val := - fn: ["w"] := - withcont: "k": - let: "w'" := !"w" in - let: "weak" := !("w'" +ₗ #1) in - if: "weak" = #1 then - delete [ #(2 + ty.(ty_size)); "w'" ];; - "k" [] - else - "w'" +ₗ #1 <- "weak" - #1;; - "k" [] - cont: "k" [] := - let: "r" := new [ #0] in - delete [ #1; "w"];; return: ["r"]. - - Lemma weak_drop_type ty : - typed_val (weak_drop ty) (fn(∅; weak ty) → unit). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>w. simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] (λ _, [w ◁ box (uninit 1)])); - [solve_typing..| |]; last first. - { simpl. iModIntro. iIntros (k arg). inv_vec arg. simpl_subst. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iIntros (k). simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (w'); simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk [Hw [Hw' _]]". - rewrite !tctx_hasty_val [[w]]lock. destruct w' as [[|lw|]|]; try done. wp_op. - iDestruct "Hw'" as (γ ν) "[#Hpersist Hwtok]". - iAssert (∃ wv : Z, (lw +ₗ 1) ↦ #wv ∗ - ((⌜wv = 1⌝ ∗ na_own tid ⊤ ∗ - ∃ s, lw ↦ s ∗ ((lw +ₗ 2) ↦∗: λ vl, ⌜length vl = ty.(ty_size)⌝) ∗ - † lw … (2 + ty_size ty)) ∨ - (⌜wv > 1⌝ ∗ ((lw +ₗ 1) ↦ #(wv - 1) ={⊤}=∗ na_own tid ⊤))))%I - with "[>Hna Hwtok]" as (wv) "[Hlw [(% & Hna & H)|[% Hclose]]]". - { iPoseProof "Hpersist" as (ty') "([>Hszeq _] & Hinv & _ & _)". - iDestruct "Hszeq" as %Hszeq. - iMod (na_inv_acc with "Hinv Hna") as "(Hrcproto & Hna & Hclose)"; [solve_ndisj..|]. - iDestruct "Hrcproto" as ([st weakc]) "[>Hrc● Hrcst]". - iDestruct (own_valid_2 with "Hrc● Hwtok") as - %[[_ Hweak%nat_included]%prod_included [Hval _]]%auth_both_valid_discrete. - destruct weakc; first by simpl in *; lia. - iMod (own_update_2 with "Hrc● Hwtok") as "Hrc●". - { apply auth_update_dealloc, prod_local_update_2, - (cancel_local_update_unit 1%nat), _. } - destruct st as [[[q'' strong]| |]|]; try done. - - iExists _. iDestruct "Hrcst" as (q0) "(Hlw & >$ & Hrcst)". iRight. - iSplitR; first by auto with lia. iIntros "!>?". iApply "Hclose". iFrame. - iExists _. iFrame. iExists _. iFrame. - by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. - - iExists _. iDestruct "Hrcst" as "[Hlw >$]". iRight. - iSplitR; first by auto with lia. iIntros "!>?". iApply "Hclose". iFrame. - iExists _. iFrame. simpl. - by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. - - iExists _. iDestruct "Hrcst" as "(>Hlw & >$ & >H† & >H)". destruct weakc. - + iLeft. iSplitR; first done. iMod ("Hclose" with "[$Hna Hrc●]") as "$". - { iExists _. iFrame. } - rewrite Hszeq. auto with iFrame. - + iRight. iSplitR; first by auto with lia. iIntros "!>?". iApply "Hclose". - iFrame. iExists _. iFrame. - by rewrite !Nat2Z.inj_succ -!Z.add_1_l Z.add_simpl_l. } - - subst. wp_read. wp_let. wp_op. wp_if. - iApply (type_type _ _ _ [ w ◁ box (uninit 1); #lw ◁ box (uninit (2 + ty.(ty_size))) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. iDestruct "H" as (?) "(? & H↦ & ?)". iDestruct "H↦" as (?) "[? %]". - rewrite /= freeable_sz_full_S. iFrame. iExists (_::_::_). - rewrite 2!heap_mapsto_vec_cons shift_loc_assoc. iFrame. - iIntros "!> !%". simpl. congruence. } - iApply type_delete; [try solve_typing..|]. - iApply type_jump; solve_typing. - - wp_read. wp_let. wp_op; case_bool_decide; first lia. wp_if. wp_op. wp_op. wp_write. - iMod ("Hclose" with "Hlw") as "Hna". - iApply (type_type _ _ _ [ w ◁ box (uninit 1) ] - with "[] LFT HE Hna HL Hk [-]"); last first. - { unlock. by rewrite tctx_interp_singleton tctx_hasty_val. } - iApply type_jump; solve_typing. - Qed. - - Definition weak_new ty : val := - fn: [] := - let: "rcbox" := new [ #(2 + ty.(ty_size))%nat ] in - let: "w" := new [ #1 ] in - "rcbox" +ₗ #0 <- #0;; - "rcbox" +ₗ #1 <- #1;; - "w" <- "rcbox";; - return: ["w"]. - - Lemma weak_new_type ty : - typed_val (weak_new ty) (fn(∅) → weak ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg. simpl_subst. - iApply (type_new (2 + ty.(ty_size))); [solve_typing..|]; iIntros (rcbox); simpl_subst. - iApply (type_new 1); [solve_typing..|]; iIntros (w); simpl_subst. - iIntros (tid) "#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]". - iDestruct (heap_mapsto_vec_cons with "Hrcbox↦1") as "[Hrcbox↦1 Hrcbox↦2]". - iDestruct (ownptr_uninit_own with "Hw") as (lw vlw) "(% & Hw↦ & Hw†)". subst w. - inv_vec vlw=>?. rewrite heap_mapsto_vec_singleton. - (* All right, we are done preparing our context. Let's get going. *) - wp_op. rewrite shift_loc_0. wp_write. wp_op. wp_write. - iMod (lft_create with "LFT") as (ν) "[Hν #Hν†]"; first done. - iPoseProof ("Hν†" with "Hν") as "H†". wp_bind (_ <- _)%E. - iApply wp_mask_mono; last iApply (wp_step_fupd with "H†"); [set_solver-..|]. - wp_write. iIntros "#Hν !>". wp_seq. - iApply (type_type _ _ _ [ #lw ◁ box (weak ty)] - with "[] LFT HE Hna HL Hk [> -]"); last first. - { rewrite tctx_interp_singleton tctx_hasty_val' //=. iFrame. - iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. - iMod (own_alloc (● _ ⋅ ◯ _)) as (γ) "[??]"; last (iExists _, _; iFrame). - { apply auth_both_valid_discrete. by split. } - iExists ty. iFrame "Hν†". iSplitR; first by iApply type_incl_refl. - iSplitL; last by iRight. iMod (na_inv_alloc with "[-]") as "$"; last done. - iExists _. iFrame. rewrite freeable_sz_full_S shift_loc_assoc. iFrame. - iExists _. iFrame. rewrite vec_to_list_length. auto. } - iApply type_jump; solve_typing. - Qed. -End code. diff --git a/theories/typing/lib/refcell/ref.v b/theories/typing/lib/refcell/ref.v deleted file mode 100644 index 220b915c..00000000 --- a/theories/typing/lib/refcell/ref.v +++ /dev/null @@ -1,125 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.refcell Require Import refcell. -Set Default Proof Using "Type". - -Definition refcell_refN := refcellN .@ "ref". - -Section ref. - Context `{!typeG Σ, !refcellG Σ}. - - (* The Rust type looks as follows (after some unfolding): - - pub struct Ref<'b, T: ?Sized + 'b> { - value: &'b T, - borrow: &'b Cell<BorrowFlag>, - } - *) - - Program Definition ref (α : lft) (ty : type) := - tc_opaque - {| ty_size := 2; - ty_lfts := α :: ty.(ty_lfts); ty_E := ty.(ty_E) ++ ty_outlives_E ty α; - ty_own tid vl := - match vl return _ with - | [ #(LitLoc lv); #(LitLoc lrc) ] => - ∃ ν q γ β ty', ty.(ty_shr) (α ⊓ ν) tid lv ∗ - α ⊑ β ∗ &na{β, tid, refcell_invN}(refcell_inv tid lrc γ β ty') ∗ - q.[ν] ∗ own γ (◯ reading_stR q ν) - | _ => False - end; - ty_shr κ tid l := - ∃ ν q γ β ty' (lv lrc : loc), - κ ⊑ ν ∗ &frac{κ} (λ q, l↦∗{q} [ #lv; #lrc]) ∗ - ▷ ty.(ty_shr) (α ⊓ ν) tid lv ∗ - ▷ (α ⊑ β) ∗ ▷ &na{β, tid, refcell_invN}(refcell_inv tid lrc γ β ty') ∗ - &na{κ, tid, refcell_refN}(own γ (◯ reading_stR q ν)) |}%I. - Next Obligation. iIntros (???[|[[]|][|[[]|][]]]) "?"; auto. Qed. - Next Obligation. - iIntros (α ty E κ l tid q ?) "#LFT #Hl Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb". done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]". done. - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦". done. - destruct vl as [|[[|lv|]|][|[[|lrc|]|][]]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - iMod (bor_exists with "LFT Hb") as (ν) "Hb". done. - iMod (bor_exists with "LFT Hb") as (q') "Hb". done. - iMod (bor_exists with "LFT Hb") as (γ) "Hb". done. - iMod (bor_exists with "LFT Hb") as (β) "Hb". done. - iMod (bor_exists with "LFT Hb") as (ty') "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. - iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ Htok]". done. - iMod (bor_sep with "LFT Hb") as "[Hinv Hb]". done. - iMod (bor_persistent with "LFT Hinv Htok") as "[#Hinv $]". done. - iMod (bor_sep with "LFT Hb") as "[Hκν Hb]". done. - iDestruct (frac_bor_lft_incl with "LFT [> Hκν]") as "#Hκν". - { iApply bor_fracture; try done. by rewrite Qp_mul_1_r. } - iMod (bor_na with "Hb") as "#Hb". done. eauto 20. - Qed. - Next Obligation. - iIntros (??????) "#? H". iDestruct "H" as (ν q γ β ty' lv lrc) "H". - iExists _, _, _, _, _, _, _. iDestruct "H" as "#(? & ? & $ & $ & $ & ?)". - iSplit; last iSplit. - - by iApply lft_incl_trans. - - by iApply frac_bor_shorten. - - by iApply na_bor_shorten. - Qed. - - Global Instance ref_type_contractive α : TypeContractive (ref α). - Proof. - split. - - apply (type_lft_morphism_add _ α [α] []) => ?. - + iApply lft_equiv_refl. - + by rewrite elctx_interp_app /= elctx_interp_ty_outlives_E - /= /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid [|[[|lv|]|][|[[|lrc|]|][]]]=>//=. - by setoid_rewrite Hs. - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. simpl. by setoid_rewrite Hs. - Qed. - Global Instance ref_ne α : NonExpansive (ref α). - Proof. solve_ne_type. Qed. - - Global Instance ref_mono E L : - Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) ref. - Proof. - iIntros (α1 α2 Hα ty1 ty2 Hty q) "HL". - iDestruct (Hty with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". - iIntros "!> #HE". iDestruct ("Hα" with "HE") as "Hα1α2". - iDestruct ("Hty" with "HE") as "(%&#Hl&#Ho&#Hs)". - iSplit; [done|iSplit; [|iSplit; iModIntro]]. - - simpl. by iApply lft_intersect_mono. - - 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. - - 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. - Qed. - Global Instance ref_mono_flip E L : - Proper (lctx_lft_incl E L ==> flip (subtype E L) ==> flip (subtype E L)) ref. - Proof. intros ??????. by apply ref_mono. Qed. - Lemma ref_mono' E L α1 α2 ty1 ty2 : - lctx_lft_incl E L α2 α1 → subtype E L ty1 ty2 → - subtype E L (ref α1 ty1) (ref α2 ty2). - Proof. intros. by eapply ref_mono. Qed. - Global Instance ref_proper E L : - Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) ref. - Proof. intros ??[]?? EQ. split; apply ref_mono'; try done; apply EQ. Qed. - Lemma ref_proper' E L α1 α2 ty1 ty2 : - lctx_lft_eq E L α1 α2 → eqtype E L ty1 ty2 → - eqtype E L (ref α1 ty1) (ref α2 ty2). - Proof. intros. by eapply ref_proper. Qed. -End ref. - -Global Hint Resolve refcell_mono' refcell_proper' : lrust_typing. diff --git a/theories/typing/lib/refcell/ref_code.v b/theories/typing/lib/refcell/ref_code.v deleted file mode 100644 index befa9dbc..00000000 --- a/theories/typing/lib/refcell/ref_code.v +++ /dev/null @@ -1,362 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree numbers. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import lifetime na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.refcell Require Import refcell ref. -Set Default Proof Using "Type". - -Section ref_functions. - Context `{!typeG Σ, !refcellG Σ}. - - Lemma refcell_inv_reading_inv tid l γ α ty q ν : - refcell_inv tid l γ α ty -∗ - own γ (◯ reading_stR q ν) -∗ - ∃ (q' : Qp) n, l ↦ #(Zpos n) ∗ ⌜(q ≤ q')%Qp⌝ ∗ - own γ (● (refcell_st_to_R $ Some (ν, false, q', n)) ⋅ ◯ reading_stR q ν) ∗ - ((1).[ν] ={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=∗ &{α}((l +ₗ 1) ↦∗: ty_own ty tid)) ∗ - (∃ q'', ⌜(q' + q'' = 1)%Qp⌝ ∗ q''.[ν]) ∗ - ty.(ty_shr) (α ⊓ ν) tid (l +ₗ 1). - Proof. - iIntros "INV H◯". - iDestruct "INV" as (st) "(H↦lrc & H● & INV)". - iAssert (∃ q' n, st ≡ Some (ν, false, q', n) ∗ ⌜q ≤ q'⌝%Qp)%I with - "[#]" as (q' n) "[%%]". - { destruct st as [[[[??]?]?]|]; - iDestruct (own_valid_2 with "H● H◯") - as %[[[=]|(?&[[? q'] n]&[=<-]&[=]&[[[Eq_ag ?%leibniz_equiv]_]|Hle])] - %option_included Hv]%auth_both_valid_discrete; simpl in *; subst. - - apply (inj to_agree), (inj2 pair) in Eq_ag. - destruct Eq_ag. setoid_subst. eauto. - - revert Hle=> /prod_included [/= /prod_included - [/= /to_agree_included [/= ??] /frac_included_weak ?] _]. - setoid_subst. eauto. } - setoid_subst. iExists q', n. rewrite own_op. by iFrame. - Qed. - - (* Cloning a ref. We need to increment the counter. *) - Definition ref_clone : val := - fn: ["x"] := - let: "x'" := !"x" in - let: "rc" := !("x'" +ₗ #1) in - let: "n" := !"rc" in - "rc" <- "n" + #1;; - letalloc: "r" <-{2} !"x'" in - delete [ #1; "x"];; return: ["r"]. - - Lemma ref_clone_type ty : - typed_val ref_clone (fn(∀ '(α, β), ∅; &shr{α}(ref β ty)) → ref β ty). - - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#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')"; - [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..|]. - iMod (na_bor_acc with "LFT H◯inv Hα2 Hna") as "(H◯ & Hna & Hcloseα2)"; [solve_ndisj..|]. - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iMod "H↦" as "[H↦1 H↦2]". wp_read. wp_let. - iDestruct (refcell_inv_reading_inv with "INV H◯") - as (q' n) "(H↦lrc & _ & [H● H◯] & H† & Hq' & Hshr')". - wp_read. wp_let. wp_op. wp_write. iDestruct "Hq'" as (q'') "(Hq'q'' & Hν1 & Hν2)". - iDestruct "Hq'q''" as %Hq'q''. iMod (own_update with "H●") as "[H● ?]". - { apply auth_update_alloc, - (op_local_update_discrete _ _ (reading_stR (q''/2)%Qp ν))=>-[Hagv _]. - split; [split|done]. - - by rewrite /= agree_idemp. - - apply frac_valid'. rewrite /= -Hq'q'' comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. } - wp_apply wp_new; [done..|]. iIntros (lr) "(?&Hlr)". - iAssert (lx' ↦∗{qlx'} [ #lv; #lrc])%I with "[H↦1 H↦2]" as "H↦". - { rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. iFrame. } - wp_let. wp_apply (wp_memcpy with "[$Hlr $H↦]"); [done..|]. - iIntros "[Hlr H↦]". wp_seq. iMod ("Hcloseα2" with "[$H◯] Hna") as "[Hα1 Hna]". - iMod ("Hcloseδ" with "[H↦lrc H● Hν1 Hshr' H†] Hna") as "[Hδ Hna]". - { iExists (Some (_, false, _, _)). rewrite Z.add_comm -Some_op -!pair_op agree_idemp. - iFrame. iExists _. iFrame. - rewrite (comm Qp_add) (assoc Qp_add) Qp_div_2 (comm Qp_add). auto. } - iMod ("Hcloseβ" with "Hδ") as "Hβ". iMod ("Hcloseα1" with "[$H↦]") as "Hα2". - iMod ("Hclose'" with "[$Hα1 $Hα2] HL") as "HL". iMod ("Hclose" with "Hβ HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(ref β ty)); #lr ◁ box(ref β ty)] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //. - rewrite /= freeable_sz_full. iFrame. iExists _. iFrame. simpl. - iExists _, _, _, _, _. iFrame "∗#". } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Turning a ref into a shared borrow. *) - Definition ref_deref : val := - fn: ["x"] := - let: "x'" := !"x" in - let: "r'" := !"x'" in - letalloc: "r" <- "r'" in - delete [ #1; "x"];; return: ["r"]. - - Lemma ref_deref_type ty : - typed_val ref_deref - (fn(∀ '(α, β), ∅; &shr{α}(ref β ty)) → &shr{α}ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=. - iDestruct "Hx'" as (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hx')". - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. - iMod (frac_bor_acc with "LFT Hfrac Hα") as (qlx') "[H↦ Hcloseα]". done. - 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..|]. - 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. } - iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. - iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Dropping a ref and decrement the count in the corresponding refcell. *) - Definition ref_drop : val := - fn: ["x"] := - let: "rc" := !("x" +ₗ #1) in - let: "n" := !"rc" in - "rc" <- "n" - #1;; - Endlft;; - delete [ #2; "x"];; - let: "r" := new [ #0] in return: ["r"]. - - Lemma ref_drop_type ty : - typed_val ref_drop (fn(∀ α, ∅; ref α ty) → unit). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val. - destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". - iDestruct "Hx" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hx"; simpl; - try iDestruct "Hx" as "[_ >[]]". - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hx" as "[[Hx↦1 Hx↦2] Hx]". wp_op. wp_read. wp_let. - iDestruct "Hx" as (ν q γ β ty') "(_ & #Hαβ & #Hinv & Hν & H◯)". - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. - 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 (refcell_inv_reading_inv with "INV [$H◯]") - as (q' n) "(H↦lrc & >% & H●◯ & H† & Hq' & Hshr)". - iDestruct "Hq'" as (q'') ">[% Hν']". - wp_read. wp_let. wp_op. wp_write. - iAssert (|={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=> refcell_inv tid lrc γ β ty')%I - with "[H↦lrc H●◯ Hν Hν' Hshr H†]" as "INV". - { iDestruct (own_valid with "H●◯") as - %[[[[_ ?]?]|[[_ [q0 Hq0]]%prod_included [n' Hn']]%prod_included] - %Some_included _]%auth_both_valid_discrete. - - simpl in *. setoid_subst. - iExists None. iFrame. iMod (own_update with "H●◯") as "$". - { apply auth_update_dealloc. rewrite -(right_id None op (Some _)). - apply cancel_local_update_unit, _. } - iApply "H†". replace 1%Qp with (q'+q'')%Qp by naive_solver. iFrame. - - simpl in *. setoid_subst. iExists (Some (ν, false, q0, n')). iFrame. - iAssert (lrc ↦ #(Z.pos n'))%I with "[H↦lrc]" as "$". - { rewrite pos_op_plus Z.sub_1_r -Pos2Z.inj_pred; last lia. - by rewrite Pos.add_1_l Pos.pred_succ. } - iMod (own_update with "H●◯") as "$". - { apply auth_update_dealloc. - rewrite -(agree_idemp (to_agree _)) !pair_op Some_op. - apply (cancel_local_update_unit (reading_stR q ν)), _. } - iApply step_fupd_intro; first set_solver-. iExists (q+q'')%Qp. iFrame. - by rewrite assoc (comm _ q0 q). } - wp_bind Endlft. iApply (wp_mask_mono _ (↑lftN ∪ ↑lft_userN)); first done. - iApply (wp_step_fupd with "INV"); [set_solver-..|]. wp_seq. - iIntros "INV !>". wp_seq. iMod ("Hcloseβ" with "[$INV] Hna") as "[Hβ Hna]". - iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ [ #lx ◁ box (uninit 2)] with "[] LFT HE Hna HL Hk"); - first last. - { rewrite tctx_interp_singleton tctx_hasty_val' //. iFrame. iExists [ #lv;#lrc]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton uninit_own. iFrame. auto. } - iApply type_delete; [solve_typing..|]. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_jump; solve_typing. - Qed. - - (* Apply a function within the ref, typically for accessing a component. *) - Definition ref_map (call_once : val) : val := - fn: ["ref"; "f"] := - let: "call_once" := call_once in - let: "x'" := !"ref" in - letalloc: "x" <- "x'" in - letcall: "r" := "call_once" ["f"; "x"]%E in - let: "r'" := !"r" in delete [ #1; "r"];; - "ref" <- "r'";; - return: ["ref"]. - - Lemma ref_map_type ty1 ty2 call_once fty : - (* fty : for<'a>, FnOnce(&'a ty1) -> &'a ty2, as witnessed by the impl call_once *) - typed_val call_once (fn(∀ α, ∅; fty, &shr{α}ty1) → &shr{α}ty2) → - typed_val (ref_map call_once) (fn(∀ α, ∅; ref α ty1, fty) → ref α ty2). - Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α ϝ ret arg). - inv_vec arg=>ref env. simpl_subst. - iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". - rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done. - iDestruct "Href" as "[Href Href†]". - iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl; - try iDestruct "Href" as "[_ >[]]". - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Href" as "[[Href↦1 Href↦2] Href]". - iDestruct "Href" as (ν qν γ β ty') "(#Hshr & #Hαβ & #Hinv & >Hν & Hγ)". - wp_read. wp_let. wp_apply wp_new; try done. - 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..|]. - iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". - iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]". - rewrite -[ϝ in (α ⊓ ν) ⊓ ϝ](right_id_L). - iApply (type_call_iris _ [α ⊓ ν; ϝ] (α ⊓ ν) [_; _] - with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H†]"); [solve_typing|done| |]. - { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. - iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. by iFrame. } - iIntros ([[|r|]|]) "Hna Hανϝ Hr //". - iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]". - iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]". - iMod ("Hclose2" with "Hϝ HL") as "HL". iMod ("Hclose1" with "Hα HL") as "HL". - wp_rec. iDestruct "Hr" as "[Hr Hr†]". - iDestruct "Hr" as ([|r'[]]) "[Hr #Hr']"; - try by iDestruct (ty_size_eq with "Hr'") as "%". - rewrite heap_mapsto_vec_singleton. wp_read. wp_let. - wp_apply (wp_delete _ _ _ [_] with "[Hr Hr†]")=>//. - { rewrite heap_mapsto_vec_singleton freeable_sz_full. iFrame. } iIntros "_". - wp_seq. wp_write. - iApply (type_type _ [_] _ [ #lref ◁ box (ref α ty2) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_singleton tctx_hasty_val' //. simpl. iFrame. - iExists [_;_]. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iFrame. destruct r' as [[]|]=>//=. auto 10 with iFrame. } - iApply type_jump; solve_typing. - Qed. - - (* Apply a function within the ref, and create two ref, - typically for splitting the reference. *) - Definition ref_map_split (call_once : val) : val := - fn: ["ref"; "f"] := - let: "call_once" := call_once in - let: "x'" := !"ref" in - - letalloc: "x" <- "x'" in - letcall: "r" := "call_once" ["f"; "x"]%E in - let: "a" := !"r" in - let: "b" := !("r" +ₗ #1) in - delete [ #2; "r"];; - - let: "rc" := !("ref" +ₗ #1) in - let: "n" := !"rc" in - "rc" <- "n" + #1;; - - delete [ #2; "ref"];; - - let: "refs" := new [ #4 ] in - "refs" <- "a";; - "refs" +ₗ #1 <- "rc";; - "refs" +ₗ #2 <- "b";; - "refs" +ₗ #3 <- "rc";; - - return: ["refs"]. - - Lemma ref_map_split_type ty ty1 ty2 call_once fty : - (* fty : for<'a>, FnOnce(&'a ty) -> (&'a ty1, &'a ty2), - as witnessed by the impl call_once *) - typed_val call_once - (fn(∀ α, ∅; fty, &shr{α}ty) → Π[&shr{α}ty1; &shr{α}ty2]) → - typed_val (ref_map_split call_once) - (fn(∀ α, ∅; ref α ty, fty) → Π[ref α ty1; ref α ty2]). - Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α ϝ ret arg). - inv_vec arg=>ref env. simpl_subst. - iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". - rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done. - iDestruct "Href" as "[Href Href†]". - iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl; - try iDestruct "Href" as "[_ >[]]". - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Href" as "[[Href↦1 Href↦2] Href]". - iDestruct "Href" as (ν qν γ β ty') "(#Hshr & #Hαβ & #Hinv & >Hν & Hγ)". - wp_read. wp_let. wp_apply wp_new; [done..|]. - 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..|]. - iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". - iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]". - rewrite -[ϝ in (α ⊓ ν) ⊓ ϝ](right_id_L). - iApply (type_call_iris _ [α ⊓ ν; ϝ] (α ⊓ ν) [_; _] - with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H†]"); [solve_typing|done| |]. - { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. - iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. by iFrame. } - iIntros ([[|r|]|]) "Hna Hανϝ Hr //". - iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]". - iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]". - iMod ("Hclose2" with "Hϝ HL") as "HL". - wp_rec. iDestruct "Hr" as "[Hr Hr†]". - iDestruct "Hr" as (?) "[Hr H]". iDestruct "H" as (vl1 vl1' ->) "[#Hr1' H]". - iDestruct "H" as (vl2 vl2' ->) "[#Hr2' ->]". - destruct vl1 as [|[[|lr1|]|] []], vl2 as [|[[|lr2|]|] []]=>//=. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hr" as "[Hr1 Hr2]". wp_read. wp_let. wp_op. wp_read. wp_let. - wp_apply (wp_delete _ _ _ [_; _] with "[Hr1 Hr2 Hr†]")=>//. - { rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton /= freeable_sz_full. - iFrame. } - iIntros "_". - iMod (lft_incl_acc with "Hαβ Hα") as (?) "[Hβ Hβclose]". done. - iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hclosena)"; [done..|]. - wp_seq. wp_op. wp_read. - iDestruct (refcell_inv_reading_inv with "INV Hγ") - as (q1 n) "(H↦lrc & _ & [H● H◯] & H† & Hq1 & Hshr')". - iDestruct "Hq1" as (q2) "(Hq1q2 & Hν1 & Hν2)". - iDestruct "Hq1q2" as %Hq1q2. iMod (own_update with "H●") as "[H● ?]". - { apply auth_update_alloc, - (op_local_update_discrete _ _ (reading_stR (q2/2)%Qp ν))=>-[Hagv _]. - split; [split|done]. - - by rewrite /= agree_idemp. - - apply frac_valid'. rewrite /= -Hq1q2 comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. } - wp_let. wp_read. wp_let. wp_op. wp_write. - wp_apply (wp_delete _ _ _ [_; _] with "[Href↦1 Href↦2 Href†]")=>//. - { rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton /= freeable_sz_full. - iFrame. } - iIntros "_". wp_seq. wp_apply wp_new; [done..|]. iIntros (lrefs) "[Hrefs† Hrefs]". - rewrite 3!heap_mapsto_vec_cons heap_mapsto_vec_singleton. wp_let. - iDestruct "Hrefs" as "(Hrefs1 & Hrefs2 & Hrefs3 & Hrefs4)". - rewrite !shift_loc_assoc. wp_write. do 3 (wp_op; wp_write). - iMod ("Hclosena" with "[H↦lrc H● Hν1 Hshr' H†] Hna") as "[Hβ Hna]". - { iExists (Some (_, false, _, _)). rewrite Z.add_comm -Some_op -!pair_op agree_idemp. - iFrame. iExists _. iFrame. - rewrite (comm Qp_add) (assoc Qp_add) Qp_div_2 (comm Qp_add). auto. } - iMod ("Hβclose" with "Hβ") as "Hα". iMod ("Hclose1" with "Hα HL") as "HL". - iApply (type_type _ [_] _ [ #lrefs ◁ box (Π[ref α ty1; ref α ty2]) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_singleton tctx_hasty_val' //= -freeable_sz_full. - simpl. iFrame. iExists [_;_;_;_]. - rewrite 3!heap_mapsto_vec_cons heap_mapsto_vec_singleton !shift_loc_assoc. - iFrame. iExists [_;_], [_;_]. iSplit=>//=. - iSplitL "Hν H◯"; [by auto 10 with iFrame|]. iExists [_;_], []. - iSplitR=>//=. rewrite right_id. auto 20 with iFrame. } - iApply type_jump; solve_typing. - Qed. -End ref_functions. diff --git a/theories/typing/lib/refcell/refcell.v b/theories/typing/lib/refcell/refcell.v deleted file mode 100644 index 8af8e2c4..00000000 --- a/theories/typing/lib/refcell/refcell.v +++ /dev/null @@ -1,212 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree numbers. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -Set Default Proof Using "Type". - -Definition refcell_stR := - optionUR (prodR (prodR - (agreeR (prodO lftO boolO)) - fracR) - positiveR). -Class refcellG Σ := - RefCellG :> inG Σ (authR refcell_stR). -Definition refcellΣ : gFunctors := #[GFunctor (authR refcell_stR)]. -Global Instance subG_refcellΣ {Σ} : subG refcellΣ Σ → refcellG Σ. -Proof. solve_inG. Qed. - -Definition refcell_st := option ((lft * Datatypes.bool) * frac * positive). -Definition refcell_st_to_R (st : refcell_st) : refcell_stR := - match st with - | None => None - | Some (x, q, n) => Some (to_agree x, q, n) - end. -Coercion refcell_st_to_R : refcell_st >-> ucmra_car. - -Definition Z_of_refcell_st (st : refcell_st) : Z := - match st with - | None => 0 - | Some (_, false, _, n) => Zpos n (* Immutably borrowed *) - | Some (_, true, _, n) => Zneg n (* Mutably borrowed *) - end. - -Definition reading_stR (q : frac) (ν : lft) : refcell_stR := - refcell_st_to_R $ Some (ν, false, q, xH). -Definition writing_stR (q : frac) (ν : lft) : refcell_stR := - refcell_st_to_R $ Some (ν, true, q, xH). - -Definition refcellN := lrustN .@ "refcell". -Definition refcell_invN := refcellN .@ "inv". - -Section refcell_inv. - Context `{!typeG Σ, !refcellG Σ}. - - Definition refcell_inv tid (l : loc) (γ : gname) (α : lft) ty : iProp Σ := - (∃ st, l ↦ #(Z_of_refcell_st st) ∗ own γ (● (refcell_st_to_R st)) ∗ - match st return _ with - | None => - (* Not borrowed. *) - ∃ depth, ⧖depth ∗ &{α}((l +ₗ 1) ↦∗: ty.(ty_own) depth tid) - | Some (ν, rw, q, _) => - (1.[ν] ={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=∗ - ∃ depth, ⧖depth ∗ &{α}((l +ₗ 1) ↦∗: ty.(ty_own) depth tid)) ∗ - (∃ q', ⌜(q + q')%Qp = 1%Qp⌝ ∗ q'.[ν]) ∗ - if rw then (* Mutably borrowed. *) True - else (* Immutably borrowed. *) ty.(ty_shr) (α ⊓ ν) tid (l +ₗ 1) - end)%I. - - Lemma refcell_inv_proper E L ty1 ty2 q : - eqtype E L ty1 ty2 → - llctx_interp L q -∗ ∀ 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 "HE") as "(% & #Hout & #Hown & #Hshr)". - iAssert (□ (&{α}((l +ₗ 1) ↦∗: ty_own ty1 tid) -∗ - &{α}((l +ₗ 1) ↦∗: ty_own ty2 tid)))%I as "#Hb". - { iIntros "!> H". iApply bor_iff; last done. - iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl; - iFrame; by iApply "Hown". } - iDestruct "H" as (st) "H"; iExists st; - iDestruct "H" as "($&$&H)"; destruct st as [[[[ν rw] q' ] n]|]; try done; - last by iApply "Hb". - iDestruct "H" as "(Hh & $ & H)". iSplitL "Hh". - { iIntros "Hν". iMod ("Hh" with "Hν") as "Hh". iModIntro. iNext. - iMod "Hh". by iApply "Hb". } - destruct rw; try done. by iApply "Hshr". - Qed. -End refcell_inv. - -Section refcell. - Context `{!typeG Σ, !refcellG Σ}. - - (* Original Rust type: - pub struct RefCell<T: ?Sized> { - borrow: Cell<BorrowFlag>, - value: UnsafeCell<T>, - } - *) - - Program Definition refcell (ty : type) := - tc_opaque - {| ty_size := S ty.(ty_size); - ty_lfts := ty.(ty_lfts); ty_E := ty.(ty_E); - ty_own tid vl := - match vl return _ with - | #(LitInt z) :: vl' => ty.(ty_own) tid vl' - | _ => False - end%I; - ty_shr κ tid l := - (∃ α γ, κ ⊑ α ∗ α ⊑ ty_lft ty ∗ - &na{α, tid, refcell_invN}(refcell_inv tid l γ α ty))%I |}. - Next Obligation. - iIntros (??[|[[]|]]); try iIntros "[]". rewrite ty_size_eq /=. by iIntros (->). - Qed. - Next Obligation. - iIntros (ty E κ l tid q ?) "#LFT #Hout Hb Htok". - iMod (bor_acc_cons with "LFT Hb Htok") as "[H Hclose]". done. - iDestruct "H" as ([|[[| |n]|]vl]) "[H↦ H]"; try iDestruct "H" as ">[]". - iDestruct "H" as "Hown". - iMod ("Hclose" $! ((∃ n:Z, l ↦ #n) ∗ - (l +ₗ 1) ↦∗: ty.(ty_own) tid) with "[] [-]")%I as "[H [Htok Htok']]". - { iIntros "!> [Hn Hvl] !>". iDestruct "Hn" as (n') "Hn". - iDestruct "Hvl" as (vl') "[H↦ Hvl]". - iExists (#n'::vl'). rewrite heap_mapsto_vec_cons. iFrame. } - { iNext. rewrite heap_mapsto_vec_cons. iDestruct "H↦" as "[Hn Hvl]". - iSplitL "Hn"; eauto with iFrame. } - iMod (bor_sep with "LFT H") as "[Hn Hvl]". done. - iMod (bor_acc_cons with "LFT Hn Htok") as "[H Hclose]". done. - iAssert ((q / 2).[κ] ∗ ▷ ∃ γ, refcell_inv tid l γ κ ty)%I with "[> -Hclose]" - as "[$ HQ]"; last first. - { iMod ("Hclose" with "[] HQ") as "[Hb $]". - - iIntros "!> H !>". iNext. iDestruct "H" as (γ st) "(H & _ & _)". eauto. - - iMod (bor_exists with "LFT Hb") as (γ) "Hb". done. - iExists κ, γ. iSplitR. by iApply lft_incl_refl. by iSplitR; [|iApply bor_na]. } - clear dependent n. iDestruct "H" as ([|n|n]) "Hn"; try lia. - - iFrame. iMod (own_alloc (● None)) as (γ) "Hst"; first by apply auth_auth_valid. - iExists γ, None. by iFrame. - - iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]". done. - iMod (own_alloc (● (refcell_st_to_R $ Some (ν, false, (1/2)%Qp, n)))) as (γ) "Hst". - { by apply auth_auth_valid. } - iApply (fupd_mask_mono (↑lftN)); first done. - iMod (rebor _ _ (κ ⊓ ν) with "LFT [] Hvl") as "[Hvl Hh]". done. - { iApply lft_intersect_incl_l. } - iDestruct (lft_intersect_acc with "Htok' Htok1") as (q') "[Htok Hclose]". - iMod (ty_share with "LFT [] Hvl Htok") as "[Hshr Htok]". done. - { iApply lft_incl_trans; [|done]. iApply lft_intersect_incl_l. } - iDestruct ("Hclose" with "Htok") as "[$ Htok]". - iExists γ, _. iFrame "Hst Hn Hshr". - iSplitR "Htok2"; last by iExists _; iFrame; rewrite Qp_div_2. - iIntros "!> !> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν". - iApply fupd_mask_mono; last iApply "Hh". 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-. } - iSplitR; [|done]. iExists (1/2)%Qp. rewrite Qp_div_2. iSplitR; [done|]. - iApply lft_tok_static. - Qed. - Next Obligation. - iIntros (?????) "#Hκ H". iDestruct "H" as (α γ) "[#??]". - iExists _, _. iFrame. by iApply lft_incl_trans. - Qed. - - Global Instance refcell_type_ne : TypeNonExpansive refcell. - Proof. - split. - - apply (type_lft_morphism_add _ static [] []) => ?. - + rewrite left_id. iApply lft_equiv_refl. - + by rewrite /elctx_interp /= left_id right_id. - - by move=>/= ?? ->. - - intros. by destruct vl as [|[[]|]]=>/=. - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. rewrite /= /refcell_inv. - repeat ((eapply dist_le; [apply Ho|lia]) || (eapply dist_le; [apply Hs|lia]) || - apply equiv_dist, lft_incl_equiv_proper_r, Hl || f_contractive || f_equiv). - Qed. - - Global Instance refcell_ne : NonExpansive refcell. - Proof. - rewrite /refcell /refcell_inv /=. intros n ty1 ty2 Hty12. split. - - by rewrite /= Hty12. - - apply Hty12. - - apply Hty12. - - rewrite /= ![X in X {| ty_own := _ |}]/ty_own. - intros. do 3 f_equiv. apply Hty12. - - simpl. intros. repeat (apply Hty12 || f_equiv). - Qed. - - Global Instance refcell_mono E L : Proper (eqtype E L ==> subtype E L) refcell. - Proof. - (* TODO : this proof is essentially [solve_proper], but within the logic. - It would easily gain from some automation. *) - iIntros (ty1 ty2 EQ 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; [by simpl; auto|iSplit; [done|iSplit; iIntros "!> * H /="]]. - - destruct vl as [|[[]|]]=>//=. by iApply "Hown". - - iDestruct "H" as (a γ) "(Ha & #? & H)". iExists a, γ. iFrame. iSplitR. - + by iApply lft_incl_trans. - + iApply na_bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H". - by iApply "Hty1ty2". by iApply "Hty2ty1". - Qed. - Lemma refcell_mono' E L ty1 ty2 : - eqtype E L ty1 ty2 → subtype E L (refcell ty1) (refcell ty2). - Proof. eapply refcell_mono. Qed. - Global Instance refcell_proper E L : Proper (eqtype E L ==> eqtype E L) refcell. - Proof. by split; apply refcell_mono. Qed. - Lemma refcell_proper' E L ty1 ty2 : - eqtype E L ty1 ty2 → eqtype E L (refcell ty1) (refcell ty2). - Proof. eapply refcell_proper. Qed. - - Global Instance refcell_send ty : - Send ty → Send (refcell ty). - Proof. move=>???[|[[]|]]//=. Qed. -End refcell. - -Global Hint Resolve refcell_mono' refcell_proper' : lrust_typing. diff --git a/theories/typing/lib/refcell/refcell_code.v b/theories/typing/lib/refcell/refcell_code.v deleted file mode 100644 index d5070402..00000000 --- a/theories/typing/lib/refcell/refcell_code.v +++ /dev/null @@ -1,319 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lang.lib Require Import memcpy. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing option. -From lrust.typing.lib.refcell Require Import refcell ref refmut. -Set Default Proof Using "Type". - -Section refcell_functions. - Context `{!typeG Σ, !refcellG Σ}. - - (* Constructing a refcell. *) - Definition refcell_new ty : val := - fn: ["x"] := - let: "r" := new [ #(S ty.(ty_size))] in - "r" +ₗ #0 <- #0;; - "r" +ₗ #1 <-{ty.(ty_size)} !"x";; - delete [ #ty.(ty_size) ; "x"];; return: ["r"]. - - Lemma refcell_new_type ty : - typed_val (refcell_new ty) (fn(∅; ty) → refcell ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new (S ty.(ty_size))); [solve_typing..|]. - iIntros (r tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hr Hx]". - iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & Hx↦ & Hx & Hx†)". subst x. - iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr↦ & Hr†)". subst r. - inv_vec vlr=>??. iDestruct (heap_mapsto_vec_cons with "Hr↦") as "[Hr↦0 Hr↦1]". wp_op. - rewrite shift_loc_0. wp_write. wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz. - wp_apply (wp_memcpy with "[$Hr↦1 $Hx↦]"); [by auto using vec_to_list_length..|]. - iIntros "[Hr↦1 Hx↦]". wp_seq. - iApply (type_type _ _ _ [ #lx ◁ box (uninit (ty_size ty)); #lr ◁ box (refcell ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=. - iFrame. iSplitL "Hx↦". - - iExists _. rewrite uninit_own. auto. - - simpl. iFrame. iExists (_::_). rewrite heap_mapsto_vec_cons. iFrame=>//=. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* The other direction: getting ownership out of a refcell. *) - Definition refcell_into_inner ty : val := - fn: ["x"] := - let: "r" := new [ #ty.(ty_size)] in - "r" <-{ty.(ty_size)} !("x" +ₗ #1);; - (* TODO: Can we make it so that the parenthesis above are mandatory? - Leaving them away is inconsistent with `let ... := !"x" +ₗ #1`. *) - delete [ #(S ty.(ty_size)) ; "x"];; return: ["r"]. - - Lemma refcell_into_inner_type ty : - typed_val (refcell_into_inner ty) (fn(∅; refcell ty) → ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_new ty.(ty_size)); [solve_typing..|]. - iIntros (r tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hr Hx]". - iDestruct (ownptr_own with "Hx") as (lx vlx) "(% & >Hx↦ & Hx & Hx†)". subst x. - inv_vec vlx=>-[[|?|?]|????] vl; simpl; try iDestruct "Hx" as ">[]". - iDestruct (heap_mapsto_vec_cons with "Hx↦") as "[Hx↦0 Hx↦1]". - iDestruct (ownptr_uninit_own with "Hr") as (lr vlr) "(% & Hr↦ & Hr†)". subst r. - wp_op. wp_apply (wp_memcpy with "[$Hr↦ $Hx↦1]"); [by auto using vec_to_list_length..|]. - iIntros "[Hr↦ Hx↦1]". wp_seq. - iApply (type_type _ _ _ [ #lx ◁ box (uninit (S (ty_size ty))); #lr ◁ box ty] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame. - iSplitR "Hr↦ Hx". - - iExists (_::_). rewrite heap_mapsto_vec_cons uninit_own. iFrame. - rewrite /= vec_to_list_length. auto. - - iExists vl. iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition refcell_get_mut : val := - fn: ["x"] := - let: "x'" := !"x" in - "x" <- "x'" +ₗ #1;; (* Get the second field *) - return: ["x"]. - - Lemma refcell_get_mut_type ty : - typed_val refcell_get_mut - (fn(∀ α, ∅; &uniq{α} (refcell ty)) → &uniq{α} ty)%T. - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL HC HT". - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx [#? Hx']]". - destruct x' as [[|lx'|]|]; try iDestruct "Hx'" as "[]". - iAssert (&{α} (∃ (z : Z), lx' ↦ #z) ∗ - (&uniq{α} ty).(ty_own) tid [ #(lx' +ₗ 1)])%I with "[> Hx']" as "[_ Hx']". - { simpl. iFrame "#". iApply bor_sep; [done..|]. iApply (bor_proper with "Hx'"). iSplit. - - iIntros "[H1 H2]". iDestruct "H1" as (z) "?". iDestruct "H2" as (vl) "[??]". - iExists (_::_). rewrite heap_mapsto_vec_cons /=. iFrame=>//=. - - iIntros "H". - iDestruct "H" as ([|[[| |z]|]vl]) "[H↦ H]"; - simpl; try iDestruct "H" as "[]". - rewrite heap_mapsto_vec_cons. iDestruct "H↦" as "[H↦1 H↦2]". - iSplitL "H↦1"; eauto. iExists _. iFrame. } - destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". - iDestruct "Hx" as (vl) "[Hx↦ Hx]". rewrite uninit_own. wp_op. - iApply (type_type _ _ _ - [ #lx ◁ box (uninit 1); #(lx' +ₗ 1) ◁ &uniq{α}ty] - with "[] LFT HE Hna HL HC [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame. - iNext. iExists _. rewrite uninit_own. iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Shared borrowing. *) - Definition refcell_try_borrow : val := - fn: ["x"] := - let: "r" := new [ #3 ] in - withcont: "k": - let: "x'" := !"x" in - let: "n" := !"x'" in - if: "n" ≤ #-1 then - "r" <-{Σ none} ();; - "k" [] - else - "x'" <- "n" + #1;; - let: "ref" := new [ #2 ] in - "ref" <- "x'" +ₗ #1;; - "ref" +ₗ #1 <- "x'";; - "r" <-{2,Σ some} !"ref";; - delete [ #2; "ref"];; - "k" [] - cont: "k" [] := - delete [ #1; "x"];; return: ["r"]. - - Lemma refcell_try_borrow_type ty : - typed_val refcell_try_borrow - (fn(∀ α, ∅; &shr{α}(refcell ty)) → option (ref α ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] (λ _, [x ◁ box (&shr{α}(refcell ty)); - r ◁ box (option (ref α ty))])); - [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; simpl_subst; last first. - { iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iApply type_deref; [solve_typing..|]. - iIntros (x' tid) "#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αβ & Hβty & Hinv)". - 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βtok1 Hβtok2] Hclose']". done. - iMod (na_bor_acc with "LFT Hinv Hβtok1 Hna") as "(INV & Hna & Hclose'')"; try done. - iDestruct "INV" as (st) "(Hlx & Hownst & Hst)". wp_read. wp_let. wp_op; case_bool_decide; wp_if. - - iMod ("Hclose''" with "[Hlx Hownst Hst] Hna") as "[Hβtok1 Hna]"; - first by iExists st; iFrame. - iMod ("Hclose'" with "[$Hβtok1 $Hβtok2]") as "Hα". - iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(refcell ty)); r ◁ box (uninit 3) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. } - iApply (type_sum_unit (option $ ref α ty)); [solve_typing..|]; first last. - simpl. iApply type_jump; solve_typing. - - wp_op. wp_write. wp_apply wp_new; [done..|]. - iIntros (lref) "(H† & Hlref)". wp_let. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hlref" as "[Hlref0 Hlref1]". wp_op. wp_write. wp_op. wp_write. - iAssert (∃ qν ν, (qβ / 2).[β] ∗ (qν).[ν] ∗ - ty_shr ty (β ⊓ ν) tid (lx +ₗ 1) ∗ - own γ (◯ reading_stR qν ν) ∗ refcell_inv tid lx γ β ty)%I - with "[> Hlx Hownst Hst Hβtok2]" as (q' ν) "(Hβtok2 & Hν & Hshr & Hreading & INV)". - { destruct st as [[[[ν []] s] n]|]. - - simpl in *; lia. - - iDestruct "Hst" as "(H† & Hst & #Hshr)". - iDestruct "Hst" as (q' Hqq') "[Hν1 Hν2]". - iExists _, _. iFrame "Hν1". iMod (own_update with "Hownst") as "[Hownst ?]". - { apply auth_update_alloc, - (op_local_update_discrete _ _ (reading_stR (q'/2)%Qp ν)) => ?. - split; [split|]. - - by rewrite /= agree_idemp. - - apply frac_valid'. rewrite /= -Hqq' comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. - - done. } - iFrame "∗#". iExists (Some (ν, false, _, _)). iFrame "∗#". - rewrite [_ ⋅ _]comm -Some_op -!pair_op agree_idemp. iFrame. - iExists _. iFrame. rewrite -(assoc Qp_add) Qp_div_2 //. - - 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 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. - { iApply lft_intersect_incl_l. } - iMod (ty_share with "LFT [] Hst Htok") as "[#Hshr Htok]". done. - { iApply lft_incl_trans; [|done]. iApply lft_intersect_incl_l. } - iFrame "Hshr". - iDestruct ("Hclose" with "Htok") as "[$ Htok2]". iExists _. iFrame "∗#". - iSplitR "Htok2". - + iIntros "!> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. - iNext. iMod "Hν". - iMod fupd_intro_mask' as "Hclose"; last iMod ("Hh" with "[Hν]") as "$". - { set_solver-. } - * rewrite -lft_dead_or. auto. - * done. - + iExists _. iFrame. by rewrite Qp_div_2. } - iMod ("Hclose''" with "[$INV] Hna") as "[Hβtok1 Hna]". - iMod ("Hclose'" with "[$Hβtok1 $Hβtok2]") as "Hα". - iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(refcell ty)); r ◁ box (uninit 3); #lref ◁ box (ref α ty)] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. - rewrite tctx_hasty_val' //. rewrite /= freeable_sz_full. iFrame. - iExists [_; _]. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iFrame. simpl. iExists _, _, _, _, _. iFrame "∗#". iApply ty_shr_mono; try by auto. - iApply lft_intersect_mono. done. iApply lft_incl_refl. } - iApply (type_sum_memcpy (option $ ref α ty)); [solve_typing..|]. - simpl. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Unique borrowing. *) - Definition refcell_try_borrow_mut : val := - fn: ["x"] := - let: "r" := new [ #3 ] in - withcont: "k": - let: "x'" := !"x" in - let: "n" := !"x'" in - if: "n" = #0 then - "x'" <- #(-1);; - let: "ref" := new [ #2 ] in - "ref" <- "x'" +ₗ #1;; - "ref" +ₗ #1 <- "x'";; - "r" <-{2,Σ some} !"ref";; - delete [ #2; "ref"];; - "k" [] - else - "r" <-{Σ none} ();; - "k" [] - cont: "k" [] := - delete [ #1; "x"];; return: ["r"]. - - Lemma refcell_try_borrow_mut_type ty : - typed_val refcell_try_borrow_mut - (fn(∀ α, ∅; &shr{α}(refcell ty)) → option (refmut α ty))%T. - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] (λ _, [x ◁ box (&shr{α}(refcell ty)); - r ◁ box (option (refmut α ty))])); - [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; - simpl_subst; last first. - { iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iApply type_deref; [solve_typing..|]. - iIntros (x' tid) "#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αβ & Hβty & Hinv)". - 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βtok Hclose']". done. - iMod (na_bor_acc with "LFT Hinv Hβtok Hna") as "(INV & Hna & Hclose'')"; try done. - iDestruct "INV" as (st) "(Hlx & Hownst & Hb)". wp_read. wp_let. wp_op; case_bool_decide; wp_if. - - wp_write. wp_apply wp_new; [done..|]. - iIntros (lref) "(H† & Hlref)". wp_let. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hlref" as "[Hlref0 Hlref1]". wp_op. wp_write. wp_op. wp_write. - destruct st as [[[[ν []] s] n]|]; try done. - iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]". done. - iMod (own_update with "Hownst") as "[Hownst ?]". - { by eapply auth_update_alloc, - (op_local_update_discrete _ _ (refcell_st_to_R $ Some (ν, true, (1/2)%Qp, xH))). } - rewrite (right_id None). iApply fupd_wp. iApply (fupd_mask_mono (↑lftN)); first done. - iMod (rebor _ _ (β ⊓ ν) with "LFT [] Hb") as "[Hb Hbh]". done. - { iApply lft_intersect_incl_l. } - iModIntro. iMod ("Hclose''" with "[Hlx Hownst Hbh Htok1] Hna") as "[Hβtok Hna]". - { iExists _. iFrame. iNext. iSplitL "Hbh". - - iIntros "Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν". - iMod fupd_intro_mask' as "Hclose"; last iMod ("Hbh" with "[Hν]") as "$". - { set_solver-. } - * rewrite -lft_dead_or. auto. - * done. - - iSplitL; [|done]. iExists _. iFrame. by rewrite Qp_div_2. } - iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(refcell ty)); r ◁ box (uninit 3); #lref ◁ box (refmut α ty)] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. - rewrite tctx_hasty_val' //. rewrite /= freeable_sz_full. iFrame. - iExists [_; _]. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iFrame. simpl. - iExists _, _, _, _, _. iFrame "#∗". iSplitL. - - iApply (bor_shorten with "[] [$Hb]"). - iApply lft_intersect_mono; first done. iApply lft_incl_refl. - - iApply lft_incl_trans; [|done]. iApply lft_incl_trans; [|done]. - iApply lft_intersect_incl_l. } - iApply (type_sum_memcpy (option $ refmut α ty)); [solve_typing..|]. - simpl. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - - iMod ("Hclose''" with "[Hlx Hownst Hb] Hna") as "[Hβtok Hna]"; - first by iExists st; iFrame. - iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(refcell ty)); r ◁ box (uninit 3) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. } - iApply (type_sum_unit (option $ refmut α ty)); [solve_typing..|]. - simpl. iApply type_jump; solve_typing. - Qed. -End refcell_functions. diff --git a/theories/typing/lib/refcell/refmut.v b/theories/typing/lib/refcell/refmut.v deleted file mode 100644 index 8e634a2e..00000000 --- a/theories/typing/lib/refcell/refmut.v +++ /dev/null @@ -1,143 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.refcell Require Import refcell. -Set Default Proof Using "Type". - -Section refmut. - Context `{!typeG Σ, !refcellG Σ}. - - (* The Rust type looks as follows (after some unfolding): - - pub struct RefMut<'b, T: ?Sized + 'b> { - value: &'b mut T, - borrow: &'b Cell<BorrowFlag>, - } - - In other words, we have a pointer to the data, and a pointer - to the refcount field of the RefCell. *) - - Program Definition refmut (α : lft) (ty : type) := - tc_opaque - {| ty_size := 2; - ty_lfts := α :: ty.(ty_lfts); ty_E := ty.(ty_E) ++ ty_outlives_E ty α; - ty_own tid vl := - match vl return _ with - | [ #(LitLoc lv); #(LitLoc lrc) ] => - ∃ ν q γ β ty', &{α ⊓ ν}(lv ↦∗: ty.(ty_own) tid) ∗ - α ⊑ β ∗ α ⊓ ν ⊑ ty_lft ty ∗ - &na{β, tid, refcell_invN}(refcell_inv tid lrc γ β ty') ∗ - q.[ν] ∗ own γ (◯ writing_stR q ν) - | _ => False - end; - ty_shr κ tid l := - ∃ (lv lrc : loc), - &frac{κ} (λ q, l↦∗{q} [ #lv; #lrc]) ∗ - □ ∀ F q, ⌜↑shrN ∪ ↑lftN ⊆ F⌝ -∗ q.[α ⊓ κ] - ={F}[F∖↑shrN]▷=∗ ty.(ty_shr) (α ⊓ κ) tid lv ∗ q.[α ⊓ κ] |}%I. - Next Obligation. - iIntros (???[|[[]|][|[[]|][]]]); try iIntros "[]". by iIntros "_". - Qed. - Next Obligation. - iIntros (α ty E κ l tid q HE) "#LFT Hl Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb". done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]". done. - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦". done. - destruct vl as [|[[|lv|]|][|[[|lrc|]|][]]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - iMod (bor_exists with "LFT Hb") as (ν) "Hb". done. - iMod (bor_exists with "LFT Hb") as (q') "Hb". done. - iMod (bor_exists with "LFT Hb") as (γ) "Hb". done. - iMod (bor_exists with "LFT Hb") as (β) "Hb". done. - iMod (bor_exists with "LFT Hb") as (ty') "Hb". done. - iMod (bor_sep with "LFT Hb") as "[Hb H]". done. - iMod (bor_sep with "LFT H") as "[Hαβ H]". done. - iMod (bor_sep with "LFT H") as "[Hαty H]". done. - iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ Htok]". done. - iMod (bor_persistent with "LFT Hαty Htok") as "[#Hαty $]". done. - iMod (bor_sep with "LFT H") as "[_ H]". done. - iMod (bor_sep with "LFT H") as "[H _]". done. - iMod (bor_fracture (λ q, (q' * q).[ν])%I with "LFT [H]") as "H". done. - { by rewrite Qp_mul_1_r. } - iDestruct (frac_bor_lft_incl with "LFT H") as "#Hκν". iClear "H". - iExists _, _. iFrame "H↦". - iApply delay_sharing_nested; [..|iApply (bor_shorten with "[] Hb")]=>//=. - - iApply lft_intersect_mono; [|done]. iApply lft_incl_refl. - - iApply lft_intersect_incl_r. - Qed. - Next Obligation. - iIntros (??????) "#? H". iDestruct "H" as (lv lrc) "[#Hf #H]". - iExists _, _. iSplit. - - by iApply frac_bor_shorten. - - iIntros "!> * % Htok". - iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. iApply lft_incl_refl. done. } - 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. iApply lft_incl_refl. done. - Qed. - - Global Instance refmut_type_contractive α : TypeContractive (refmut α). - Proof. - split. - - apply (type_lft_morphism_add _ α [α] []) => ?. - + iApply lft_equiv_refl. - + by rewrite elctx_interp_app /= elctx_interp_ty_outlives_E - /= /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid [|[[|lv|]|][|[[|lrc|]|][]]]=>//=. - repeat (apply Ho || apply equiv_dist, lft_incl_equiv_proper_r, Hl || - f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. simpl. - repeat (apply Hs || f_contractive || f_equiv). - Qed. - Global Instance refmut_ne α : NonExpansive (refmut α). - Proof. solve_ne_type. Qed. - - Global Instance refmut_mono E L : - Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) refmut. - Proof. - intros α1 α2 Hα ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (q) "HL". - iDestruct (Hty with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". - iIntros "!> #HE". iDestruct ("Hα" with "HE") as "Hα1α2". - iDestruct ("Hty" with "HE") as "(%&#[??]&#Ho&#Hs)". - iSplit; [done|iSplit; [|iSplit; iModIntro]]. - - simpl. by iApply lft_intersect_mono. - - iIntros (tid [|[[]|][|[[]|][]]]) "H"=>//=. - iDestruct "H" as (ν γ q' β ty') "(Hb & #H⊑ & #Hl & #Hinv & Hν & Hown)". - iExists ν, γ, q', β, ty'. iFrame "∗#". iSplit; [|iSplit]. - + iApply bor_shorten; last iApply bor_iff; last done. - * iApply lft_intersect_mono; first done. 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; [|done]. iApply lft_incl_trans; [|iApply "Hl"]. - iApply lft_intersect_mono; [done|]. iApply lft_incl_refl. - - iIntros (κ tid l) "H /=". iDestruct "H" as (lv lrc) "H". iExists lv, lrc. - iDestruct "H" as "[$ #H]". iIntros "!> * % Htok". - iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. done. 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". - Qed. - Global Instance refmut_mono_flip E L : - Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) refmut. - Proof. intros ??????. by apply refmut_mono. Qed. - Lemma refmut_mono' E L α1 α2 ty1 ty2 : - lctx_lft_incl E L α2 α1 → eqtype E L ty1 ty2 → - subtype E L (refmut α1 ty1) (refmut α2 ty2) . - Proof. intros. by eapply refmut_mono. Qed. - Global Instance refmut_proper E L : - Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) refmut. - Proof. intros ??[]???. split; by apply refmut_mono'. Qed. - Lemma refmut_proper' E L α1 α2 ty1 ty2 : - lctx_lft_eq E L α1 α2 → eqtype E L ty1 ty2 → - eqtype E L (refmut α1 ty1) (refmut α2 ty2). - Proof. intros. by eapply refmut_proper. Qed. -End refmut. - -Global Hint Resolve refmut_mono' refmut_proper' : lrust_typing. diff --git a/theories/typing/lib/refcell/refmut_code.v b/theories/typing/lib/refcell/refmut_code.v deleted file mode 100644 index b23d8fa6..00000000 --- a/theories/typing/lib/refcell/refmut_code.v +++ /dev/null @@ -1,340 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree numbers. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.refcell Require Import refcell refmut. -Set Default Proof Using "Type". - -Section refmut_functions. - Context `{!typeG Σ, !refcellG Σ}. - - (* Turning a refmut into a shared borrow. *) - Definition refmut_deref : val := - fn: ["x"] := - let: "x'" := !"x" in - let: "r'" := !"x'" in - letalloc: "r" <- "r'" in - delete [ #1; "x"];; return: ["r"]. - - Lemma refmut_deref_type ty : - typed_val refmut_deref (fn(∀ '(α, β), ∅; &shr{α}(refmut β ty)) → &shr{α}ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]=>//=. - iDestruct "Hx'" as (lv lrc) "#(Hfrac & #Hshr)". - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "([Hα1 Hα2] & HL & Hclose')"; - [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'')"; - [solve_typing..|]. - iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hcloseβα2]". - wp_bind (!#lx')%E. iApply (wp_step_fupd with "[Hα2β]"); - [|by iApply ("Hshr" with "[] Hα2β")|]; first done. - iMod "H↦" as "[H↦1 H↦2]". wp_read. iIntros "[#Hshr' Hα2β] !>". wp_let. - iDestruct ("Hcloseβα2" with "Hα2β") as "[Hβ Hα2]". - iMod ("Hcloseα1" with "[$H↦1 $H↦2]") as "Hα1". - iMod ("Hclose''" with "Hβ HL") as "HL". iMod ("Hclose'" with "[$] HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [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 (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. - iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Turning a refmut into a unique borrow. *) - Definition refmut_derefmut : val := refmut_deref. - - Lemma refmut_derefmut_type ty : - typed_val refmut_derefmut - (fn(∀ '(α, β), ∅; &uniq{α}(refmut β ty)) → &uniq{α}ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "(Hx & #Hαty & Hx')". destruct x' as [[|lx'|]|]; try done. - iMod (bor_exists with "LFT Hx'") as (vl) "H". done. - iMod (bor_sep with "LFT H") as "[H↦ H]". done. - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose')"; - [solve_typing..|]. - destruct vl as [|[[|lv|]|][|[[|lrc|]|][]]]; simpl; - try by iMod (bor_persistent with "LFT H Hα") as "[>[] _]". - iMod (bor_exists with "LFT H") as (ν) "H". done. - iMod (bor_exists with "LFT H") as (q) "H". done. - 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 (ty') "H". done. - iMod (bor_sep with "LFT H") as "[Hb H]". done. - iMod (bor_sep with "LFT H") as "[Hβδ H]". done. - iMod (bor_sep with "LFT H") as "[Hβty H]". done. - iMod (bor_persistent with "LFT Hβδ Hα") as "[#Hβδ Hα]". done. - iMod (bor_persistent with "LFT Hβty Hα") as "[#Hβty Hα]". done. - iMod (bor_sep with "LFT H") as "[_ H]". done. - iMod (bor_sep with "LFT H") as "[H _]". done. - iMod (bor_fracture (λ q', (q * q').[ν])%I with "LFT [H]") as "H". done. - { by rewrite Qp_mul_1_r. } - iDestruct (frac_bor_lft_incl _ _ q with "LFT H") as "#Hαν". iClear "H". - rewrite heap_mapsto_vec_cons. iMod (bor_sep with "LFT H↦") as "[H↦ _]". done. - iMod (bor_acc with "LFT H↦ Hα") as "[H↦ Hcloseα]". done. - wp_bind (!#lx')%E. iMod (bor_unnest with "LFT Hb") as "Hb"; first done. - wp_read. wp_let. iMod "Hb". - iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]". iMod ("Hclose'" with "Hα HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [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 "Hx". iSplitR. - - iApply lft_incl_trans; [|done]. by iApply lft_incl_glb. - - iApply (bor_shorten with "[] Hb"). - by iApply lft_incl_glb; [iApply lft_incl_glb|iApply lft_incl_refl]. } - iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|]. - iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Dropping a refmut and set the count to 0 in the corresponding refcell. *) - Definition refmut_drop : val := - fn: ["x"] := - let: "rc" := !("x" +ₗ #1) in - let: "n" := !"rc" in - "rc" <- "n" + #1;; - Endlft;; - delete [ #2; "x"];; - let: "r" := new [ #0] in return: ["r"]. - - Lemma refmut_drop_type ty : - typed_val refmut_drop (fn(∀ α, ∅; refmut α ty) → unit). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk Hx". rewrite tctx_interp_singleton tctx_hasty_val. - destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". - iDestruct "Hx" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hx"; simpl; - try iDestruct "Hx" as "[_ >[]]". - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hx" as "[[Hx↦1 Hx↦2] Hx]". wp_op. wp_read. wp_let. - iDestruct "Hx" as (ν q γ β ty') "(Hb & #Hαβ & #Hβty & #Hinv & Hν & H◯)". - 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 (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hcloseβ)"; [done..|]. - iDestruct "INV" as (st) "(H↦lrc & H● & INV)". wp_read. wp_let. wp_op. wp_write. - iAssert (|={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=> 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 _] - %auth_both_valid_discrete. - destruct st as [[[[??]?]?]|]; [|done]. move: Hst=>-[= ???]. subst. - destruct INCL as [[[[ν' ?]%to_agree_inj ?] ?]| - [[[??]%to_agree_included [q'' Hq'']]%prod_included [n' Hn']]%prod_included]. - - simpl in *. setoid_subst. iExists None. iFrame. - iMod (own_update_2 with "H● H◯") as "$". - { apply auth_update_dealloc. rewrite -(right_id None op (Some _)). - apply cancel_local_update_unit, _. } - iDestruct "INV" as "(H† & Hq & _)". - iApply "H†". - iDestruct "Hq" as (q) "(<- & ?)". iFrame. - - simpl in *. setoid_subst. iExists (Some (_, _, _, _)). - iMod (own_update_2 with "H● H◯") as "$". - { apply auth_update_dealloc. rewrite -(agree_idemp (to_agree _)) !pair_op Some_op. - apply (cancel_local_update_unit (writing_stR q _)), _. } - iDestruct "INV" as "(H† & Hq & _)". - rewrite /= (_:Z.neg (1%positive ⋅ n') + 1 = Z.neg n'); - last (rewrite pos_op_plus; lia). iFrame. - iApply step_fupd_intro; [set_solver-|]. iSplitL; [|done]. - iDestruct "Hq" as (q' ?) "?". iExists (q+q')%Qp. iFrame. - rewrite assoc (comm _ q'' q) //. } - wp_bind Endlft. iApply (wp_mask_mono _ (↑lftN ∪ ↑lft_userN)); first done. - iApply (wp_step_fupd with "INV"); [set_solver-|]. wp_seq. iIntros "{Hb} Hb !>". - iMod ("Hcloseβ" with "Hb Hna") as "[Hβ Hna]". - iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". wp_seq. - iApply (type_type _ _ _ [ #lx ◁ box (uninit 2)] - with "[] LFT HE Hna HL Hk"); last first. - { rewrite tctx_interp_singleton tctx_hasty_val' //. iFrame. iExists [ #lv;#lrc]. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton uninit_own. iFrame. auto. } - iApply type_delete; [solve_typing..|]. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_jump; solve_typing. - Qed. - - (* Apply a function within the refmut, typically for accessing a component. *) - Definition refmut_map (call_once : val) : val := - fn: ["ref"; "f"] := - let: "call_once" := call_once in - let: "x'" := !"ref" in - letalloc: "x" <- "x'" in - letcall: "r" := "call_once" ["f"; "x"]%E in - let: "r'" := !"r" in delete [ #1; "r"];; - "ref" <- "r'";; - return: ["ref"]. - - Lemma refmut_map_type ty1 ty2 call_once fty : - (* fty : for<'a>, FnOnce(&'a ty1) -> &'a ty2, as witnessed by the impl call_once *) - typed_val call_once (fn(∀ α, ∅; fty, &uniq{α}ty1) → &uniq{α}ty2) → - typed_val (refmut_map call_once) (fn(∀ α, ∅; refmut α ty1, fty) → refmut α ty2). - Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α ϝ ret arg). - inv_vec arg=>ref env. simpl_subst. - iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Href & Henv & _)". - rewrite (tctx_hasty_val _ ref). destruct ref as [[|lref|]|]; try done. - iDestruct "Href" as "[Href Href†]". - iDestruct "Href" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Href"; simpl; - try iDestruct "Href" as "[_ >[]]". - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Href" as "[[Href↦1 Href↦2] Href]". - iDestruct "Href" as (ν q γ β ty') "(Hbor & #Hαβ & #Hβty & #Hinv & >Hν & Hγ)". - wp_read. wp_let. wp_apply wp_new; first done. done. - 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..|]. - iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". - iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]". - rewrite -[ϝ in (α ⊓ ν) ⊓ ϝ](right_id_L). - iApply (type_call_iris _ [α ⊓ ν; ϝ] (α ⊓ ν) [_; _] - with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H† Hbor]"); [solve_typing|done| |]. - { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. - iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. by iFrame. } - iIntros ([[|r|]|]) "Hna Hανϝ Hr //". - iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]". - iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]". - iMod ("Hclose2" with "Hϝ HL") as "HL". iMod ("Hclose1" with "Hα HL") as "HL". - wp_rec. iDestruct "Hr" as "[Hr Hr†]". - iDestruct "Hr" as ([|[[]|][]]) "(Hr & #Hl & Hr')"=>//. - rewrite heap_mapsto_vec_singleton. wp_read. wp_let. - wp_apply (wp_delete _ _ _ [_] with "[Hr Hr†]")=>//. - { rewrite heap_mapsto_vec_singleton freeable_sz_full. iFrame. } iIntros "_". - wp_seq. wp_write. - iApply (type_type _ [_] _ [ #lref ◁ box (refmut α ty2) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_singleton tctx_hasty_val' //. simpl. iFrame. - iExists [_;_]. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iFrame. simpl. auto 20 with iFrame. } - iApply type_jump; solve_typing. - Qed. - - (* Apply a function within the refmut, and create two refmuts, - typically for splitting the reference. *) - Definition refmut_map_split (call_once : val) : val := - fn: ["refmut"; "f"] := - let: "call_once" := call_once in - let: "x'" := !"refmut" in - - letalloc: "x" <- "x'" in - letcall: "r" := "call_once" ["f"; "x"]%E in - let: "a" := !"r" in - let: "b" := !("r" +ₗ #1) in - delete [ #2; "r"];; - - let: "rc" := !("refmut" +ₗ #1) in - let: "n" := !"rc" in - "rc" <- "n" - #1;; - - delete [ #2; "refmut"];; - - let: "refmuts" := new [ #4 ] in - "refmuts" <- "a";; - "refmuts" +ₗ #1 <- "rc";; - "refmuts" +ₗ #2 <- "b";; - "refmuts" +ₗ #3 <- "rc";; - - return: ["refmuts"]. - - Lemma refmut_map_split_type ty ty1 ty2 call_once fty : - (* fty : for<'a>, FnOnce(&mut'a ty) -> (&mut'a ty1, &mut'a ty2), - as witnessed by the impl call_once *) - typed_val call_once - (fn(∀ α, ∅; fty, &uniq{α}ty) → Π[&uniq{α}ty1; &uniq{α}ty2]) → - typed_val (refmut_map_split call_once) - (fn(∀ α, ∅; refmut α ty, fty) → Π[refmut α ty1; refmut α ty2]). - Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α ϝ ret arg). - inv_vec arg=>refmut env. simpl_subst. - iApply type_let; [apply Hf | solve_typing |]. iIntros (f'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk (#Hf' & Hrefmut & Henv & _)". - rewrite (tctx_hasty_val _ refmut). destruct refmut as [[|lrefmut|]|]; try done. - iDestruct "Hrefmut" as "[Hrefmut Hrefmut†]". - iDestruct "Hrefmut" as ([|[[|lv|]|][|[[|lrc|]|][]]]) "Hrefmut"; simpl; - try iDestruct "Hrefmut" as "[_ >[]]". - rewrite {1}heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hrefmut" as "[[Hrefmut↦1 Hrefmut↦2] Hrefmut]". - iDestruct "Hrefmut" as (ν qν γ β ty') "(Hbor & #Hαβ & #Hl & #Hinv & >Hν & Hγ)". - wp_read. wp_let. wp_apply wp_new; [done..|]. - 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..|]. - iDestruct (lft_intersect_acc with "Hα Hν") as (?) "[Hαν Hclose3]". - iDestruct (lft_intersect_acc with "Hαν Hϝ") as (?) "[Hανϝ Hclose4]". - rewrite -[ϝ in (α ⊓ ν) ⊓ ϝ](right_id_L). - iApply (type_call_iris _ [α ⊓ ν; ϝ] (α ⊓ ν) [_; _] - with "LFT HE Hna [Hανϝ] Hf' [$Henv Hlx H† Hbor]"); [solve_typing|done| |]. - { rewrite big_sepL_singleton tctx_hasty_val' //. rewrite /= freeable_sz_full. - iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame. auto. } - iIntros ([[|r|]|]) "Hna Hανϝ Hr //". - iDestruct ("Hclose4" with "Hανϝ") as "[Hαν Hϝ]". - iDestruct ("Hclose3" with "Hαν") as "[Hα Hν]". - iMod ("Hclose2" with "Hϝ HL") as "HL". - wp_rec. iDestruct "Hr" as "[Hr Hr†]". - iDestruct "Hr" as (?) "[Hr H]". iDestruct "H" as (vl1 vl1' ->) "[[#? Hr1'] H]". - iDestruct "H" as (vl2 vl2' ->) "[[#? Hr2'] ->]". - destruct vl1 as [|[[|lr1|]|] []], vl2 as [|[[|lr2|]|] []]=>//=. - rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton. - iDestruct "Hr" as "[Hr1 Hr2]". wp_read. wp_let. wp_op. wp_read. wp_let. - wp_apply (wp_delete _ _ _ [_; _] with "[Hr1 Hr2 Hr†]")=>//. - { rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton /= freeable_sz_full. - iFrame. } - iIntros "_". - iMod (lft_incl_acc with "Hαβ Hα") as (?) "[Hβ Hβclose]". done. - iMod (na_bor_acc with "LFT Hinv Hβ Hna") as "(INV & Hna & Hclosena)"; [done..|]. - wp_seq. wp_op. wp_read. - iDestruct "INV" as (st) "(Hlrc & H● & Hst)". - iDestruct (own_valid_2 with "H● Hγ") as %[Hst _]%auth_both_valid_discrete. - destruct st as [[[[??]?]?]|]; last by destruct Hst as [[?|] Hst]; inversion Hst. - move:Hst=>/Some_pair_included [/Some_pair_included_total_1 - [/to_agree_included /(leibniz_equiv _ _) [= <- <-] _] _]. - iDestruct "Hst" as "(H† & Hq & _)". iDestruct "Hq" as (q1 Hqq1) "[Hν1 Hν2]". - iMod (own_update with "H●") as "[H● ?]". - { apply auth_update_alloc, - (op_local_update_discrete _ _ (writing_stR (q1/2)%Qp ν))=>-[Hagv _]. - split; [split|done]. - - by rewrite /= agree_idemp. - - apply frac_valid'. rewrite /= -Hqq1 comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. } - wp_let. wp_read. wp_let. wp_op. wp_write. - wp_apply (wp_delete _ _ _ [_; _] with "[Hrefmut↦1 Hrefmut↦2 Hrefmut†]")=>//. - { rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton /= freeable_sz_full. - iFrame. } - iIntros "_". wp_seq. wp_apply wp_new; [done..|]. iIntros (lrefmuts) "[Hrefmuts† Hrefmuts]". - rewrite 3!heap_mapsto_vec_cons heap_mapsto_vec_singleton. wp_let. - iDestruct "Hrefmuts" as "(Hrefmuts1 & Hrefmuts2 & Hrefmuts3 & Hrefmuts4)". - rewrite !shift_loc_assoc. wp_write. do 3 (wp_op; wp_write). - iMod ("Hclosena" with "[Hlrc H● Hν1 H†] Hna") as "[Hβ Hna]". - { iExists (Some (_, true, _, _)). - rewrite -Some_op -!pair_op agree_idemp /= (comm _ xH _). - iFrame. iSplitL; [|done]. iExists _. iFrame. - rewrite (comm Qp_add) (assoc Qp_add) Qp_div_2 (comm Qp_add). auto. } - iMod ("Hβclose" with "Hβ") as "Hα". iMod ("Hclose1" with "Hα HL") as "HL". - iApply (type_type _ [_] _ [ #lrefmuts ◁ box (Π[refmut α ty1; refmut α ty2]) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_singleton tctx_hasty_val' //= -freeable_sz_full. - simpl. iFrame. iExists [_;_;_;_]. - rewrite 3!heap_mapsto_vec_cons heap_mapsto_vec_singleton !shift_loc_assoc. - iFrame. iExists [_;_], [_;_]. iSplit=>//=. - iSplitL "Hν Hγ Hr1'"; [by auto 10 with iFrame|]. iExists [_;_], []. - iSplitR=>//=. rewrite right_id. auto 20 with iFrame. } - iApply type_jump; solve_typing. - Qed. -End refmut_functions. diff --git a/theories/typing/lib/rwlock/rwlock.v b/theories/typing/lib/rwlock/rwlock.v deleted file mode 100644 index 1b48d479..00000000 --- a/theories/typing/lib/rwlock/rwlock.v +++ /dev/null @@ -1,226 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree excl numbers. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import at_borrow. -From lrust.typing Require Import typing. -Set Default Proof Using "Type". - -Definition rwlock_stR := - optionUR (csumR (exclR unitO) (prodR (prodR (agreeR lftO) fracR) positiveR)). -Class rwlockG Σ := - RwLockG :> inG Σ (authR rwlock_stR). -Definition rwlockΣ : gFunctors := #[GFunctor (authR rwlock_stR)]. -Global Instance subG_rwlockΣ {Σ} : subG rwlockΣ Σ → rwlockG Σ. -Proof. solve_inG. Qed. - -Definition Z_of_rwlock_st (st : rwlock_stR) : Z := - match st with - | None => 0 - | Some (Cinr (_, _, n)) => Zpos n - | Some _ => -1 - end. - -Definition reading_st (q : frac) (ν : lft) : rwlock_stR := - Some (Cinr (to_agree ν, q, xH)). -Definition writing_st : rwlock_stR := - Some (Cinl (Excl ())). - -Definition rwlockN := lrustN .@ "rwlock". - -Section rwlock_inv. - Context `{!typeG Σ, !rwlockG Σ}. - - Definition rwlock_inv tid_own tid_shr (l : loc) (γ : gname) (α : lft) ty - : iProp Σ := - (∃ st, l ↦ #(Z_of_rwlock_st st) ∗ own γ (● st) ∗ - match st return _ with - | None => - (* Not locked. *) - &{α}((l +ₗ 1) ↦∗: ty.(ty_own) tid_own) - | Some (Cinr (agν, q, n)) => - (* Locked for read. *) - ∃ (ν : lft) q', agν ≡ to_agree ν ∗ - □ (1.[ν] ={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=∗ [†ν]) ∗ - ([†ν] ={↑lftN}=∗ &{α}((l +ₗ 1) ↦∗: ty.(ty_own) tid_own)) ∗ - ty.(ty_shr) (α ⊓ ν) tid_shr (l +ₗ 1) ∗ - ⌜(q + q')%Qp = 1%Qp⌝ ∗ q'.[ν] - | _ => (* Locked for write. *) True - end)%I. - - Lemma rwlock_inv_proper E L ty1 ty2 q : - eqtype E L ty1 ty2 → - llctx_interp L q -∗ ∀ tid_own tid_shr l γ α, □ (elctx_interp E -∗ - rwlock_inv tid_own tid_shr l γ α ty1 -∗ - rwlock_inv tid_own tid_shr l γ α ty2). - Proof. - (* TODO : this proof is essentially [solve_proper], but within the logic. - It would easily gain from some automation. *) - rewrite eqtype_unfold. iIntros (Hty) "HL". - iDestruct (Hty with "HL") as "#Hty". iIntros "* !> #HE H". - iDestruct ("Hty" with "HE") as "(% & #Hl & #Hown & #Hshr)". - iAssert (□ (&{α}((l +ₗ 1) ↦∗: ty_own ty1 tid_own) -∗ - &{α}((l +ₗ 1) ↦∗: ty_own ty2 tid_own)))%I as "#Hb". - { iIntros "!> H". iApply bor_iff; last done. - iNext; iModIntro; iSplit; iIntros "H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl; - iFrame; by iApply "Hown". } - iDestruct "H" as (st) "H"; iExists st; - iDestruct "H" as "($&$&H)"; destruct st as [[|[[agν ?]?]|]|]; try done; - last by iApply "Hb". - iDestruct "H" as (ν q') "(Hag & #Hend & Hh & ? & ? & ?)". iExists ν, q'. - iFrame. iSplitR. done. iSplitL "Hh"; last by iApply "Hshr". - iIntros "Hν". iApply "Hb". iApply ("Hh" with "Hν"). - Qed. - - Lemma rwlock_inv_change_tid_own tid_own1 tid_own2 tid_shr l γ α ty : - Send ty → - rwlock_inv tid_own1 tid_shr l γ α ty ≡ rwlock_inv tid_own2 tid_shr l γ α ty. - Proof. - intros ?. apply bi.exist_proper=>?; do 7 f_equiv; first do 7 f_equiv. - - do 5 f_equiv. iApply send_change_tid'. - - iApply send_change_tid'. - Qed. - - Lemma rwlock_inv_change_tid_shr tid_own tid_shr1 tid_shr2 l γ α ty : - Sync ty → - rwlock_inv tid_own tid_shr1 l γ α ty ≡ rwlock_inv tid_own tid_shr2 l γ α ty. - Proof. - intros ?. apply bi.exist_proper=>?; do 7 f_equiv; first do 7 f_equiv. - iApply sync_change_tid'. - Qed. -End rwlock_inv. - -Section rwlock. - Context `{!typeG Σ, !rwlockG Σ}. - - (* Original Rust type (we ignore poisoning): - pub struct RwLock<T: ?Sized> { - inner: Box<sys::RWLock>, - poison: poison::Flag, - data: UnsafeCell<T>, - } - *) - - Program Definition rwlock (ty : type) := - {| ty_size := S ty.(ty_size); ty_lfts := ty.(ty_lfts); ty_E := ty.(ty_E); - ty_own tid vl := - match vl return _ with - | #(LitInt z) :: vl' => ⌜-1 ≤ z⌝ ∗ ty.(ty_own) tid vl' - | _ => False - end%I; - ty_shr κ tid l := - (∃ α γ, κ ⊑ α ∗ α ⊑ ty_lft ty ∗ - &at{α,rwlockN}(rwlock_inv tid tid l γ α ty))%I |}. - Next Obligation. - iIntros (??[|[[]|]]); try iIntros "[]". rewrite ty_size_eq. - iIntros "[_ %] !% /=". congruence. - Qed. - Next Obligation. - iIntros (ty E κ l tid q ?) "#LFT #? Hb Htok". - iMod (bor_acc_cons with "LFT Hb Htok") as "[H Hclose]". done. - iDestruct "H" as ([|[[| |n]|]vl]) "[H↦ H]"; try iDestruct "H" as ">[]". - iDestruct "H" as "[>% Hown]". - iMod ("Hclose" $! ((∃ n:Z, l ↦ #n ∗ ⌜-1 ≤ n⌝) ∗ - (l +ₗ 1) ↦∗: ty.(ty_own) tid) with "[] [-]")%I as "[H [Htok Htok']]". - { iIntros "!> [Hn Hvl] !>". iDestruct "Hn" as (n') "[Hn >%]". - iDestruct "Hvl" as (vl') "[H↦ Hvl]". - iExists (#n'::vl'). rewrite heap_mapsto_vec_cons. iFrame "∗%". } - { iNext. rewrite heap_mapsto_vec_cons. iDestruct "H↦" as "[Hn Hvl]". - iSplitL "Hn"; [eauto|iExists _; iFrame]. } - iMod (bor_sep with "LFT H") as "[Hn Hvl]". done. - iMod (bor_acc_cons with "LFT Hn Htok") as "[H Hclose]". done. - iAssert ((q / 2).[κ] ∗ ▷ ∃ γ, rwlock_inv tid tid l γ κ ty)%I with "[> -Hclose]" - as "[$ HQ]"; last first. - { iMod ("Hclose" with "[] HQ") as "[Hb $]". - - iIntros "!> H !>". iNext. iDestruct "H" as (γ st) "(H & _ & _)". - iExists _. iIntros "{$H}!%". destruct st as [[|[[]?]|]|]; simpl; lia. - - iMod (bor_exists with "LFT Hb") as (γ) "Hb". done. - iExists κ, γ. iSplitR; [by iApply lft_incl_refl|]. iSplitR; [done|]. - iApply bor_share; try done. solve_ndisj. } - clear dependent n. iDestruct "H" as ([|n|[]]) "[Hn >%]"; try lia. - - iFrame. iMod (own_alloc (● None)) as (γ) "Hst"; first by apply auth_auth_valid. - iExists γ, None. by iFrame. - - iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]". done. - iMod (own_alloc (● Some (Cinr (to_agree ν, (1/2)%Qp, n)))) as (γ) "Hst". - { by apply auth_auth_valid. } - iMod (rebor _ _ (κ ⊓ ν) with "LFT [] Hvl") as "[Hvl Hh]". done. - { iApply lft_intersect_incl_l. } - iDestruct (lft_intersect_acc with "Htok' Htok1") as (q') "[Htok Hclose]". - iMod (ty_share with "LFT [] Hvl Htok") as "[Hshr Htok]". done. - { iApply lft_incl_trans; [|done]. iApply lft_intersect_incl_l. } - iDestruct ("Hclose" with "Htok") as "[$ Htok]". - iExists γ, _. iFrame "Hst Hn". iExists _, _. iIntros "{$Hshr}". - iSplitR; first by auto. iFrame "Htok2". iSplitR; first done. - rewrite Qp_div_2. iSplitL; last by auto. - iIntros "!> !> Hν". iDestruct (lft_tok_dead with "Htok Hν") as "[]". - - iMod (own_alloc (● writing_st)) as (γ) "Hst". { by apply auth_auth_valid. } - iFrame. iExists _, _. eauto with iFrame. - Qed. - Next Obligation. - iIntros (?????) "#Hκ H". iDestruct "H" as (α γ) "[#??]". - iExists _, _. iFrame. iApply lft_incl_trans; auto. - Qed. - - Global Instance rwlock_type_ne : TypeNonExpansive rwlock. - Proof. - split. - - apply (type_lft_morphism_add _ static [] []) => ?. - + rewrite left_id. iApply lft_equiv_refl. - + by rewrite /elctx_interp /= left_id right_id. - - by move=>/= ?? ->. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid [|[[| |]|]vl]=>//=. by rewrite Ho. - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. rewrite /= /rwlock_inv. - repeat (apply dist_S, Ho || apply dist_S, Hs || - apply equiv_dist, lft_incl_equiv_proper_r, Hl || - f_contractive || f_equiv). - Qed. - - Global Instance rwlock_ne : NonExpansive rwlock. - Proof. - unfold rwlock, rwlock_inv. intros n ty1 ty2 Hty12. - split; simpl; try by rewrite Hty12. - - intros tid [|[[| |]|]vl]=>//=. by rewrite Hty12. - - intros κ tid l. repeat (apply Hty12 || f_equiv). - Qed. - - Global Instance rwlock_mono E L : Proper (eqtype E L ==> subtype E L) rwlock. - Proof. - (* TODO : this proof is essentially [solve_proper], but within the logic. - It would easily gain from some automation. *) - iIntros (ty1 ty2 EQ qL) "HL". generalize EQ. rewrite eqtype_unfold=>EQ'. - iDestruct (EQ' with "HL") as "#EQ'". - iDestruct (rwlock_inv_proper with "HL") as "#Hty1ty2"; first done. - iDestruct (rwlock_inv_proper with "HL") as "#Hty2ty1"; first by symmetry. - iIntros "!> #HE". iDestruct ("EQ'" with "HE") as "(% & #[??] & #Hown & #Hshr)". - iSplit; [|iSplit; [done|iSplit; iIntros "!> * H"]]. - - iPureIntro. simpl. congruence. - - destruct vl as [|[[]|]]; try done. iDestruct "H" as "[$ H]". by iApply "Hown". - - iDestruct "H" as (a γ) "[Ha #[??]]". iExists a, γ. iFrame. iSplit. - + by iApply lft_incl_trans. - + iApply at_bor_iff; last done. iNext; iModIntro; iSplit; iIntros "H". - by iApply "Hty1ty2". by iApply "Hty2ty1". - Qed. - Lemma rwlock_mono' E L ty1 ty2 : - eqtype E L ty1 ty2 → subtype E L (rwlock ty1) (rwlock ty2). - Proof. eapply rwlock_mono. Qed. - Global Instance rwlock_proper E L : Proper (eqtype E L ==> eqtype E L) rwlock. - Proof. by split; apply rwlock_mono. Qed. - Lemma rwlock_proper' E L ty1 ty2 : - eqtype E L ty1 ty2 → eqtype E L (rwlock ty1) (rwlock ty2). - Proof. eapply rwlock_proper. Qed. - - (* Apparently, we don't need to require ty to be sync, - contrarily to Rust's implementation. *) - Global Instance rwlock_send ty : - Send ty → Send (rwlock ty). - Proof. move=>???[|[[]|]]//=??. iIntros "[$?]". by iApply send_change_tid. Qed. - - Global Instance rwlock_sync ty : - Send ty → Sync ty → Sync (rwlock ty). - Proof. - move=>??????/=. do 2 apply bi.exist_mono=>?. apply bi.sep_mono_r. - apply bi.equiv_spec. f_equiv. - by rewrite rwlock_inv_change_tid_own rwlock_inv_change_tid_shr. - Qed. -End rwlock. - -Global Hint Resolve rwlock_mono' rwlock_proper' : lrust_typing. diff --git a/theories/typing/lib/rwlock/rwlock_code.v b/theories/typing/lib/rwlock/rwlock_code.v deleted file mode 100644 index b8ab832b..00000000 --- a/theories/typing/lib/rwlock/rwlock_code.v +++ /dev/null @@ -1,303 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lang.lib Require Import memcpy. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing option. -From lrust.typing.lib.rwlock Require Import rwlock rwlockreadguard rwlockwriteguard. -Set Default Proof Using "Type". - -Section rwlock_functions. - Context `{!typeG Σ, !rwlockG Σ}. - - (* Constructing a rwlock. *) - Definition rwlock_new ty : val := - fn: ["x"] := - let: "r" := new [ #(S ty.(ty_size))] in - "r" +ₗ #0 <- #0;; - "r" +ₗ #1 <-{ty.(ty_size)} !"x";; - delete [ #ty.(ty_size) ; "x"];; return: ["r"]. - - Lemma rwlock_new_type ty : - typed_val (rwlock_new ty) (fn(∅; ty) → rwlock ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ret ϝ arg). inv_vec arg=>x. simpl_subst. - iApply (type_new (S ty.(ty_size))); [solve_typing..|]. - iIntros (r tid) "#LFT HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done. - iDestruct "Hx" as "[Hx Hx†]". iDestruct "Hx" as (vl) "[Hx↦ Hx]". - destruct r as [[|lr|]|]; try done. iDestruct "Hr" as "[Hr Hr†]". - iDestruct "Hr" as (vl') "Hr". rewrite uninit_own. - iDestruct "Hr" as "[Hr↦ >SZ]". destruct vl' as [|]; iDestruct "SZ" as %[=]. - rewrite heap_mapsto_vec_cons. iDestruct "Hr↦" as "[Hr↦0 Hr↦1]". wp_op. - rewrite shift_loc_0. wp_write. wp_op. iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz. - wp_apply (wp_memcpy with "[$Hr↦1 $Hx↦]"); [done||lia..|]. - iIntros "[Hr↦1 Hx↦]". wp_seq. - iApply (type_type _ _ _ [ #lx ◁ box (uninit (ty_size ty)); #lr ◁ box (rwlock ty)] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //=. iFrame. - iSplitL "Hx↦". - - iExists _. rewrite uninit_own. auto. - - iExists (_::_). rewrite heap_mapsto_vec_cons. iFrame. simpl. iFrame. auto. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* The other direction: getting ownership out of a rwlock. - Note: as we ignore poisonning, this cannot fail. *) - Definition rwlock_into_inner ty : val := - fn: ["x"] := - let: "r" := new [ #ty.(ty_size)] in - "r" <-{ty.(ty_size)} !("x" +ₗ #1);; - delete [ #(S ty.(ty_size)) ; "x"];; return: ["r"]. - - Lemma rwlock_into_inner_type ty : - typed_val (rwlock_into_inner ty) (fn(∅; rwlock ty) → ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (_ ret ϝ arg). inv_vec arg=>x. simpl_subst. - iApply (type_new ty.(ty_size)); [solve_typing..|]. - iIntros (r tid) "#LFT HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done. - iDestruct "Hx" as "[Hx Hx†]". - iDestruct "Hx" as ([|[[]|]vl]) "[Hx↦ Hx]"; try iDestruct "Hx" as ">[]". - destruct r as [[|lr|]|]; try done. iDestruct "Hr" as "[Hr Hr†]". - iDestruct "Hr" as (vl') "Hr". rewrite uninit_own heap_mapsto_vec_cons. - iDestruct "Hr" as "[Hr↦ >%]". iDestruct "Hx↦" as "[Hx↦0 Hx↦1]". wp_op. - iDestruct "Hx" as "[% Hx]". iDestruct (ty.(ty_size_eq) with "Hx") as %Hsz. - wp_apply (wp_memcpy with "[$Hr↦ $Hx↦1]"); [done||lia..|]. - iIntros "[Hr↦ Hx↦1]". wp_seq. - iApply (type_type _ _ _ [ #lx ◁ box (uninit (S (ty_size ty))); #lr ◁ box ty] - with "[] LFT HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame. - iSplitR "Hr↦ Hx". - - iExists (_::_). rewrite heap_mapsto_vec_cons uninit_own -Hsz. iFrame. auto. - - iExists vl. iFrame. } - iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - Definition rwlock_get_mut : val := - fn: ["x"] := - let: "x'" := !"x" in - "x" <- "x'" +ₗ #1;; - return: ["x"]. - - Lemma rwlock_get_mut_type ty : - typed_val rwlock_get_mut (fn(∀ α, ∅; &uniq{α} (rwlock ty)) → &uniq{α} ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iIntros (tid) "#LFT HE Hna HL HC HT". - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx [#? Hx']]". - destruct x' as [[|lx'|]|]; try iDestruct "Hx'" as "[]". - iAssert (&{α} (∃ (z : Z), lx' ↦ #z ∗ ⌜-1 ≤ z⌝) ∗ - (&uniq{α} ty).(ty_own) tid [ #(lx' +ₗ 1)])%I with "[> Hx']" as "[_ Hx']". - { iFrame "#". iApply bor_sep; [done..|]. iApply (bor_proper with "Hx'"). iSplit. - - iIntros "[H1 H2]". iDestruct "H1" as (z) "[??]". iDestruct "H2" as (vl) "[??]". - iExists (_::_). rewrite heap_mapsto_vec_cons. iFrame. iFrame. - - iIntros "H". iDestruct "H" as ([|[[| |z]|]vl]) "[H↦ H]"; try done. - rewrite heap_mapsto_vec_cons. - iDestruct "H" as "[H1 H2]". iDestruct "H↦" as "[H↦1 H↦2]". - iSplitL "H1 H↦1"; eauto. iExists _. iFrame. } - destruct x as [[|lx|]|]; try done. iDestruct "Hx" as "[Hx Hx†]". - iDestruct "Hx" as (vl) "[Hx↦ Hx]". rewrite uninit_own. wp_op. - iApply (type_type _ _ _ - [ #lx ◁ box (uninit 1); #(lx' +ₗ 1) ◁ &uniq{α}ty] - with "[] LFT HE Hna HL HC [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //. iFrame. - iNext. iExists _. rewrite uninit_own. iFrame. } - iApply type_assign; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Acquiring a read lock. *) - Definition rwlock_try_read : val := - fn: ["x"] := - let: "r" := new [ #2 ] in - let: "x'" := !"x" in - withcont: "k": - withcont: "loop": - "loop" [] - cont: "loop" [] := - let: "n" := !ˢᶜ"x'" in - if: "n" ≤ #-1 then - "r" <-{Σ none} ();; - "k" [] - else - if: CAS "x'" "n" ("n" + #1) then - "r" <-{Σ some} "x'";; - "k" [] - else "loop" [] - cont: "k" [] := - delete [ #1; "x"];; return: ["r"]. - - Lemma rwlock_try_read_type ty : - typed_val rwlock_try_read - (fn(∀ α, ∅; &shr{α}(rwlock ty)) → option (rwlockreadguard α ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] (λ _, [x ◁ box (&shr{α}(rwlock ty)); - r ◁ box (option (rwlockreadguard α ty))])); - [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg]; - simpl_subst; last first. - { iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. } - iApply (type_cont [] [ϝ ⊑ₗ []] (λ _, [x ◁ _; x' ◁ _; r ◁ _])); - [iIntros (loop)|iIntros "/= !>"; iIntros (loop arg); inv_vec arg]; - simpl_subst. - { iApply type_jump; solve_typing. } - iIntros (tid) "#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αβ & Hβty & Hinv)". - 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βtok1 Hβtok2] Hclose']". done. - wp_bind (!ˢᶜ(LitV lx))%E. - iMod (at_bor_acc_tok with "LFT Hinv Hβtok1") as "[INV Hclose'']"; try done. - iDestruct "INV" as (st) "(Hlx & INV)". wp_read. - iMod ("Hclose''" with "[Hlx INV]") as "Hβtok1"; first by iExists _; iFrame. - iModIntro. wp_let. wp_op; case_bool_decide as Hm1; wp_if. - - iMod ("Hclose'" with "[$]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(rwlock ty)); r ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. } - iApply (type_sum_unit (option $ rwlockreadguard α ty)); - [solve_typing..|]; first last. - simpl. iApply type_jump; solve_typing. - - wp_op. wp_bind (CAS _ _ _). - iMod (at_bor_acc_tok with "LFT Hinv Hβtok1") as "[INV Hclose'']"; try done. - iDestruct "INV" as (st') "(Hlx & Hownst & Hst)". revert Hm1. - destruct (decide (Z_of_rwlock_st st = Z_of_rwlock_st st')) as [->|?]=>?. - + iApply (wp_cas_int_suc with "Hlx"). iNext. iIntros "Hlx". - iAssert (∃ qν ν, (qβ / 2).[β] ∗ (qν).[ν] ∗ - ty_shr ty (β ⊓ ν) tid (lx +ₗ 1) ∗ - own γ (◯ reading_st qν ν) ∗ rwlock_inv tid tid lx γ β ty ∗ - ((1).[ν] ={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=∗ [†ν]))%I - with "[> Hlx Hownst Hst Hβtok2]" - as (q' ν) "(Hβtok2 & Hν & Hshr & Hreading & INV & H†)". - { destruct st' as [[|[[agν q] n]|]|]; try done. - - iDestruct "Hst" as (ν q') "(Hag & #H† & Hh & #Hshr & #Hqq' & [Hν1 Hν2])". - iExists _, _. iFrame "Hν1". iDestruct "Hag" as %Hag. iDestruct "Hqq'" as %Hqq'. - iMod (own_update with "Hownst") as "[Hownst ?]". - { apply auth_update_alloc, - (op_local_update_discrete _ _ (reading_st (q'/2)%Qp ν))=>-[Hagv _]. - split; [split|]. - - by rewrite /= Hag agree_idemp. - - apply frac_valid'. rewrite /= -Hqq' comm_L. - by apply Qp_add_le_mono_l, Qp_div_le. - - done. } - iFrame "∗#". iExists _. rewrite Z.add_comm /=. iFrame. iExists _, _. - iFrame "∗#". iSplitR; first by rewrite /= Hag agree_idemp. - rewrite (comm Qp_add) (assoc Qp_add) Qp_div_2 (comm Qp_add). auto. - - iMod (lft_create with "LFT") as (ν) "[[Htok1 Htok2] #Hhν]". solve_ndisj. - iMod (own_update with "Hownst") as "[Hownst Hreading]"; first by apply - auth_update_alloc, (op_local_update_discrete _ _ (reading_st (1/2)%Qp ν)). - rewrite right_id. iExists _, _. iFrame "Htok1 Hreading". - iDestruct (lft_intersect_acc with "Hβtok2 Htok2") as (q) "[Htok Hclose]". - iApply (fupd_mask_mono (↑lftN)). solve_ndisj. - iMod (rebor _ _ (β ⊓ ν) with "LFT [] Hst") as "[Hst Hh]". solve_ndisj. - { iApply lft_intersect_incl_l. } - iMod (ty_share with "LFT [] Hst Htok") as "[#Hshr Htok]". solve_ndisj. - { iApply lft_incl_trans; [|done]. iApply lft_intersect_incl_l. } - iFrame "#". iDestruct ("Hclose" with "Htok") as "[$ Htok2]". - iExists _. iFrame. iExists _, _. iSplitR; first done. iFrame "#∗". - rewrite Qp_div_2. iSplitL; last done. - iIntros "!> Hν". iApply "Hh". rewrite -lft_dead_or. auto. } - iMod ("Hclose''" with "[$INV]") as "Hβtok1". iModIntro. wp_case. - iMod ("Hclose'" with "[$]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(rwlock ty)); r ◁ box (uninit 2); - #lx ◁ rwlockreadguard α ty] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. iFrame. - iExists _, _, _, _, _. iFrame "∗#". iApply ty_shr_mono; last done. - iApply lft_intersect_mono; first done. iApply lft_incl_refl. } - iApply (type_sum_assign (option $ rwlockreadguard α ty)); [solve_typing..|]. - simpl. iApply type_jump; solve_typing. - + iApply (wp_cas_int_fail with "Hlx"); try done. iNext. iIntros "Hlx". - iMod ("Hclose''" with "[Hlx Hownst Hst]") as "Hβtok1"; first by iExists _; iFrame. - iModIntro. wp_case. iMod ("Hclose'" with "[$]") as "Hα". - iMod ("Hclose" with "Hα HL") as "HL". - iSpecialize ("Hk" with "[]"); first solve_typing. - iApply ("Hk" $! [#] with "Hna HL"). - rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. - iExists _. iSplit. done. simpl. eauto. - Qed. - - (* Acquiring a write lock. *) - Definition rwlock_try_write : val := - fn: ["x"] := - withcont: "k": - let: "r" := new [ #2 ] in - let: "x'" := !"x" in - if: CAS "x'" #0 #(-1) then - "r" <-{Σ some} "x'";; - "k" ["r"] - else - "r" <-{Σ none} ();; - "k" ["r"] - cont: "k" ["r"] := - delete [ #1; "x"];; return: ["r"]. - - Lemma rwlock_try_write_type ty : - typed_val rwlock_try_write - (fn(∀ α, ∅; &shr{α}(rwlock ty)) → option (rwlockwriteguard α ty)). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply (type_cont [_] [ϝ ⊑ₗ []] (λ r, [x ◁ box (&shr{α}(rwlock ty)); - (r!!!0%fin:val) ◁ box (option (rwlockwriteguard α ty))])); - [iIntros (k)|iIntros "/= !>"; iIntros (k arg); inv_vec arg=>r]; - simpl_subst; last first. - { iApply type_delete; [solve_typing..|]. iApply type_jump; solve_typing. } - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_deref; [solve_typing..|]. - iIntros (x' tid) "#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αβ & Hβty & Hinv)". - 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βtok Hclose']". done. - wp_bind (CAS _ _ _). - iMod (at_bor_acc_tok with "LFT Hinv Hβtok") as "[INV Hclose'']"; try done. - iDestruct "INV" as (st) "(Hlx & >Hownst & Hst)". destruct st. - - iApply (wp_cas_int_fail with "Hlx"). by destruct c as [|[[]]|]. - iNext. iIntros "Hlx". - iMod ("Hclose''" with "[Hlx Hownst Hst]") as "Hβtok"; first by iExists _; iFrame. - iMod ("Hclose'" with "Hβtok") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iModIntro. wp_case. - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(rwlock ty)); r ◁ box (uninit 2) ] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. } - iApply (type_sum_unit (option $ rwlockwriteguard α ty)); - [solve_typing..|]; first last. - rewrite /option /=. iApply type_jump; solve_typing. - - iApply (wp_cas_int_suc with "Hlx"). iIntros "!> Hlx". - iMod (own_update with "Hownst") as "[Hownst ?]". - { by eapply auth_update_alloc, (op_local_update_discrete _ _ writing_st). } - iMod ("Hclose''" with "[Hlx Hownst]") as "Hβtok". { iExists _. iFrame. } - iModIntro. wp_case. iMod ("Hclose'" with "Hβtok") as "Hα". - iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ - [ x ◁ box (&shr{α}(rwlock ty)); r ◁ box (uninit 2); - #lx ◁ rwlockwriteguard α ty] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val - tctx_hasty_val' //. iFrame. iExists _, _, _. iFrame "∗#". } - iApply (type_sum_assign (option $ rwlockwriteguard α ty)); [solve_typing..|]. - simpl. iApply type_jump; solve_typing. - Qed. -End rwlock_functions. diff --git a/theories/typing/lib/rwlock/rwlockreadguard.v b/theories/typing/lib/rwlock/rwlockreadguard.v deleted file mode 100644 index 83b652e1..00000000 --- a/theories/typing/lib/rwlock/rwlockreadguard.v +++ /dev/null @@ -1,152 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.rwlock Require Import rwlock. -Set Default Proof Using "Type". - -Section rwlockreadguard. - Context `{!typeG Σ, !rwlockG Σ}. - - (* Original Rust type: - pub struct RwLockReadGuard<'a, T: ?Sized + 'a> { - __lock: &'a RwLock<T>, - } - *) - - Program Definition rwlockreadguard (α : lft) (ty : type) := - {| ty_size := 1; - ty_lfts := α :: ty.(ty_lfts); ty_E := ty.(ty_E) ++ ty_outlives_E ty α; - ty_own tid vl := - match vl return _ with - | [ #(LitLoc l) ] => - ∃ ν q γ β tid_own, ty.(ty_shr) (α ⊓ ν) tid (l +ₗ 1) ∗ - α ⊑ β ∗ &at{β,rwlockN}(rwlock_inv tid_own tid l γ β ty) ∗ - q.[ν] ∗ own γ (◯ reading_st q ν) ∗ - (1.[ν] ={↑lftN ∪ ↑lft_userN}[↑lft_userN]▷=∗ [†ν]) - | _ => False - end; - ty_shr κ tid l := - ∃ (l' : loc), - &frac{κ} (λ q, l↦∗{q} [ #l']) ∗ - ▷ ty.(ty_shr) (α ⊓ κ) tid (l' +ₗ 1) |}%I. - Next Obligation. intros α ty tid [|[[]|] []]; auto. Qed. - Next Obligation. - iIntros (α ty E κ l tid q ?) "#LFT #Hl Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb". done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]". done. - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦". done. - destruct vl as [|[[|l'|]|][]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - iMod (bor_exists with "LFT Hb") as (ν) "Hb". done. - iMod (bor_exists with "LFT Hb") as (q') "Hb". done. - iMod (bor_exists with "LFT Hb") as (γ) "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. - iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ Htok]". done. - iMod (bor_sep with "LFT Hb") as "[Hinv Hb]". done. - iMod (bor_persistent with "LFT Hinv Htok") as "[#Hinv $]". done. - iMod (bor_sep with "LFT Hb") as "[Hκν _]". done. - iDestruct (frac_bor_lft_incl with "LFT [> Hκν]") as "#Hκν". - { iApply bor_fracture; try done. by rewrite Qp_mul_1_r. } - iExists _. iFrame "#". iApply ty_shr_mono; last done. - iApply lft_intersect_mono; last done. iApply lft_incl_refl. - Qed. - Next Obligation. - iIntros (α ty κ κ' tid l) "#Hκ'κ H". iDestruct "H" as (l') "[#Hf #Hshr]". - iExists l'. iSplit; first by iApply frac_bor_shorten. - iApply ty_shr_mono; last done. iApply lft_intersect_mono; last done. - iApply lft_incl_refl. - Qed. - - Global Instance rwlockreadguard_type_contractive α : TypeContractive (rwlockreadguard α). - Proof. - split. - - apply (type_lft_morphism_add _ α [α] []) => ?. - + iApply lft_equiv_refl. - + by rewrite elctx_interp_app elctx_interp_ty_outlives_E - /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid [|[[]|][]]=>//=. unfold rwlock_inv. - repeat (apply Hs || apply dist_S, Hs || apply dist_S, Ho || - f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. simpl. by setoid_rewrite Hs. - Qed. - - Global Instance rwlockreadguard_ne α : NonExpansive (rwlockreadguard α). - Proof. - unfold rwlockreadguard, rwlock_inv. intros n ty1 ty2 Hty12. - split=>//=; try by rewrite Hty12. - - intros tid [|[[]|][]]=>//=. repeat (apply Hty12 || f_equiv). - - intros κ tid l. repeat (apply Hty12 || f_equiv). - Qed. - - (* The rust type is not covariant, although it probably could (cf. refcell). - This would require changing the definition of the type, though. *) - Global Instance rwlockreadguard_mono E L : - Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockreadguard. - Proof. - iIntros (α1 α2 Hα ty1 ty2 Hty q) "HL". - iDestruct (proj1 Hty with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". - iDestruct (rwlock_inv_proper with "HL") as "#Hty1ty2"; first done. - iDestruct (rwlock_inv_proper with "HL") as "#Hty2ty1"; first by symmetry. - iIntros "!> #HE". iDestruct ("Hα" with "HE") as "Hα1α2". - iDestruct ("Hty" with "HE") as "(%&#?&#Ho&#Hs)". - iSplit; [done|iSplit; [|iSplit; iModIntro]]. - - by iApply lft_intersect_mono. - - iIntros (tid [|[[]|][]]) "H"; try done. - iDestruct "H" as (ν q' γ β tid_own) "(#Hshr & #H⊑ & #Hinv & Htok & Hown)". - iExists ν, q', γ, β, tid_own. iFrame "∗#". iSplit; last iSplit. - + iApply ty_shr_mono; last by iApply "Hs". - iApply lft_intersect_mono. done. iApply lft_incl_refl. - + by iApply lft_incl_trans. - + iApply (at_bor_iff with "[] Hinv"). - iIntros "!> !>"; iSplit; iIntros "H". by iApply "Hty1ty2". by iApply "Hty2ty1". - - iIntros (κ tid l) "H". iDestruct "H" as (l') "[Hf Hshr]". iExists l'. - iFrame. iApply ty_shr_mono; last by iApply "Hs". - iApply lft_intersect_mono. done. iApply lft_incl_refl. - Qed. - Global Instance rwlockreadguard_mono_flip E L : - Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) - rwlockreadguard. - Proof. intros ??????. by apply rwlockreadguard_mono. Qed. - Lemma rwlockreadguard_mono' E L α1 α2 ty1 ty2 : - lctx_lft_incl E L α2 α1 → eqtype E L ty1 ty2 → - subtype E L (rwlockreadguard α1 ty1) (rwlockreadguard α2 ty2). - Proof. intros. by eapply rwlockreadguard_mono. Qed. - Global Instance rwlockreadguard_proper E L : - Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) rwlockreadguard. - Proof. intros ??[]?? EQ. split; apply rwlockreadguard_mono'; try done; apply EQ. Qed. - Lemma rwlockreadguard_proper' E L α1 α2 ty1 ty2 : - lctx_lft_eq E L α1 α2 → eqtype E L ty1 ty2 → - eqtype E L (rwlockreadguard α1 ty1) (rwlockreadguard α2 ty2). - Proof. intros. by eapply rwlockreadguard_proper. Qed. - - (* Rust requires the type to also be Send. Not sure why. *) - Global Instance rwlockreadguard_sync α ty : - Sync ty → Sync (rwlockreadguard α ty). - Proof. - move=>?????/=. apply bi.exist_mono=>?. by rewrite sync_change_tid. - Qed. - - (* POSIX requires the unlock to occur from the thread that acquired - the lock, so Rust does not implement Send for RwLockWriteGuard. We can - prove this for our spinlock implementation, however. *) - Global Instance rwlockreadguard_send α ty : - Sync ty → Send (rwlockreadguard α ty). - Proof. - iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H". - iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_spec. - repeat lazymatch goal with - | |- (ty_shr _ _ _ _) ≡ (ty_shr _ _ _ _) => by apply sync_change_tid' - | |- (rwlock_inv _ _ _ _ _ _) ≡ _ => by apply rwlock_inv_change_tid_shr - | |- _ => f_equiv - end. - Qed. -End rwlockreadguard. - -Global Hint Resolve rwlockreadguard_mono' rwlockreadguard_proper' : lrust_typing. diff --git a/theories/typing/lib/rwlock/rwlockreadguard_code.v b/theories/typing/lib/rwlock/rwlockreadguard_code.v deleted file mode 100644 index acd7ae81..00000000 --- a/theories/typing/lib/rwlock/rwlockreadguard_code.v +++ /dev/null @@ -1,148 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree numbers. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import lifetime na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.rwlock Require Import rwlock rwlockreadguard. -Set Default Proof Using "Type". - -Section rwlockreadguard_functions. - Context `{!typeG Σ, !rwlockG Σ}. - - (* Turning a rwlockreadguard into a shared borrow. *) - Definition rwlockreadguard_deref : val := - fn: ["x"] := - let: "x'" := !"x" in - let: "r'" := !"x'" +ₗ #1 in - letalloc: "r" <- "r'" in - delete [ #1; "x"];; return: ["r"]. - - Lemma rwlockreadguard_deref_type ty : - typed_val rwlockreadguard_deref - (fn(∀ '(α, β), ∅; &shr{α}(rwlockreadguard β ty)) → &shr{α} ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. - iDestruct "Hx'" as (l') "#[Hfrac Hshr]". - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. - 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..|]. - 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. - by iApply lft_incl_refl. } - iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. - iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Dropping a rwlockreadguard and releasing the corresponding lock. *) - Definition rwlockreadguard_drop : val := - fn: ["x"] := - let: "x'" := !"x" in - withcont: "loop": - "loop" [] - cont: "loop" [] := - let: "n" := !ˢᶜ"x'" in - if: CAS "x'" "n" ("n" - #1) then - delete [ #1; "x"];; - let: "r" := new [ #0] in return: ["r"] - else "loop" []. - - Lemma rwlockreadguard_drop_type ty : - typed_val rwlockreadguard_drop (fn(∀ α, ∅; rwlockreadguard α ty) → unit). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iApply (type_cont [] [ϝ ⊑ₗ []] (λ _, [x ◁ _; x' ◁ _])); - [iIntros (loop)|iIntros "/= !>"; iIntros (loop arg); inv_vec arg]; - simpl_subst. - { iApply type_jump; solve_typing. } - iIntros (tid) "#LFT #HE Hna HL Hk HT". - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx Hx']". - destruct x' as [[|lx'|]|]; try done. simpl. - iDestruct "Hx'" as (ν q γ β tid_own) "(Hx' & #Hαβ & #Hinv & Hν & H◯ & H†)". - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. - iMod (lft_incl_acc with "Hαβ Hα") as (qβ) "[Hβ Hcloseα]". done. - wp_bind (!ˢᶜ#lx')%E. - iMod (at_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|]. - iDestruct "INV" as (st) "[H↦ INV]". wp_read. - iMod ("Hcloseβ" with "[H↦ INV]") as "Hβ"; first by iExists _; iFrame. - iModIntro. wp_let. wp_op. wp_bind (CAS _ _ _). - iMod (at_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|]. - iDestruct "INV" as (st') "(Hlx & >H● & Hst)". - destruct (decide (Z_of_rwlock_st st = Z_of_rwlock_st st')) as [->|?]. - + iAssert (|={⊤ ∖ ↑rwlockN}[⊤ ∖ ↑rwlockN ∖ ↑lftN ∖ ↑lft_userN]▷=> - (lx' ↦ #(Z_of_rwlock_st st'-1) ==∗ rwlock_inv tid_own tid lx' γ β ty))%I - with "[H● H◯ Hx' Hν Hst H†]" as "INV". - { iDestruct (own_valid_2 with "H● H◯") as %[[[=]| (? & st0 & [=<-] & -> & [Heq|Hle])] - %option_included Hv]%auth_both_valid_discrete. - - destruct st0 as [|[[agν q']n]|]; try by inversion Heq. revert Heq. simpl. - intros [[EQ <-%leibniz_equiv]%(inj2 pair) <-%leibniz_equiv] - %(inj Cinr)%(inj2 pair). - iDestruct "Hst" as (ν' q') "(>EQν & _ & Hh & _ & >Hq & >Hν')". - rewrite -EQ. iDestruct "EQν" as %<-%(inj to_agree)%leibniz_equiv. - iCombine "Hν" "Hν'" as "Hν". iDestruct "Hq" as %->. - iApply (step_fupd_mask_mono ((↑lftN ∪ ↑lft_userN) ∪ (⊤ ∖ ↑rwlockN ∖ ↑lftN ∖ ↑lft_userN))); - last iApply (step_fupd_mask_frame_r _ (↑lft_userN)). - { set_solver-. } - { solve_ndisj. } - { rewrite difference_difference. apply: disjoint_difference_r1. done. } - { (* FIXME [solve_ndisj] fails. *) - apply: disjoint_difference_r1. done. } - iMod ("H†" with "Hν") as "H†". iModIntro. iNext. iMod "H†". - iMod fupd_intro_mask' as "Hclose"; last iMod ("Hh" with "H†") as "Hb". - { set_solver-. } - iMod "Hclose" as "_". iIntros "!> Hlx". iExists None. iFrame. - iApply (own_update_2 with "H● H◯"). apply auth_update_dealloc. - rewrite -(right_id None op (Some _)). apply cancel_local_update_unit, _. - - iApply step_fupd_intro. set_solver. iNext. iIntros "Hlx". - apply csum_included in Hle. - destruct Hle as [|[(?&?&[=]&?)|(?&[[agν q']n]&[=<-]&->&Hle%prod_included)]]; - [by subst|]. - destruct Hle as [[Hag [q0 Hq]]%prod_included Hn%pos_included]. - iDestruct "Hst" as (ν' q'') "(EQν & H†' & Hh & Hshr & Hq & Hν')". - iDestruct "EQν" as %EQν. revert Hag Hq. rewrite /= EQν to_agree_included. - intros <-%leibniz_equiv ->%leibniz_equiv. - iExists (Some (Cinr (to_agree ν, q0, Pos.predl n))). - iSplitL "Hlx"; first by destruct n. - replace (q ⋅ q0 + q'')%Qp with (q0 + (q + q''))%Qp by - by rewrite (comm _ q q0) assoc. iCombine "Hν" "Hν'" as "Hν". - iSplitL "H● H◯"; last by eauto with iFrame. - iApply (own_update_2 with "H● H◯"). apply auth_update_dealloc. - assert (n = 1%positive ⋅ Pos.predl n) as EQn. - { rewrite pos_op_plus -Pplus_one_succ_l Pos.succ_pred // =>?. by subst. } - rewrite {1}EQn -{1}(agree_idemp (to_agree _)) 2!pair_op Cinr_op Some_op. - apply (cancel_local_update_unit (reading_st q ν)) , _. } - iApply (wp_step_fupd with "INV"). set_solver. - iApply (wp_cas_int_suc with "Hlx"); try done. iNext. iIntros "Hlx INV !>". - iMod ("INV" with "Hlx") as "INV". iMod ("Hcloseβ" with "[$INV]") as "Hβ". - iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iModIntro. wp_if. - iApply (type_type _ _ _ [ x ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_singleton tctx_hasty_val //. } - iApply type_delete; [solve_typing..|]. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_jump; solve_typing. - + iApply (wp_cas_int_fail with "Hlx"); try done. iNext. iIntros "Hlx". - iMod ("Hcloseβ" with "[Hlx H● Hst]") as "Hβ". by iExists _; iFrame. - iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". - iModIntro. wp_if. - iApply (type_type _ _ _ [ x ◁ box (uninit 1); #lx' ◁ rwlockreadguard α ty] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val - tctx_hasty_val' //=; simpl. auto 10 with iFrame. } - iApply type_jump; solve_typing. - Qed. -End rwlockreadguard_functions. diff --git a/theories/typing/lib/rwlock/rwlockwriteguard.v b/theories/typing/lib/rwlock/rwlockwriteguard.v deleted file mode 100644 index 69620a05..00000000 --- a/theories/typing/lib/rwlock/rwlockwriteguard.v +++ /dev/null @@ -1,162 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.rwlock Require Import rwlock. -Set Default Proof Using "Type". - -Section rwlockwriteguard. - Context `{!typeG Σ, !rwlockG Σ}. - - (* Original Rust type (we ignore poisoning): - pub struct RwLockWriteGuard<'a, T: ?Sized + 'a> { - __lock: &'a RwLock<T>, - __poison: poison::Guard, - } - *) - - Program Definition rwlockwriteguard (α : lft) (ty : type) := - {| ty_size := 1; - ty_lfts := α :: ty.(ty_lfts); ty_E := ty.(ty_E) ++ ty_outlives_E ty α; - ty_own tid vl := - match vl return _ with - | [ #(LitLoc l) ] => - ∃ γ β tid_shr, &{β}((l +ₗ 1) ↦∗: ty.(ty_own) tid) ∗ - α ⊑ β ∗ β ⊑ ty_lft ty ∗ - &at{β,rwlockN}(rwlock_inv tid tid_shr l γ β ty) ∗ - own γ (◯ writing_st) - | _ => False - end; - ty_shr κ tid l := - ∃ (l' : loc), - &frac{κ} (λ q, l↦∗{q} [ #l']) ∗ - □ ∀ F q, ⌜↑shrN ∪ ↑lftN ⊆ F⌝ -∗ q.[α ⊓ κ] ={F}[F∖↑shrN]▷=∗ - ty.(ty_shr) (α ⊓ κ) tid (l' +ₗ 1) ∗ q.[α ⊓ κ] |}%I. - Next Obligation. by iIntros (???[|[[]|][]]) "?". Qed. - Next Obligation. - iIntros (α ty E κ l tid q HE) "#LFT #Hl Hb Htok". - iMod (bor_exists with "LFT Hb") as (vl) "Hb". done. - iMod (bor_sep with "LFT Hb") as "[H↦ Hb]". done. - iMod (bor_fracture (λ q, l ↦∗{q} vl)%I with "LFT H↦") as "#H↦". done. - destruct vl as [|[[|l'|]|][]]; - try by iMod (bor_persistent with "LFT Hb Htok") as "[>[] _]". - iMod (bor_exists with "LFT Hb") as (γ) "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αβ H]". done. - iMod (bor_sep with "LFT H") as "[Hβty _]". done. - iMod (bor_persistent with "LFT Hαβ Htok") as "[#Hαβ Htok]". done. - iMod (bor_persistent with "LFT Hβty Htok") as "[#Hβty $]". done. - iExists _. iFrame "H↦". iApply delay_sharing_nested=>//. - - iNext. iApply lft_incl_trans; [|done]. iApply lft_intersect_incl_l. - - iApply bor_shorten; [|done]. iApply lft_intersect_incl_r. - Qed. - Next Obligation. - iIntros (??????) "#? H". iDestruct "H" as (l') "[#Hf #H]". - iExists _. iSplit. - - by iApply frac_bor_shorten. - - iIntros "!> * % Htok". - iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. iApply lft_incl_refl. done. } - 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. iApply lft_incl_refl. done. - Qed. - - Global Instance rwlockwriteguard_type_contractive α : - TypeContractive (rwlockwriteguard α). - Proof. - split. - - apply (type_lft_morphism_add _ α [α] []) => ?. - + iApply lft_equiv_refl. - + by rewrite elctx_interp_app elctx_interp_ty_outlives_E - /elctx_interp /= left_id right_id. - - done. - - intros n ty1 ty2 Hsz Hl HE Ho Hs tid [|[[]|][]]=>//=. unfold rwlock_inv. - repeat (apply dist_S, Hs || apply dist_S, Ho || apply Ho || - apply equiv_dist, lft_incl_equiv_proper_r, Hl || - f_contractive || f_equiv). - - intros n ty1 ty2 Hsz Hl HE Ho Hs κ tid l. simpl. - repeat (apply Hs || f_contractive || f_equiv). - Qed. - - Global Instance rwlockwriteguard_ne α : NonExpansive (rwlockwriteguard α). - Proof. - unfold rwlockwriteguard, rwlock_inv. intros n ty1 ty2 Hty12. - split=>//=; try by rewrite Hty12. - - intros tid [|[[]|][]]=>//=. repeat (apply Hty12 || f_equiv). - - intros κ tid l. repeat (apply Hty12 || f_equiv). - Qed. - - Global Instance rwlockwriteguard_mono E L : - Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockwriteguard. - Proof. - intros α1 α2 Hα ty1 ty2 Hty. generalize Hty. rewrite eqtype_unfold. iIntros (Hty' q) "HL". - iDestruct (Hty' with "HL") as "#Hty". iDestruct (Hα with "HL") as "#Hα". - iDestruct (rwlock_inv_proper with "HL") as "#Hty1ty2"; first done. - iDestruct (rwlock_inv_proper with "HL") as "#Hty2ty1"; first by symmetry. - iIntros "!> #HE". iDestruct ("Hα" with "HE") as "Hα1α2". - iDestruct ("Hty" with "HE") as "(%&#[??]&#Ho&#Hs)". - iSplit; [done|iSplit; [|iSplit; iModIntro]]. - - by iApply lft_intersect_mono. - - iIntros (tid [|[[]|][]]) "H"; try done. - iDestruct "H" as (γ β tid_shr) "(Hb & #H⊑ & #Hβty & #Hinv & Hown)". - iExists γ, β, tid_shr. iFrame "∗#". iSplit; [|iSplit; [|iSplit]]. - + 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. - + by iApply lft_incl_trans. - + iApply at_bor_iff; try done. - iIntros "!>!>"; iSplit; iIntros "H". by iApply "Hty1ty2". by iApply "Hty2ty1". - - iIntros (κ tid l) "H". iDestruct "H" as (l') "H". iExists l'. - iDestruct "H" as "[$ #H]". iIntros "!> * % Htok". - iMod (lft_incl_acc with "[] Htok") as (q') "[Htok Hclose]"; first solve_ndisj. - { iApply lft_intersect_mono. done. 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". - Qed. - Global Instance rwlockwriteguard_mono_flip E L : - Proper (lctx_lft_incl E L ==> eqtype E L ==> flip (subtype E L)) rwlockwriteguard. - Proof. intros ??????. by apply rwlockwriteguard_mono. Qed. - Lemma rwlockwriteguard_mono' E L α1 α2 ty1 ty2 : - lctx_lft_incl E L α2 α1 → eqtype E L ty1 ty2 → - subtype E L (rwlockwriteguard α1 ty1) (rwlockwriteguard α2 ty2) . - Proof. intros. by eapply rwlockwriteguard_mono. Qed. - Global Instance rwlockwriteguard_proper E L : - Proper (lctx_lft_eq E L ==> eqtype E L ==> eqtype E L) rwlockwriteguard. - Proof. intros ??[]???. split; by apply rwlockwriteguard_mono'. Qed. - Lemma rwlockwriteguard_proper' E L α1 α2 ty1 ty2 : - lctx_lft_eq E L α1 α2 → eqtype E L ty1 ty2 → - eqtype E L (rwlockwriteguard α1 ty1) (rwlockwriteguard α2 ty2). - Proof. intros. by eapply rwlockwriteguard_proper. Qed. - - (* Rust requires the type to also be Send. Not sure why. *) - Global Instance rwlockwriteguard_sync α ty : - Sync ty → Sync (rwlockwriteguard α ty). - Proof. - move=>?????/=. apply bi.exist_mono=>?. do 7 f_equiv. - by rewrite sync_change_tid. - Qed. - - (* POSIX requires the unlock to occur from the thread that acquired - the lock, so Rust does not implement Send for RwLockWriteGuard. We can - prove this for our spinlock implementation, however. *) - Global Instance rwlockwriteguard_send α ty : - Send ty → Send (rwlockwriteguard α ty). - Proof. - iIntros (??? [|[[]|][]]) "H"; try done. simpl. iRevert "H". - iApply bi.exist_mono. iIntros (κ); simpl. apply bi.equiv_spec. - repeat lazymatch goal with - | |- (ty_own _ _ _) ≡ (ty_own _ _ _) => by apply send_change_tid' - | |- (rwlock_inv _ _ _ _ _ _) ≡ _ => by apply rwlock_inv_change_tid_own - | |- _ => f_equiv - end. - Qed. -End rwlockwriteguard. - -Global Hint Resolve rwlockwriteguard_mono' rwlockwriteguard_proper' : lrust_typing. diff --git a/theories/typing/lib/rwlock/rwlockwriteguard_code.v b/theories/typing/lib/rwlock/rwlockwriteguard_code.v deleted file mode 100644 index 9539be53..00000000 --- a/theories/typing/lib/rwlock/rwlockwriteguard_code.v +++ /dev/null @@ -1,148 +0,0 @@ -From iris.proofmode Require Import proofmode. -From iris.algebra Require Import auth csum frac agree. -From iris.bi Require Import fractional. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Import typing. -From lrust.typing.lib.rwlock Require Import rwlock rwlockwriteguard. -Set Default Proof Using "Type". - -Section rwlockwriteguard_functions. - Context `{!typeG Σ, !rwlockG Σ}. - - (* Turning a rwlockwriteguard into a shared borrow. *) - Definition rwlockwriteguard_deref : val := - fn: ["x"] := - let: "x'" := !"x" in - let: "r'" := !"x'" +ₗ #1 in - letalloc: "r" <- "r'" in - delete [ #1; "x"];; return: ["r"]. - - Lemma rwlockwriteguard_deref_type ty : - typed_val rwlockwriteguard_deref - (fn(∀ '(α, β), ∅; &shr{α}(rwlockwriteguard β ty)) → &shr{α} ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#LFT #HE Hna HL Hk HT". simpl_subst. - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx Hx']". destruct x' as [[|lx'|]|]; try done. - iDestruct "Hx'" as (l') "#[Hfrac Hshr]". - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "([Hα1 Hα2] & HL & Hclose)"; - [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')"; - [solve_typing..|]. - iDestruct (lft_intersect_acc with "Hβ Hα2") as (qβα) "[Hα2β Hcloseβα2]". - wp_bind (!(LitV lx'))%E. iApply (wp_step_fupd with "[Hα2β]"); - [|by iApply ("Hshr" with "[] Hα2β")|]; first done. - iMod "H↦" as "[H↦1 H↦2]". wp_read. iIntros "[#Hshr' Hα2β]!>". wp_op. wp_let. - iDestruct ("Hcloseβα2" with "Hα2β") as "[Hβ Hα2]". - iMod ("Hcloseα1" with "[$H↦1 $H↦2]") as "Hα1". iMod ("Hclose'" with "Hβ HL") as "HL". - iMod ("Hclose" with "[$] HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [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. } - iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|]. - iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Turning a rwlockwriteguard into a unique borrow. *) - Definition rwlockwriteguard_derefmut : val := - fn: ["x"] := - let: "x'" := !"x" in - let: "r'" := !"x'" +ₗ #1 in - letalloc: "r" <- "r'" in - delete [ #1; "x"];; return: ["r"]. - - Lemma rwlockwriteguard_derefmut_type ty : - typed_val rwlockwriteguard_derefmut - (fn(∀ '(α, β), ∅; &uniq{α}(rwlockwriteguard β ty)) → &uniq{α}ty). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). - iIntros (tid) "#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. - iMod (bor_sep with "LFT H") as "[H↦ H]". done. - iMod (lctx_lft_alive_tok α with "HE HL") as (qα) "(Hα & HL & Hclose)"; [solve_typing..|]. - destruct vl as [|[[|l|]|][]]; - try by iMod (bor_persistent with "LFT H Hα") as "[>[] _]". - 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βδ H]". done. - iMod (bor_sep with "LFT H") as "[Hβty _]". done. - iMod (bor_persistent with "LFT Hβδ Hα") as "[#Hβδ Hα]". done. - iMod (bor_persistent with "LFT Hβty Hα") as "[#Hβty Hα]". done. - iMod (bor_acc with "LFT H↦ Hα") as "[H↦ Hcloseα]". done. - wp_bind (!(LitV lx'))%E. iMod (bor_unnest with "LFT Hb") as "Hb"; first done. - wp_read. wp_op. wp_let. iMod "Hb". - iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]". iMod ("Hclose" with "Hα HL") as "HL". - iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [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. iSplitR. - - repeat iApply lft_incl_trans=>//. - - iApply (bor_shorten with "[] Hb"). iApply lft_incl_glb. - by iApply lft_incl_trans. by iApply lft_incl_refl. } - iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|]. - iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|]. - iApply type_jump; solve_typing. - Qed. - - (* Drop a rwlockwriteguard and release the lock. *) - Definition rwlockwriteguard_drop : val := - fn: ["x"] := - let: "x'" := !"x" in - "x'" <-ˢᶜ #0;; - delete [ #1; "x"];; - let: "r" := new [ #0] in return: ["r"]. - - Lemma rwlockwriteguard_drop_type ty : - typed_val rwlockwriteguard_drop - (fn(∀ α, ∅; rwlockwriteguard α ty) → unit). - Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". - iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst. - iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst. - iIntros (tid) "#LFT #HE Hna HL Hk HT". - rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. - iDestruct "HT" as "[Hx Hx']". - destruct x' as [[|lx'|]|]; try done. simpl. - iDestruct "Hx'" as (γ β tid_own) "(Hx' & #Hαβ & #Hβty & #Hinv & H◯)". - 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. - iMod (at_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|]. - iDestruct "INV" as (st) "(H↦ & H● & INV)". wp_write. - iMod ("Hcloseβ" with "[> H↦ H● H◯ INV Hx']") as "Hβ". - { iDestruct (own_valid_2 with "H● H◯") as %[[[=]| (? & st0 & [=<-] & -> & - [Heq|Hle])]%option_included Hv]%auth_both_valid_discrete; - last by destruct (exclusive_included _ _ Hle). - destruct st0 as [[[]|]| |]; try by inversion Heq. - iExists None. iFrame. iMod (own_update_2 with "H● H◯") as "$"; last done. - apply auth_update_dealloc. rewrite -(right_id None op (Some _)). - apply cancel_local_update_unit, _. } - iModIntro. wp_seq. iMod ("Hcloseα" with "Hβ") as "Hα". - iMod ("Hclose" with "Hα HL") as "HL". - iApply (type_type _ _ _ [ x ◁ box (uninit 1)] - with "[] LFT HE Hna HL Hk"); first last. - { rewrite tctx_interp_singleton tctx_hasty_val //. } - iApply type_delete; [solve_typing..|]. - iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst. - iApply type_jump; solve_typing. - Qed. -End rwlockwriteguard_functions. diff --git a/theories/typing/lib/slice/slice_basic.v b/theories/typing/lib/slice/slice_basic.v index 6aec9578..3eecba61 100644 --- a/theories/typing/lib/slice/slice_basic.v +++ b/theories/typing/lib/slice/slice_basic.v @@ -15,7 +15,7 @@ Section slice_basic. Proof. split; [by apply (type_lft_morphism_add_one κ)|done| |]. - move=> > EqSz */=. rewrite EqSz. by do 12 f_equiv. - - move=> > EqSz */=. rewrite EqSz. do 16 (f_contractive || f_equiv). by simpl in *. + - move=> > EqSz */=. rewrite EqSz. do 16 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance shr_slice_send {𝔄} κ (ty: type 𝔄) : Sync ty → Send (shr_slice κ ty). @@ -68,8 +68,8 @@ Section slice_basic. - move=> > EqSz EqLft */=. f_equiv. + apply equiv_dist. iDestruct EqLft as "#[??]". iSplit; iIntros "#In"; (iApply lft_incl_trans; [iApply "In"|done]). - + rewrite EqSz /uniq_body. do 23 (f_contractive || f_equiv). by simpl in *. - - move=> > EqSz */=. rewrite EqSz. do 17 (f_contractive || f_equiv). by simpl in *. + + rewrite EqSz /uniq_body. do 23 (f_contractive || f_equiv). by eapply dist_later_dist_lt. + - move=> > EqSz */=. rewrite EqSz. do 17 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance uniq_slice_send {𝔄} κ (ty: type 𝔄) : Send ty → Send (uniq_slice κ ty). @@ -162,7 +162,7 @@ Section slice_basic. - iSplit; [|done]. rewrite (tctx_hasty_val #n)/=. iExists _. iFrame "⧖". by iExists n. - iApply proph_obs_eq; [|done]=>/= π. f_equal. - by rewrite -vec_to_list_apply vec_to_list_length. + by rewrite -vec_to_list_apply length_vec_to_list. Qed. End slice_basic. diff --git a/theories/typing/lib/slice/slice_split.v b/theories/typing/lib/slice/slice_split.v index d2dbd45d..950e97e6 100644 --- a/theories/typing/lib/slice/slice_split.v +++ b/theories/typing/lib/slice/slice_split.v @@ -43,7 +43,7 @@ Section slice_split. iDestruct "↦r" as "(↦₀ & ↦₁ & ↦₂ & ↦₃ &_)". wp_write. wp_op. wp_write. do 3 wp_op. wp_write. do 2 wp_op. wp_write. do 2 wp_seq. iMod (proph_obs_sat with "PROPH Obs") as %(?& inat &->& Lt &_); [done|]. - rewrite -vec_to_list_apply vec_to_list_length in Lt. set ifin := nat_to_fin Lt. + rewrite -vec_to_list_apply length_vec_to_list in Lt. set ifin := nat_to_fin Lt. have Eqi: inat = ifin by rewrite fin_to_nat_to_fin. rewrite cctx_interp_singleton. iApply ("C" $! [# #_] -[λ π, (_, _)] with "Na L [-] []"). - iSplitL; [|done]. rewrite tctx_hasty_val. iExists _. iFrame "⧖". @@ -84,7 +84,7 @@ Section slice_split. iDestruct "↦r" as "(↦₀ & ↦₁ & ↦₂ & ↦₃ &_)". wp_write. wp_op. wp_write. do 3 wp_op. wp_write. do 2 wp_op. wp_write. do 2 wp_seq. iMod (proph_obs_sat with "PROPH Obs") as %(?& inat &->& Lt &_); [done|]. - rewrite -vec_to_list_apply vec_to_list_length in Lt. set ifin := nat_to_fin Lt. + rewrite -vec_to_list_apply length_vec_to_list in Lt. set ifin := nat_to_fin Lt. have Eqi: inat = ifin by rewrite fin_to_nat_to_fin. rewrite cctx_interp_singleton. iApply ("C" $! [# #_] -[λ π, (_, _)] with "Na L [-] []"). - iSplitL; [|done]. rewrite tctx_hasty_val. iExists _. iFrame "⧖". diff --git a/theories/typing/lib/smallvec/smallvec.v b/theories/typing/lib/smallvec/smallvec.v index 31d09ccb..a4600b0d 100644 --- a/theories/typing/lib/smallvec/smallvec.v +++ b/theories/typing/lib/smallvec/smallvec.v @@ -55,7 +55,7 @@ Section smallvec. iDestruct (big_sepL_ty_own_length with "tys") as %<-. rewrite heap_mapsto_vec_app trans_big_sepL_mt_ty_own shift_loc_assoc. iDestruct "↦" as "[? ↦ex]". iSplitR "↦ex"; iExists _; iFrame. - by rewrite -app_length. + by rewrite -length_app. - iDestruct 1 as (b ?????) "(↦hd & big)". case b=>/=. + rewrite trans_big_sepL_mt_ty_own. iDestruct "big" as "[(%wll & ↦ar & tys) (%wl' &->& ↦ex)]". @@ -63,7 +63,7 @@ Section smallvec. iExists ([_;_;_;_]++_++_). rewrite !heap_mapsto_vec_cons heap_mapsto_vec_app !shift_loc_assoc -Eqsz. iDestruct "↦hd" as "($&$&$&$&_)". iFrame "↦ar ↦ex". - iExists true, _, _, _, _, _. iSplit; [by rewrite -app_length|]. + iExists true, _, _, _, _, _. iSplit; [by rewrite -length_app|]. iExists _, _. by iFrame. + iDestruct "big" as "[(%&%&↦tl) ?]". iExists ([_;_;_;_]++_). rewrite !heap_mapsto_vec_cons !shift_loc_assoc. diff --git a/theories/typing/lib/smallvec/smallvec_basic.v b/theories/typing/lib/smallvec/smallvec_basic.v index 58598751..c07e28bb 100644 --- a/theories/typing/lib/smallvec/smallvec_basic.v +++ b/theories/typing/lib/smallvec/smallvec_basic.v @@ -148,7 +148,7 @@ Section smallvec_basic. iDestruct "↦" as "(↦₀ & ↦₁ & ↦₂ & ↦₃ &_)". case b=>/=; wp_read; wp_case. - rewrite trans_big_sepL_mt_ty_own. iDestruct "big" as "[(%wll & ↦ar & tys) (%wl' & -> & ↦ex)]". - iDestruct (big_sepL_ty_own_length with "tys") as %<-. rewrite -app_length. + iDestruct (big_sepL_ty_own_length with "tys") as %<-. rewrite -length_app. wp_apply (wp_delete (_::_::_::_::_) with "[↦₀ ↦₁ ↦₂ ↦₃ ↦ar ↦ex †]"); [done|..]. { rewrite !heap_mapsto_vec_cons heap_mapsto_vec_app !shift_loc_assoc freeable_sz_full. iFrame. } @@ -163,7 +163,7 @@ Section smallvec_basic. iDestruct "big" as "((%&<-& ↦tl) & (%& ↦ar & tys) & (%& %Eq & ↦ex) & †')". iDestruct (big_sepL_ty_own_length with "tys") as %Eq'. do 2 (wp_op; wp_read). do 3 wp_op. wp_read. - rewrite -Nat2Z.inj_add -Nat2Z.inj_mul !Nat.mul_add_distr_r -Eq -Eq' -app_length. + rewrite -Nat2Z.inj_add -Nat2Z.inj_mul !Nat.mul_add_distr_r -Eq -Eq' -length_app. wp_apply (wp_delete (_++_) with "[↦ar ↦ex †']"); [done|..]. { rewrite heap_mapsto_vec_app. iFrame. } iIntros "_". wp_seq. wp_apply (wp_delete (_::_::_::_::_) with "[↦₀ ↦₁ ↦₂ ↦₃ ↦tl †]"); [done| |]. @@ -209,7 +209,7 @@ Section smallvec_basic. iFrame "⧖ †r". iNext. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame "↦r". by iExists _. - iApply proph_obs_eq; [|done]=>/= ?. f_equal. - by rewrite -vec_to_list_apply vec_to_list_length. + by rewrite -vec_to_list_apply length_vec_to_list. Qed. End smallvec_basic. diff --git a/theories/typing/lib/smallvec/smallvec_pop.v b/theories/typing/lib/smallvec/smallvec_pop.v index fe118ba4..f11b86fa 100644 --- a/theories/typing/lib/smallvec/smallvec_pop.v +++ b/theories/typing/lib/smallvec/smallvec_pop.v @@ -85,7 +85,7 @@ Section smallvec_pop. iFrame "↦₀ ↦₁ ↦₂ ↦₃ ↦tys". iSplit; [done|]. iExists (_++_). rewrite heap_mapsto_vec_app shift_loc_assoc -Z.add_assoc -Nat2Z.inj_add -Lvl Nat.add_comm. - iFrame "↦last ↦tl". iPureIntro. rewrite app_length. lia. } + iFrame "↦last ↦tl". iPureIntro. rewrite length_app. lia. } iMod ("ToL" with "α L") as "L". iApply (type_type +[#v' ◁ &uniq{α} (smallvec n ty); #r ◁ box (option_ty ty)] -[pπ'; Some ∘ _] with "[] LFT TIME PROPH UNIQ E Na L C [-] []"). @@ -119,7 +119,7 @@ Section smallvec_pop. rewrite !heap_mapsto_vec_cons heap_mapsto_vec_nil !shift_loc_assoc. iFrame "↦₀ ↦₁ ↦₂ ↦₃ ↦tl ↦tys". have ->: len + S ex = S (len + ex) by lia. iFrame "†". iSplit; [done|]. iExists (_++_). rewrite heap_mapsto_vec_app. - iFrame "↦last". rewrite shift_loc_assoc_nat app_length Nat.add_comm Lvl. + iFrame "↦last". rewrite shift_loc_assoc_nat length_app Nat.add_comm Lvl. iFrame. iPureIntro=>/=. lia. } iMod ("ToL" with "α L") as "L". iApply (type_type +[#v' ◁ &uniq{α} (smallvec n ty); #r ◁ box (option_ty ty)] diff --git a/theories/typing/lib/smallvec/smallvec_push.v b/theories/typing/lib/smallvec/smallvec_push.v index 54414306..148ea341 100644 --- a/theories/typing/lib/smallvec/smallvec_push.v +++ b/theories/typing/lib/smallvec/smallvec_push.v @@ -51,7 +51,7 @@ Section smallvec_push. have Lwl: length wl = ty.(ty_size) + (n - S len) * ty.(ty_size). { rewrite Nat.mul_sub_distr_r Nat.add_sub_assoc; [lia|]. apply Nat.mul_le_mono_r. lia. } - move: (app_length_ex wl _ _ Lwl)=> [?[?[->[Lul ?]]]]. + move: (length_app_ex wl _ _ Lwl)=> [?[?[->[Lul ?]]]]. rewrite heap_mapsto_vec_app Lul !shift_loc_assoc. iDestruct "↦tl" as "[↦new ↦tl]". iApply (wp_memcpy with "[$↦new $↦x]"); [lia..|]. iIntros "!> [↦new ↦x]". iApply "ToΦ". iSplitR "↦x"; last first. { iExists _. by iFrame. } @@ -65,7 +65,7 @@ Section smallvec_push. rewrite vec_to_list_snoc big_sepL_app big_sepL_singleton. iSplitL "↦tys". + iStopProof. do 6 f_equiv. apply ty_own_depth_mono. lia. + iExists _. iSplitL "↦new". - * rewrite vec_to_list_length Nat.add_0_r shift_loc_assoc. iFrame. + * rewrite length_vec_to_list Nat.add_0_r shift_loc_assoc. iFrame. * iApply ty_own_depth_mono; [|done]. lia. - do 2 wp_op. wp_apply wp_new; [lia|done|]. iIntros (?) "[† ↦l]". have ->: ∀sz: nat, ((len + 1) * sz)%Z = len * sz + sz by lia. @@ -83,13 +83,13 @@ Section smallvec_push. iFrame "↦₀ ↦₁ ↦₂ ↦₃"=>/=. iSplit. { iPureIntro. fun_ext=> ?. by rewrite vec_to_list_snoc lapply_app. } rewrite Nat.add_comm Nat.add_0_r. iFrame "†". iSplitL "↦o ↦tl". - { iExists (_++_). rewrite EqLen app_length heap_mapsto_vec_app shift_loc_assoc. + { iExists (_++_). rewrite EqLen length_app heap_mapsto_vec_app shift_loc_assoc. iFrame "↦o". rewrite Lwll. by iFrame "↦tl". } iSplitL; last first. { iExists []. by rewrite heap_mapsto_vec_nil. } rewrite vec_to_list_snoc big_sepL_app big_sepL_singleton. iSplitL "tys ↦l". + rewrite trans_big_sepL_mt_ty_own. iExists _. iFrame "↦l". iStopProof. do 3 f_equiv. apply ty_own_depth_mono. lia. - + iExists _. rewrite Nat.add_0_r vec_to_list_length. iFrame "↦new". + + iExists _. rewrite Nat.add_0_r length_vec_to_list. iFrame "↦new". iApply ty_own_depth_mono; [|done]. lia. Qed. diff --git a/theories/typing/lib/take_mut.v b/theories/typing/lib/take_mut.v deleted file mode 100644 index 0b71a85d..00000000 --- a/theories/typing/lib/take_mut.v +++ /dev/null @@ -1,77 +0,0 @@ -From iris.proofmode Require Import proofmode. -From lrust.lang.lib Require Import memcpy. -From lrust.lifetime Require Import na_borrow. -From lrust.typing Require Export type. -From lrust.typing Require Import typing. -Set Default Proof Using "Type". - -Section code. - Context `{!typeG Σ}. - - Definition take ty (call_once : val) : val := - fn: ["x"; "f"] := - let: "x'" := !"x" in - let: "call_once" := call_once in - letalloc: "t" <-{ty.(ty_size)} !"x'" in - letcall: "r" := "call_once" ["f"; "t"]%E in - "x'" <-{ty.(ty_size)} !"r";; - delete [ #1; "x"];; delete [ #ty.(ty_size); "r"];; - let: "r" := new [ #0] in return: ["r"]. - - Lemma take_type ty fty call_once : - (* fty : FnOnce(ty) -> ty, as witnessed by the impl call_once *) - typed_val call_once (fn(∅; fty, ty) → ty) → - typed_val (take ty call_once) (fn(∀ α, ∅; &uniq{α} ty, fty) → unit). - Proof. - intros Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !>". iIntros (α ϝ ret arg). - inv_vec arg=>x env. simpl_subst. - iApply type_deref; [solve_typing..|]; iIntros (x'); simpl_subst. - iApply type_let; [apply Hf|solve_typing|]; iIntros (f'); simpl_subst. - iApply (type_new ty.(ty_size)); [solve_typing..|]; iIntros (t); simpl_subst. - (* Switch to Iris. *) - iIntros (tid) "#LFT #TIME #HE Hna HL Hk (Ht & Hf' & Hx & Hx' & Henv & _)". - rewrite !tctx_hasty_val [[x]]lock [[env]]lock [fn _]lock. - iDestruct "Hx'" as (depthx) "[Hdepthx Hx']". iDestruct "Hx'" as "[#Houtx Hx']". - iDestruct "Ht" as (deptht) "[Hdeptht Ht]". - iDestruct (ownptr_uninit_own with "Ht") as (tl tvl) "(% & % & >Htl & Htl†)". subst t. simpl. - destruct x' as [[|lx'|]|], depthx as [|depthx]=>//=. - iDestruct "Hx'" as (depthx' γ ?) "[H◯ Hx']". - 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 (bor_acc with "LFT Hx' Hα") as "[Hx' Hclose3]"; first done. - iDestruct "Hx'" as (depthx'') "(H● & Hdepthx'' & Hx')". - 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..|]. - iIntros "[Htl Hx'↦]". wp_seq. - (* Call the function. *) - wp_let. unlock. - iDestruct "Hf'" as (?) "[_ Hf']". - iApply (type_call_iris _ [ϝ] () [_; _] - with "LFT TIME HE Hna [Hϝ] Hf' [Henv Htl Htl† Hx'vl Hdepthx'']"); [solve_typing| | |]. - { by rewrite /= (right_id static). } - { rewrite /= !tctx_hasty_val tctx_hasty_val' //. iFrame. iSplitL; [|done]. - iExists _. iFrame. iExists _. iFrame. } - (* Prove the continuation of the function call. *) - iIntros (r depth') "Hna Hϝ Hdepth' Hr". simpl. - iDestruct (ownptr_own with "Hr") as (lr rvl) "(% & % & Hlr & Hrvl & Hlr†)". subst r. - wp_rec. - rewrite (right_id static). destruct depth'; [lia|]. - wp_apply (wp_memcpy with "[$Hx'↦ $Hlr]"); [by auto using vec_to_list_length..|]. - iIntros "[Hlx' Hlr]". wp_seq. - iMod (own_update_2 with "H● H◯") as "[H● H◯]"; [by apply excl_auth_update|]. - iMod ("Hclose3" with "[Hlx' Hrvl Hdepth' H●]") as "[Hlx' Hα]". - { iExists _. iFrame. iExists _. iFrame. } - iMod ("Hclose2" with "Hϝ HL") as "HL". - iMod ("Hclose1" with "Hα HL") as "HL". - (* Finish up the proof. *) - iApply (type_type _ _ _ [ x ◁ _; #lr ◁ box (uninit ty.(ty_size)) ] - with "[] LFT TIME HE Hna HL Hk [-]"); last first. - { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //. - unlock. iFrame. iExists (S depthx). iFrame. iExists _. - rewrite uninit_own. iFrame. auto using vec_to_list_length. } - iApply type_delete; [solve_typing..|]. - iApply type_delete; [solve_typing..|]. - iApply type_new; [solve_typing..|]; iIntros (r); simpl_subst. - iApply type_jump; solve_typing. - Qed. -End code. diff --git a/theories/typing/lib/vec/vec_basic.v b/theories/typing/lib/vec/vec_basic.v index af9dd03d..b94b5791 100644 --- a/theories/typing/lib/vec/vec_basic.v +++ b/theories/typing/lib/vec/vec_basic.v @@ -12,8 +12,8 @@ Section vec_basic. Global Instance vec_type_contractive 𝔄 : TypeContractive (vec_ty (𝔄:=𝔄)). Proof. split; [by apply type_lft_morphism_id_like|done| |]. - - move=>/= > ->*. do 19 (f_contractive || f_equiv). by simpl in *. - - move=>/= > ->*. do 16 (f_contractive || f_equiv). by simpl in *. + - move=>/= > ->*. do 19 (f_contractive || f_equiv). by eapply dist_later_dist_lt. + - move=>/= > ->*. do 16 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance vec_send {𝔄} (ty: type 𝔄) : Send ty → Send (vec_ty ty). @@ -124,9 +124,9 @@ Section vec_basic. do 2 (wp_op; wp_read). do 2 wp_op. wp_read. rewrite trans_big_sepL_mt_ty_own. iDestruct "big" as "((%& ↦old & tys) & (%& %Eq & ↦ex) & †')". iDestruct (big_sepL_ty_own_length with "tys") as %Eq'. - rewrite -Nat2Z.inj_add -Nat2Z.inj_mul !Nat.mul_add_distr_r -Eq -Eq' -app_length. + rewrite -Nat2Z.inj_add -Nat2Z.inj_mul !Nat.mul_add_distr_r -Eq -Eq' -length_app. wp_apply (wp_delete (_++_) with "[↦old ↦ex †']"); [done|..]. - { rewrite heap_mapsto_vec_app app_length. iFrame. } + { rewrite heap_mapsto_vec_app length_app. iFrame. } iIntros "_". wp_seq. wp_apply (wp_delete [_;_;_] with "[↦₀ ↦₁ ↦₂ †]"); [done| |]. { rewrite !heap_mapsto_vec_cons shift_loc_assoc heap_mapsto_vec_nil freeable_sz_full. iFrame. } @@ -171,7 +171,7 @@ Section vec_basic. iFrame "⧖ †r". iNext. iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame "↦r". by iExists _. - iApply proph_obs_eq; [|done]=>/= ?. f_equal. - by rewrite -vec_to_list_apply vec_to_list_length. + by rewrite -vec_to_list_apply length_vec_to_list. Qed. End vec_basic. diff --git a/theories/typing/lib/vec/vec_pushpop.v b/theories/typing/lib/vec/vec_pushpop.v index fe455138..ecacb382 100644 --- a/theories/typing/lib/vec/vec_pushpop.v +++ b/theories/typing/lib/vec/vec_pushpop.v @@ -125,7 +125,7 @@ Section vec_pushpop. have ->: (ex + 1)%Z = S ex by lia. iExists _, _, _, _. rewrite !heap_mapsto_vec_cons heap_mapsto_vec_nil shift_loc_assoc. iFrame "↦₀ ↦₁ ↦₂ ↦tys †". iSplit; [done|]. iExists (vl ++ wl). - rewrite app_length heap_mapsto_vec_app shift_loc_assoc_nat plus_comm Eqvl. + rewrite length_app heap_mapsto_vec_app shift_loc_assoc_nat (comm plus) Eqvl. iSplit; [iPureIntro; lia|]. iFrame. } iMod ("ToL" with "α L") as "L". iApply (type_type +[#v' ◁ &uniq{α} (vec_ty ty); #r ◁ box (option_ty ty)] diff --git a/theories/typing/lib/vec_util.v b/theories/typing/lib/vec_util.v index 47e25983..690c0c1e 100644 --- a/theories/typing/lib/vec_util.v +++ b/theories/typing/lib/vec_util.v @@ -48,7 +48,7 @@ Section vec_util. wp_rec. wp_read. wp_let. do 2 (wp_op; wp_read; wp_let). do 2 wp_op. wp_write. wp_op. wp_case. have ->: (len + 1)%Z = S len by lia. move: Lvl. case ex as [|ex]=>/= Lvl; last first. - { move: {Lvl}(app_length_ex vl _ _ Lvl)=> [vl'[?[->[Lvl ?]]]]. + { move: {Lvl}(length_app_ex vl _ _ Lvl)=> [vl'[?[->[Lvl ?]]]]. rewrite heap_mapsto_vec_app Lvl shift_loc_assoc_nat Nat.add_comm. iDestruct "↦ex" as "[↦new ↦ex]". have ->: len + S ex = S len + ex by lia. do 2 wp_op. have ->: (S ex - 1)%Z = ex by lia. wp_write. do 2 wp_op. @@ -60,9 +60,9 @@ Section vec_util. rewrite vec_to_list_snoc big_sepL_app. iSplitL "↦l tys". - rewrite trans_big_sepL_mt_ty_own. iExists _. iFrame "↦l". iStopProof. do 3 f_equiv. apply ty_own_depth_mono. lia. - - rewrite big_sepL_singleton. iExists _. rewrite Nat.add_0_r vec_to_list_length. + - rewrite big_sepL_singleton. iExists _. rewrite Nat.add_0_r length_vec_to_list. iFrame "↦new". iApply ty_own_depth_mono; [|done]. lia. } - rewrite plus_0_r. wp_op. wp_write. do 3 wp_op. + replace (len + 0) with len by lia. wp_op. wp_write. do 3 wp_op. wp_apply wp_new; [lia|done|]. iIntros (?) "[†' ↦l']". wp_let. have ->: ∀sz: nat, ((2 * len + 1) * sz)%Z = len * sz + (sz + len * sz) by lia. rewrite Nat2Z.id !repeat_app !heap_mapsto_vec_app !repeat_length shift_loc_assoc_nat. @@ -81,7 +81,7 @@ Section vec_util. rewrite vec_to_list_snoc big_sepL_app. iSplitL "↦l' tys". - rewrite trans_big_sepL_mt_ty_own. iExists _. iFrame "↦l'". iStopProof. do 3 f_equiv. apply ty_own_depth_mono. lia. - - rewrite big_sepL_singleton. iExists _. rewrite Nat.add_0_r vec_to_list_length. + - rewrite big_sepL_singleton. iExists _. rewrite Nat.add_0_r length_vec_to_list. iFrame "↦new". iApply ty_own_depth_mono; [|done]. lia. Qed. End vec_util. diff --git a/theories/typing/own.v b/theories/typing/own.v index d06e2d1b..e3d23997 100644 --- a/theories/typing/own.v +++ b/theories/typing/own.v @@ -20,7 +20,7 @@ Section own. Proof. case n=> [|n']. { iSplit; iIntros; by [iRight|]. } have ->: Z.of_nat (S n') = 0 ↔ False by done. - by rewrite right_id /freeable_sz Qp_div_diag. + by rewrite right_id /freeable_sz Qp.div_diag. Qed. Lemma freeable_sz_full_S n l : freeable_sz (S n) (S n) l ⊣⊢ †{1}l…(S n). @@ -37,7 +37,7 @@ Section own. - by rewrite right_id Nat.add_0_r. - iSplit; by [iIntros "[??]"|iIntros]. - rewrite heap_freeable_op_eq. f_equiv; [|done]. - by rewrite -Qp_div_add_distr pos_to_Qp_add -Nat2Pos.inj_add. + by rewrite -Qp.div_add_distr pos_to_Qp_add -Nat2Pos.inj_add. Qed. Lemma freeable_sz_eq n sz sz' l : @@ -103,8 +103,8 @@ Section own. Global Instance own_type_contractive 𝔄 n : TypeContractive (@own_ptr 𝔄 n). Proof. split; [by apply type_lft_morphism_id_like|done| |]. - - move=>/= > ->*. do 9 (f_contractive || f_equiv). by simpl in *. - - move=>/= > *. do 6 (f_contractive || f_equiv). by simpl in *. + - move=>/= > ->*. do 9 (f_contractive || f_equiv). by eapply dist_later_dist_lt. + - move=>/= > *. do 6 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance own_send {𝔄} n (ty: type 𝔄) : Send ty → Send (own_ptr n ty). @@ -172,8 +172,8 @@ Section box. Global Instance box_type_contractive 𝔄 : TypeContractive (@box 𝔄). Proof. split; [by apply type_lft_morphism_id_like|done| |]. - - move=>/= > ->*. do 9 (f_contractive || f_equiv). by simpl in *. - - move=>/= *. do 6 (f_contractive || f_equiv). by simpl in *. + - move=>/= > ->*. do 9 (f_contractive || f_equiv). by eapply dist_later_dist_lt. + - move=>/= *. do 6 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Lemma box_type_incl {𝔄 𝔅} (f: 𝔄 → 𝔅) ty ty': diff --git a/theories/typing/product.v b/theories/typing/product.v index 0cb50dac..5ebc6a45 100644 --- a/theories/typing/product.v +++ b/theories/typing/product.v @@ -76,7 +76,7 @@ Section product. ty'.(ty_shr) (snd ∘ vπ) d κ tid (l +ₗ ty.(ty_size)) |}%I. Next Obligation. - iIntros "* (%&%&->& H)". rewrite app_length !ty_size_eq. by iDestruct "H" as "[->->]". + iIntros "* (%&%&->& H)". rewrite length_app !ty_size_eq. by iDestruct "H" as "[->->]". Qed. Next Obligation. move=>/= *. do 6 f_equiv; by apply ty_own_depth_mono. Qed. Next Obligation. move=>/= *. f_equiv; by apply ty_shr_depth_mono. Qed. @@ -121,7 +121,7 @@ Section product. iDestruct (ty_shr_proph with "LFT In [] ty' κ₊") as "> Toκ₊"; first done. { iApply lft_incl_trans; [done|]. rewrite lft_intersect_list_app. iApply lft_intersect_incl_r. } - iIntros "!>!>". iCombine "Toκ Toκ₊" as ">Toκ2". + iIntros "!>!>". iMod "Toκ". iMod "Toκ₊". iCombine "Toκ Toκ₊" as "Toκ2". iApply (step_fupdN_wand with "Toκ2"). iIntros "!> [Toκ Toκ₊]". iMod "Toκ" as (ξl q ?) "[ξl Toκ]". iMod "Toκ₊" as (ξl' q' ?) "[ξl' Toκ₊]". iDestruct (proph_tok_combine with "ξl ξl'") as (q0) "[ξl Toξl]". @@ -244,10 +244,10 @@ Section typing. iMod (copy_shr_acc with "LFT ty' Na κ₊") as (q' wl') "(Na & ↦' & #ty' & Toκ₊)"; first done. { apply subseteq_difference_r. { symmetry. apply shr_locsE_disj. } - move: HF. rewrite -plus_assoc shr_locsE_shift. set_solver. } + move: HF. rewrite -assoc shr_locsE_shift. set_solver. } iDestruct (na_own_acc with "Na") as "[$ ToNa]". { rewrite shr_locsE_shift. set_solver. } - case (Qp_lower_bound q q')=> [q''[?[?[->->]]]]. iExists q'', (wl ++ wl'). + case (Qp.lower_bound q q')=> [q''[?[?[->->]]]]. iExists q'', (wl ++ wl'). rewrite heap_mapsto_vec_app. iDestruct (ty_size_eq with "ty") as ">->". iDestruct "↦" as "[$ ↦r]". iDestruct "↦'" as "[$ ↦r']". iSplitR. { iIntros "!>!>". iExists wl, wl'. iSplit; by [|iSplit]. } @@ -318,8 +318,8 @@ Section typing. - iIntros (?? vπ) "*% #LFT #E [L L₊] [ty ty']". iMod (Rls with "LFT E L ty") as "Upd"; [done|]. iMod (Rls' with "LFT E L₊ ty'") as "Upd'"; [done|]. iIntros "!>!>". - iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand with "Upd"). - iIntros "[>(%Eq &$&$) >(%Eq' &$&$)] !%". + iMod "Upd". iMod "Upd'". iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand with "Upd"). + iIntros "!> [>(%Eq &$&$) >(%Eq' &$&$)] !%". move: Eq=> [a Eq]. move: Eq'=> [b Eq']. exists (a, b). fun_ext=>/= π. move: (equal_f Eq π) (equal_f Eq' π)=>/=. by case (vπ π)=>/= ??<-<-. diff --git a/theories/typing/product_split.v b/theories/typing/product_split.v index 707e54ac..81c9dc12 100644 --- a/theories/typing/product_split.v +++ b/theories/typing/product_split.v @@ -394,11 +394,6 @@ Global Hint Resolve tctx_extract_split_shr_prod tctx_extract_split_shr_xprod [tctx_extract_elt_further]. Moreover, it is placed in a difference hint db. The reason is that it can make the proof search diverge if the type is an evar. - - Unfortunately, priorities are not taken into account accross hint - databases with [typeclasses eauto], so this is useless, and some - solve_typing get slow because of that. See: - https://coq.inria.fr/bugs/show_bug.cgi?id=5304 *) Global Hint Resolve tctx_extract_merge_own_prod tctx_extract_merge_own_xprod - | 40 : lrust_typing_merge. + | 60 : lrust_typing_merge. diff --git a/theories/typing/programs.v b/theories/typing/programs.v index 913252cc..57a1ec58 100644 --- a/theories/typing/programs.v +++ b/theories/typing/programs.v @@ -15,7 +15,7 @@ Section typing. lft_ctx -∗ time_ctx -∗ proph_ctx -∗ uniq_ctx -∗ elctx_interp E -∗ na_own tid ⊤ -∗ llctx_interp L 1 -∗ cctx_interp tid postπ C -∗ tctx_interp tid T vπl -∗ ⟨π, tr (postπ π) (vπl -$ π)⟩ -∗ WP e {{ _, cont_postcondition }}. - Global Arguments typed_body {_ _} _ _ _ _ _%E _%type. + Global Arguments typed_body {_ _} _ _ _ _ _%_E _%_type. Global Instance typed_body_proper 𝔄l 𝔅 E L C T e : Proper ((≡) ==> (≡)) (@typed_body 𝔄l 𝔅 E L C T e). @@ -24,7 +24,7 @@ Section typing. Lemma typed_body_impl {𝔄l 𝔅} (tr tr': predl_trans' 𝔄l 𝔅) E L (C: cctx 𝔅) (T: tctx 𝔄l) e : (∀post vl, tr post vl → tr' post vl) → - typed_body E L C T e tr' -∗ typed_body E L C T e tr. + typed_body E L C T e tr' ⊢ typed_body E L C T e tr. Proof. move=> Imp. rewrite /typed_body. do 16 f_equiv=>/=. do 2 f_equiv. move=> ?. by apply Imp. @@ -47,7 +47,7 @@ Section typing. na_own tid ⊤ -∗ llctx_interp L 1 -∗ tctx_interp tid T vπl -∗ ⟨π, tr (postπ π) (vπl -$ π)⟩ -∗ WP e {{ v, ∃vπl', na_own tid ⊤ ∗ llctx_interp L 1 ∗ tctx_interp tid (T' v) vπl' ∗ ⟨π, postπ π (vπl' -$ π)⟩ }}. - Global Arguments typed_instr {_ _} _ _ _ _%E _ _%type. + Global Arguments typed_instr {_ _} _ _ _ _%_E _ _%_type. (** Writing and Reading *) @@ -59,7 +59,7 @@ Section typing. ⌜v = #l⌝ ∗ ▷ l ↦∗: tyb.(ty_own) (gt ∘ vπ) d tid ∗ ∀wπ db', ▷ l ↦∗: tyb'.(ty_own) wπ db' tid -∗ ⧖(S db') ={⊤}=∗ llctx_interp L qL ∗ ty'.(ty_own) (st ∘ vπ ⊛ wπ) (S db') tid [v]. - Global Arguments typed_write {_ _ _ _} _ _ _%T _%T _%T _%T _%type _%type. + Global Arguments typed_write {_ _ _ _} _ _ _%_T _%_T _%_T _%_T _%_type _%_type. (* Technically speaking, we could remvoe the vl quantifiaction here and use mapsto_pred instead (i.e., l ↦∗: ty.(ty_own) tid). However, that would @@ -72,20 +72,20 @@ Section typing. ty.(ty_own) vπ d tid [v] ={⊤}=∗ ∃(l: loc) vl q, ⌜v = #l⌝ ∗ l ↦∗{q} vl ∗ ▷ tyb.(ty_own) (gt ∘ vπ) d tid vl ∗ (l ↦∗{q} vl ={⊤}=∗ na_own tid ⊤ ∗ llctx_interp L qL ∗ ty'.(ty_own) (st ∘ vπ) d tid [v]). - Global Arguments typed_read {_ _ _} _ _ _%T _%T _%T _ _%type. + Global Arguments typed_read {_ _ _} _ _ _%_T _%_T _%_T _ _%_type. Definition typed_instr_ty {𝔄l 𝔅} (E: elctx) (L: llctx) (T: tctx 𝔄l) (e: expr) (ty: type 𝔅) (tr: pred' 𝔅 → predl 𝔄l) : Prop := typed_instr E L T e (λ v, +[v ◁ ty]) (λ post al, tr (λ b, post -[b]) al). - Global Arguments typed_instr_ty {_ _} _ _ _ _%E _%T _%type. + Global Arguments typed_instr_ty {_ _} _ _ _ _%_E _%_T _%_type. Definition typed_val {𝔄} (v: val) (ty: type 𝔄) (a: 𝔄) : Prop := ∀E L, typed_instr_ty E L +[] (of_val v) ty (λ post _, post a). - Global Arguments typed_val {_} _%V _%T _%type. + Global Arguments typed_val {_} _%_V _%_T _%_type. (* This lemma is helpful for specifying the predicate transformer. *) Lemma type_with_tr 𝔄l 𝔅 tr E L (C: cctx 𝔅) (T: tctx 𝔄l) e : - typed_body E L C T e tr -∗ typed_body E L C T e tr. + typed_body E L C T e tr ⊢ typed_body E L C T e tr. Proof. done. Qed. (* This lemma is helpful when switching from proving unsafe code in Iris diff --git a/theories/typing/shr_bor.v b/theories/typing/shr_bor.v index aa564980..5c87e083 100644 --- a/theories/typing/shr_bor.v +++ b/theories/typing/shr_bor.v @@ -42,7 +42,7 @@ Section typing. Proof. split; [by apply (type_lft_morphism_add_one κ)|done| |]. - move=>/= *. by do 4 f_equiv. - - move=>/= *. do 8 (f_contractive || f_equiv). by simpl in *. + - move=>/= *. do 8 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance shr_send {𝔄} κ (ty: type 𝔄) : Sync ty → Send (&shr{κ} ty). diff --git a/theories/typing/soundness.v b/theories/typing/soundness.v index fa15c7fa..cc9b375a 100644 --- a/theories/typing/soundness.v +++ b/theories/typing/soundness.v @@ -7,12 +7,12 @@ From lrust.typing Require Import typing. Set Default Proof Using "Type". Class typePreG Σ := PreTypeG { - type_preG_lrustGS :> lrustGpreS Σ; - type_preG_lftGS :> lftGpreS Σ; - type_preG_na_invG :> na_invG Σ; - type_preG_frac_borG :> frac_borG Σ; - type_preG_prophG:> prophPreG Σ; - type_preG_uniqG:> uniqPreG Σ; + type_preG_lrustGS :: lrustGpreS Σ; + type_preG_lftGS :: lftGpreS Σ; + type_preG_na_invG :: na_invG Σ; + type_preG_frac_borG :: frac_borG Σ; + type_preG_prophG:: prophPreG Σ; + type_preG_uniqG:: uniqPreG Σ; }. Definition typeΣ : gFunctors := diff --git a/theories/typing/sum.v b/theories/typing/sum.v index 0d0f09f9..b19c6ad7 100644 --- a/theories/typing/sum.v +++ b/theories/typing/sum.v @@ -27,12 +27,12 @@ Section sum. iDestruct (heap_mapsto_vec_app with "↦") as "[↦ ↦']". iSplitL "↦'"; [|iExists vl'; by iFrame]. iExists vl''. rewrite (shift_loc_assoc_nat _ 1) Eq'. iFrame "↦'". iPureIntro. - by rewrite -Eq' -app_length. + by rewrite -Eq' -length_app. - iDestruct 1 as (i vπ' ->) "[[↦ (%vl'' & ↦'' &%)] (%vl' & ↦' & ty)]". iDestruct (ty_size_eq with "ty") as "%Eq". iExists (#i :: vl' ++ vl''). rewrite heap_mapsto_vec_cons heap_mapsto_vec_app (shift_loc_assoc_nat _ 1) Eq. iFrame "↦ ↦' ↦''". iExists i, vπ', vl', vl''. iFrame "ty". iPureIntro. - do 2 (split; [done|]). rewrite/= app_length Eq. by f_equal. + do 2 (split; [done|]). rewrite/= length_app Eq. by f_equal. Qed. Lemma ty_lfts_lookup_incl {𝔄l} (tyl: typel 𝔄l) i : @@ -196,7 +196,7 @@ Section typing. iDestruct (na_own_acc with "Na") as "[$ ToNa]". { apply difference_mono_l. trans (shr_locsE (l +ₗ 1) (max_ty_size tyl)). { apply shr_locsE_subseteq. lia. } set_solver+. } - case (Qp_lower_bound q q')=> [q''[?[?[->->]]]]. + case (Qp.lower_bound q q')=> [q''[?[?[->->]]]]. iExists q'', (#i :: vl ++ vl'). rewrite heap_mapsto_vec_cons heap_mapsto_vec_app shift_loc_assoc -Nat.add_1_l Nat2Z.inj_add. @@ -204,7 +204,7 @@ Section typing. iDestruct (ty_size_eq with "ty") as ">%Eq". rewrite Eq. iDestruct "↦pad" as "[$ ↦pad]". iSplitR. { iIntros "!>!>". iExists i, _, vl, vl'. iFrame "ty". iPureIntro. - do 2 (split; [done|]). rewrite/= app_length Eq. by f_equal. } + do 2 (split; [done|]). rewrite/= length_app Eq. by f_equal. } iIntros "!> Na (↦i' & ↦' & ↦pad')". iDestruct ("ToNa" with "Na") as "Na". iMod ("Toκ₊" with "Na [$↦ $↦']") as "[$$]". iApply "Toκ". iFrame "↦i ↦i'". iExists vl'. by iFrame. diff --git a/theories/typing/type.v b/theories/typing/type.v index 47be8272..1d8e38e1 100644 --- a/theories/typing/type.v +++ b/theories/typing/type.v @@ -11,12 +11,12 @@ Set Default Proof Using "Type". Implicit Type (𝔄 𝔅 ℭ: syn_type) (𝔄l 𝔅l: syn_typel). Class typeG Σ := TypeG { - type_lrustGS :> lrustGS Σ; - type_prophG :> prophG Σ; - type_uniqG :> uniqG Σ; - type_lftGS :> lftGS Σ; - type_na_invG :> na_invG Σ; - type_frac_borG :> frac_borG Σ; + type_lrustGS :: lrustGS Σ; + type_prophG :: prophG Σ; + type_uniqG :: uniqG Σ; + type_lftGS :: lftGS Σ; + type_na_invG :: na_invG Σ; + type_frac_borG :: frac_borG Σ; }. Definition lrustN := nroot .@ "lrust". @@ -33,11 +33,11 @@ Record type `{!typeG Σ} 𝔄 := { ty_shr_persistent vπ d κ tid l : Persistent (ty_shr vπ d κ tid l); - ty_size_eq vπ d tid vl : ty_own vπ d tid vl -∗ ⌜length vl = ty_size⌝; + ty_size_eq vπ d tid vl : ty_own vπ d tid vl ⊢ ⌜length vl = ty_size⌝; ty_own_depth_mono d d' vπ tid vl : - (d ≤ d')%nat → ty_own vπ d tid vl -∗ ty_own vπ d' tid vl; + (d ≤ d')%nat → ty_own vπ d tid vl ⊢ ty_own vπ d' tid vl; ty_shr_depth_mono d d' vπ κ tid l : - (d ≤ d')%nat → ty_shr vπ d κ tid l -∗ ty_shr vπ d' κ tid l; + (d ≤ d')%nat → ty_shr vπ d κ tid l ⊢ ty_shr vπ d' κ tid l; ty_shr_lft_mono κ κ' vπ d tid l : κ' ⊑ κ -∗ ty_shr vπ d κ tid l -∗ ty_shr vπ d κ' tid l; @@ -167,9 +167,9 @@ Record simple_type `{!typeG Σ} 𝔄 := { st_size: nat; st_lfts: list lft; st_E: elctx; st_own: proph 𝔄 → nat → thread_id → list val → iProp Σ; st_own_persistent vπ d tid vl : Persistent (st_own vπ d tid vl); - st_size_eq vπ d tid vl : st_own vπ d tid vl -∗ ⌜length vl = st_size⌝; + st_size_eq vπ d tid vl : st_own vπ d tid vl ⊢ ⌜length vl = st_size⌝; st_own_depth_mono d d' vπ tid vl : - (d ≤ d')%nat → st_own vπ d tid vl -∗ st_own vπ d' tid vl; + (d ≤ d')%nat → st_own vπ d tid vl ⊢ st_own vπ d' tid vl; st_own_proph E vπ d tid vl κ q : ↑lftN ⊆ E → lft_ctx -∗ κ ⊑ lft_intersect_list st_lfts -∗ st_own vπ d tid vl -∗ q.[κ] ={E}=∗ |={E}▷=>^d |={E}=> ∃ξl q', ⌜vπ ./ ξl⌝ ∗ @@ -238,7 +238,7 @@ Program Definition st_of_pt `{!typeG Σ} {𝔄} (pt: plain_type 𝔄) : simple_t st_own vπ d tid vl := ∃v, ⌜vπ = const v⌝ ∗ pt.(pt_own) v tid vl; |}%I. Next Obligation. move=> >. iIntros "[%[_?]]". by iApply pt_size_eq. Qed. -Next Obligation. done. Qed. +Next Obligation. eauto. Qed. Next Obligation. move=> * /=. iIntros "_ _[%[->?]]". iIntros "$ !>". iApply step_fupdN_full_intro. iModIntro. iExists [], 1%Qp. @@ -537,7 +537,7 @@ Qed. End type_lft_morphism. Class TypeNonExpansive `{!typeG Σ} {𝔄 𝔅} (T: type 𝔄 → type 𝔅) : Prop := { - type_ne_type_lft_morphism :> TypeLftMorphism T; + type_ne_type_lft_morphism :: TypeLftMorphism T; type_ne_ty_size ty ty' : ty.(ty_size) = ty'.(ty_size) → (T ty).(ty_size) = (T ty').(ty_size); type_ne_ty_own n ty ty' : @@ -587,39 +587,61 @@ Section type_contractive. Proof. move=> HT. split; [by apply HT|move=> *; by apply HT| |]. - move=> *. apply HT=>// *; by [apply dist_dist_later|apply dist_S]. - - move=> n *. apply HT=>// *; [|by apply dist_dist_later]. - case n as [|[|]]=>//. simpl in *. by apply dist_S. + - move=> n ????? DIST *. apply HT=>// *; [|by apply dist_dist_later]. + case n as [|[|]]=>//. apply DIST. lia. Qed. Global Instance type_ne_ne_compose {𝔄 𝔅 ℭ} (T: type 𝔅 → type ℭ) (T': type 𝔄 → type 𝔅) : TypeNonExpansive T → TypeNonExpansive T' → TypeNonExpansive (T ∘ T'). Proof. - move=> HT HT'. split; [by apply _|move=> *; by apply HT, HT'| |]; - (move=> n *; apply HT; try (by apply HT'); - first (by iApply type_lft_morphism_lft_equiv_proper); - first (apply type_lft_morphism_elctx_interp_proper=>//; apply _)). - move=> *. case n as [|]=>//. by apply HT'. + move=> HT HT'. split; [by apply _|move=> *; by apply HT, HT'| |]. + - move=> n ????? DIST1 DIST2; apply HT; try (by apply HT'). + + by iApply type_lft_morphism_lft_equiv_proper. + + apply type_lft_morphism_elctx_interp_proper=>//; apply _. + + intros. eapply HT'=>//. intros. apply dist_later_S. done. + - move=> n ????? DIST1 DIST2; apply HT; try (by apply HT'). + + by iApply type_lft_morphism_lft_equiv_proper. + + apply type_lft_morphism_elctx_interp_proper=>//; apply _. + + intros. destruct n; [apply dist_later_0|]. apply dist_later_S, HT'=>// *. + eapply dist_later_dist_lt; [|done]. lia. Qed. Global Instance type_contractive_compose_right {𝔄 𝔅 ℭ} (T: type 𝔅 → type ℭ) (T': type 𝔄 → type 𝔅) : TypeContractive T → TypeNonExpansive T' → TypeContractive (T ∘ T'). Proof. - move=> HT HT'. split; [by apply _|move=> *; by apply HT| |]; - (move=> n *; apply HT; try (by apply HT'); - first (by iApply type_lft_morphism_lft_equiv_proper); - first (apply type_lft_morphism_elctx_interp_proper=>//; apply _)); - move=> *; case n as [|[|]]=>//; by apply HT'. + move=> HT HT'. split; [by apply _|move=> *; by apply HT| |]. + - move=> n ????? DIST1 DIST2; apply HT; try (by apply HT'). + + by iApply type_lft_morphism_lft_equiv_proper. + + apply type_lft_morphism_elctx_interp_proper=>//; apply _. + + intros. destruct n; [apply dist_later_0|]. apply dist_later_S, HT'=>// *. + eapply dist_later_dist_lt; [|done]. lia. + - move=> n ????? DIST1 DIST2; apply HT; try (by apply HT'). + + by iApply type_lft_morphism_lft_equiv_proper. + + apply type_lft_morphism_elctx_interp_proper=>//; apply _. + + intros. destruct n as [|[]]=>//. apply HT'=>// *. + eapply dist_later_dist_lt; [|done]. lia. + + intros. destruct n; [apply dist_later_0|]; apply dist_later_S, HT'=>// *. + * destruct n; [apply dist_later_0|]; apply dist_later_S=>//. + * eapply dist_later_dist_lt; [|done]. lia. Qed. Global Instance type_contractive_compose_left {𝔄 𝔅 ℭ} (T: type 𝔅 → type ℭ) (T': type 𝔄 → type 𝔅) : TypeNonExpansive T → TypeContractive T' → TypeContractive (T ∘ T'). Proof. - move=> HT HT'. split; [by apply _|move=> *; by apply HT, HT'| |]; - (move=> n *; apply HT; try (by apply HT'); - first (by iApply type_lft_morphism_lft_equiv_proper); - first (apply type_lft_morphism_elctx_interp_proper=>//; apply _)); - move=> *; case n as [|]=>//; by apply HT'. + move=> HT HT'. split; [by apply _|move=> *; by apply HT, HT'| |]. + - move=> n ????? DIST1 DIST2; apply HT; try (by apply HT'). + + by iApply type_lft_morphism_lft_equiv_proper. + + apply type_lft_morphism_elctx_interp_proper=>//; apply _. + + intros. apply HT'=>// *. + * destruct n=>//. eapply dist_later_dist_lt; [|done]. lia. + * by apply dist_later_S. + - move=> n ????? DIST1 DIST2; apply HT; try (by apply HT'). + + by iApply type_lft_morphism_lft_equiv_proper. + + apply type_lft_morphism_elctx_interp_proper=>//; apply _. + + intros. destruct n; [apply dist_later_0|]. apply dist_later_S, HT'=>// *. + * destruct n; [apply dist_later_0|]; apply dist_later_S=>//. + * eapply dist_later_dist_lt; [|done]. lia. Qed. Global Instance const_type_contractive {𝔄 𝔅} (ty: type 𝔄) : diff --git a/theories/typing/type_context.v b/theories/typing/type_context.v index e1da9699..7b3b8005 100644 --- a/theories/typing/type_context.v +++ b/theories/typing/type_context.v @@ -70,7 +70,7 @@ Section type_context. eval_path p = Some v → ⊢ WP p @ E {{ v', ⌜v' = v⌝ }}. Proof. move: v. elim: p=>//. - - move=> > [=?]. by iApply wp_value. + - move=> > [=] ?. by iApply wp_value. - move=> > ?? /of_to_val ?. by iApply wp_value. - case=>// e Wp. case=>//. case=>//= ?. move: Wp. case (eval_path e)=>//. case=>//. case=>// ? Wp _ ?[=<-]. @@ -205,11 +205,10 @@ Section lemmas. iIntros ([Rl _] Rl' ??[??]) "#LFT #E [L L₊] /=[(%&%&%& ⧖ & ty) T]". iMod (Rl with "LFT E L ty") as "Upd"; [done|]. iMod (Rl' with "LFT E L₊ T") as (?) "[⧖' Upd']". iCombine "⧖ ⧖'" as "#⧖". - iCombine "Upd Upd'" as "Upd". iExists _. iFrame "⧖". + iCombine "Upd Upd'" as "Upd". iFrame "⧖". iApply (step_fupdN_wand with "Upd"). iIntros "!> [>(%Eq &$& ty) >(%Eq' &$&$)] !>". iSplit; last first. - { iExists _, _. iFrame "ty". iSplit; [done|]. - iApply persistent_time_receipt_mono; [|done]. lia. } + { iDestruct (ty_own_depth_mono with "ty") as "ty"; [|by iFrame]. lia. } iPureIntro. move: Eq=> [b Eq]. move: Eq'=> [bl Eq']. exists (b -:: bl). fun_ext=> π. by move: (equal_f Eq π) (equal_f Eq' π)=>/= <-<-. Qed. diff --git a/theories/typing/type_sum.v b/theories/typing/type_sum.v index 464e99b6..677bf919 100644 --- a/theories/typing/type_sum.v +++ b/theories/typing/type_sum.v @@ -449,7 +449,7 @@ Section type_sum. move: (max_hlist_with_ge (λ _, ty_size) tyl i). set tyisz := tyi.(ty_size). set maxsz := max_hlist_with _ _=> ?. have LenvlAdd: length vl = tyisz + (maxsz - tyisz) by lia. - move: LenvlAdd=> /app_length_ex[wl[wl'[->[Lenwl ?]]]]. + move: LenvlAdd=> /length_app_ex[wl[wl'[->[Lenwl ?]]]]. rewrite heap_mapsto_vec_app Lenwl. iDestruct "↦" as "[↦ ↦p]". iApply wp_fupd. iApply (wp_persistent_time_receipt with "TIME ⧖'"); [done|]. iApply (wp_memcpy with "[$↦ $↦']"); [lia..|]. iIntros "!> [↦ ↦'] #⧖'S". @@ -457,7 +457,7 @@ Section type_sum. iMod ("Totyw'" with "[↦i ↦ ↦p tyi] ⧖'S") as "($& tyw')". { iNext. iExists (_::_++_). rewrite heap_mapsto_vec_cons heap_mapsto_vec_app. rewrite /tyisz -Lenvl'. iFrame "↦i ↦ ↦p". iExists _, _, _, _. iFrame "tyi". - iPureIntro. do 2 (split; [done|]). rewrite/= app_length. lia. } + iPureIntro. do 2 (split; [done|]). rewrite/= length_app. lia. } iModIntro. iExists -[λ π, _;λ π, _]. rewrite right_id. iSplit; last first. { iApply proph_obs_impl; [|done]=>/= ?[? Imp]. by apply Imp. } iSplitL "tyw'"; iExists _, _; (iSplit; [done|]); by iFrame. diff --git a/theories/typing/uninit.v b/theories/typing/uninit.v index ab5fbc76..8d786522 100644 --- a/theories/typing/uninit.v +++ b/theories/typing/uninit.v @@ -28,7 +28,7 @@ Section uninit. Lemma uninit_shr_add κ l m n : uninit_shr κ l (m + n) 0 ⊣⊢ uninit_shr κ l m 0 ∗ uninit_shr κ (l +ₗ m) n 0. Proof. - move: l. elim m. { move=> ?. by rewrite shift_loc_0 left_id. } + move: l. elim m. { move=> ?. by rewrite shift_loc_0 /= left_id. } move=> ? IH ? /=. rewrite uninit_shr_off IH -assoc. f_equiv. f_equiv; [by rewrite -uninit_shr_off|]. by rewrite shift_loc_assoc -Nat2Z.inj_add. @@ -69,7 +69,7 @@ Section uninit. iDestruct "shr" as "[(%& Bor) shr]". iDestruct "κ" as "[κ κ₊]". iMod (frac_bor_acc with "LFT Bor κ") as (q1) "[>↦ Toκ]"; [solve_ndisj|]. rewrite uninit_shr_off. iMod ("IH" with "shr κ₊") as (q2 vl Eq) "[↦s Toκ₊]". - iModIntro. case (Qp_lower_bound q1 q2) as (q0 &?&?&->&->). + iModIntro. case (Qp.lower_bound q1 q2) as (q0 &?&?&->&->). iExists q0, (_ :: vl)=>/=. iSplit; [by rewrite Eq|]. rewrite heap_mapsto_vec_cons shift_loc_0. iDestruct "↦" as "[$ ↦r]". iDestruct "↦s" as "[$ ↦sr]". iIntros "[↦ ↦s]". @@ -150,8 +150,8 @@ Section typing. iIntros (?) "_!>_". iSplit; [done|]. iSplit; [iApply lft_equiv_refl|]. iSplit; iModIntro=>/=. - iIntros (??? vl). iSplit; last first. - { iIntros ((?&?&->&?&?)) "!%". rewrite app_length. by f_equal. } - iIntros (?) "!%". by apply app_length_ex. + { iIntros ((?&?&->&?&?)) "!%". rewrite length_app. by f_equal. } + iIntros (?) "!%". by apply length_app_ex. - iIntros. rewrite -uninit_shr_add. iSplit; by iIntros. Qed. Lemma uninit_plus_prod_1 E L m n : diff --git a/theories/typing/uniq_array_util.v b/theories/typing/uniq_array_util.v index 0fd1dfe3..35100de4 100644 --- a/theories/typing/uniq_array_util.v +++ b/theories/typing/uniq_array_util.v @@ -23,9 +23,10 @@ Section uniq_array_util. by iMod (bor_create _ _ True%I with "LFT [//]") as "[$ _]". } iDestruct "Bors" as "[Bor Bors]". iDestruct "κ'" as "[κ' κ'₊]". iMod (ty_share_uniq_body with "LFT In In' Bor κ'") as "Upd"; [done|]. - setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "κ'₊ Bors") as "Upd'". - iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ (S _) with "Upd"). - iIntros "!> [>(ξ &$&$) >(ξl &$&$)]". + simpl. iMod "Upd". + setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "κ'₊ Bors") as ">Upd'". + iIntros "!>!>!>". iMod "Upd". iMod "Upd'". iCombine "Upd Upd'" as "Upd". iModIntro. + iApply (step_fupdN_wand with "Upd"). iIntros "[>(ξ &$&$) >(ξl &$&$)]". by iMod (bor_combine with "LFT ξ ξl") as "$". Qed. @@ -43,11 +44,12 @@ Section uniq_array_util. iIntros (?) "#LFT #In #In' uniqs κ'". iInduction vπξil as [|] "IH" forall (q l). { iApply step_fupdN_full_intro. iIntros "!>!>". iExists [], 1%Qp. do 2 (iSplit; [done|]). iIntros. by iFrame. } - iDestruct "uniqs" as "[uniq uniqs]". iDestruct "κ'" as "[κ' κ'₊]"=>/=. + simpl. iDestruct "uniqs" as "[uniq uniqs]". iDestruct "κ'" as "[κ' κ'₊]"=>/=. iMod (ty_own_proph_uniq_body with "LFT In In' uniq κ'") as "Upd"; [done|]. setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "uniqs κ'₊") as "Upd'". - iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ (S _) with "Upd"). - iIntros "!> [>(%&%&%& [ζl ξ] & Touniq) >(%&%&%& [ζl' ξl] & Touniqs)] !>". + simpl. iMod "Upd". iMod "Upd'". iIntros "!>!>!>". iMod "Upd". iMod "Upd'". iModIntro. + iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ _ with "Upd"). + iIntros "[>(%&%&%& [ζl ξ] & Touniq) >(%&%&%& [ζl' ξl] & Touniqs)] !>". iDestruct (proph_tok_combine with "ζl ζl'") as (?) "[ζζl Toζζl]". iDestruct (proph_tok_combine with "ξ ξl") as (?) "[ξl Toξl]". iDestruct (proph_tok_combine with "ζζl ξl") as (?) "[ζζξl Toζζξl]". @@ -73,7 +75,8 @@ Section uniq_array_util. iDestruct "uniqs" as "[uniq uniqs]". iDestruct "L" as "[L L₊]"=>/=. iMod (resolve_uniq_body with "LFT PROPH In E L uniq") as "Upd"; [done..|]. setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "L₊ uniqs") as "Upd'". - iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ (S _) with "Upd"). + simpl. iMod "Upd". iMod "Upd'". iIntros "!>!>!>". iMod "Upd". iMod "Upd'". + iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ _ with "Upd"). iIntros "!> [>[Obs $] >[Obs' $]]". by iCombine "Obs Obs'" as "$". Qed. @@ -92,7 +95,8 @@ Section uniq_array_util. iDestruct "uniqs" as "[uniq uniqs]". iDestruct "L" as "[L L₊]"=>/=. iMod (real_uniq_body with "LFT E L uniq") as "Upd"; [done..|]. setoid_rewrite <-shift_loc_assoc_nat. iMod ("IH" with "L₊ uniqs") as "Upd'". - iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ (S _) with "Upd"). + simpl. iMod "Upd". iMod "Upd'". iIntros "!>!>!>". iMod "Upd". iMod "Upd'". + iCombine "Upd Upd'" as "Upd". iApply (step_fupdN_wand _ _ _ with "Upd"). iIntros "!> [>([%v ->]&$&$) >([%vl %Eq]&$&$)] !%". exists (v ::: vl). fun_ext=>/= ?. by rewrite Eq. Qed. diff --git a/theories/typing/uniq_bor.v b/theories/typing/uniq_bor.v index 458c9517..2465e7dc 100644 --- a/theories/typing/uniq_bor.v +++ b/theories/typing/uniq_bor.v @@ -102,8 +102,8 @@ Section typing. - move=> > ? Hl * /=. f_equiv. + apply equiv_dist. iDestruct Hl as "#[??]". iSplit; iIntros "#H"; (iApply lft_incl_trans; [iApply "H"|done]). - + rewrite /uniq_body. do 18 (f_contractive || f_equiv). by simpl in *. - - move=> */=. do 10 (f_contractive || f_equiv). by simpl in *. + + rewrite /uniq_body. do 18 (f_contractive || f_equiv). by eapply dist_later_dist_lt. + - move=> */=. do 10 (f_contractive || f_equiv). by eapply dist_later_dist_lt. Qed. Global Instance uniq_send {𝔄} κ (ty: type 𝔄) : Send ty → Send (&uniq{κ} ty). diff --git a/theories/typing/uniq_cmra.v b/theories/typing/uniq_cmra.v index 33dc2fb5..60e3e8d7 100644 --- a/theories/typing/uniq_cmra.v +++ b/theories/typing/uniq_cmra.v @@ -22,8 +22,8 @@ Local Definition add_line ξ q vπ d (S: uniq_smryUR) : uniq_smryUR := .<[ξ.(pv_ty) := <[ξ.(pv_id) := item q vπ d]> (S ξ.(pv_ty))]> S. Definition uniqΣ: gFunctors := #[GFunctor uniqUR]. -Class uniqPreG Σ := UniqPreG { uniq_preG_inG :> inG Σ uniqUR }. -Class uniqG Σ := UniqG { uniq_inG :> uniqPreG Σ; uniq_name: gname }. +Class uniqPreG Σ := UniqPreG { uniq_preG_inG :: inG Σ uniqUR }. +Class uniqG Σ := UniqG { uniq_inG :: uniqPreG Σ; uniq_name: gname }. Global Instance subG_uniqPreG Σ : subG uniqΣ Σ → uniqPreG Σ. Proof. solve_inG. Qed. @@ -74,13 +74,13 @@ Proof. iPureIntro. move: Val. rewrite -auth_frag_op auth_frag_valid discrete_fun_singleton_op discrete_fun_singleton_valid singleton_op singleton_valid. - by move/frac_agree_op_valid=> [?[=??]]. + by move/frac_agree_op_valid=> [?[=]]. Qed. Local Lemma vo_vo2 ξ vπ d : .VO[ξ] vπ d ∗ .VO[ξ] vπ d ⊣⊢ .VO2[ξ] vπ d. Proof. by rewrite -own_op -auth_frag_op discrete_fun_singleton_op singleton_op /item - -frac_agree_op Qp_half_half. + -frac_agree_op Qp.half_half. Qed. Local Lemma vo_pc ξ vπ d vπ' d' : @@ -107,7 +107,7 @@ Lemma uniq_intro {𝔄} (vπ: proph 𝔄) d E : let ξ := PrVar (𝔄 ↾ prval_to_inh vπ) ξi in .VO[ξ] vπ d ∗ .PC[ξ] vπ d. Proof. iIntros (?) "PROPH ?". iInv uniqN as (S) ">●S". - set 𝔄i := 𝔄 ↾ prval_to_inh vπ. set I := dom (gset _) (S 𝔄i). + set 𝔄i := 𝔄 ↾ prval_to_inh vπ. set I := dom (S 𝔄i). iMod (proph_intro 𝔄i I with "PROPH") as (ξi NIn) "ξ"; [by solve_ndisj|]. set ξ := PrVar 𝔄i ξi. set S' := add_line ξ 1 vπ d S. move: NIn=> /not_elem_of_dom ?. diff --git a/theories/util/basic.v b/theories/util/basic.v index d9f8ec7d..702bb012 100644 --- a/theories/util/basic.v +++ b/theories/util/basic.v @@ -13,7 +13,7 @@ Lemma semi_iso' {A B} f g `{!@SemiIso A B f g} x : g (f x) = x. Proof. have ->: g (f x) = (g ∘ f) x by done. by rewrite semi_iso. Qed. Class Iso {A B} (f: A → B) (g: B → A) := - { iso_semi_iso_l :> SemiIso f g; iso_semi_iso_r :> SemiIso g f }. + { iso_semi_iso_l :: SemiIso f g; iso_semi_iso_r :: SemiIso g f }. Global Instance iso_id {A} : Iso (@id A) id. Proof. done. Qed. @@ -97,11 +97,11 @@ Proof. split; fun_ext; by [case=> [?[]]|]. Qed. (** * Utility for Lists *) -Lemma app_length_ex {A} (xl: list A) m n : +Lemma length_app_ex {A} (xl: list A) m n : length xl = m + n → ∃yl zl, xl = yl ++ zl ∧ length yl = m ∧ length zl = n. Proof. move=> ?. exists (take m xl), (drop m xl). - rewrite take_length drop_length take_drop. split; [done|lia]. + rewrite length_take length_drop take_drop. split; [done|lia]. Qed. Lemma zip_fst_snd_fun {A B C} (f: C → list (A * B)) : diff --git a/theories/util/fancy_lists.v b/theories/util/fancy_lists.v index 5be053ea..3cb6ad4d 100644 --- a/theories/util/fancy_lists.v +++ b/theories/util/fancy_lists.v @@ -70,9 +70,9 @@ End hlist. Ltac inv_hlist xl := let A := type of xl in match eval hnf in A with hlist _ ?Xl => match eval hnf in Xl with - | [] => revert dependent xl; + | [] => generalize dependent xl; match goal with |- ∀xl, @?P xl => apply (hlist_nil_inv P) end - | _ :: _ => revert dependent xl; + | _ :: _ => generalize dependent xl; match goal with |- ∀xl, @?P xl => apply (hlist_cons_inv P) end; (* Try going on recursively. *) try (let x := fresh "x" in intros x xl; inv_hlist xl; revert x) @@ -371,7 +371,7 @@ Fixpoint pinj {Xl: list A} : ∀i, F (Xl !!ₗ i) → psum Xl := end. Lemma pinj_inj {Xl} i j (x: F (Xl !!ₗ i)) y : - pinj i x = pinj j y → i = j ∧ x ~= y. + pinj i x = pinj j y → i = j ∧ JMeq x y. Proof. move: Xl i j x y. elim; [move=> i; by inv_fin i|]=>/= ?? IH i j. inv_fin i; inv_fin j=>//=. { by move=> ??[=->]. } @@ -381,7 +381,7 @@ Qed. Global Instance pinj_Inj {Xl} i : Inj eq eq (@pinj Xl i). Proof. move: i. elim Xl; [move=> i; by inv_fin i|]=>/= ?? IH i. - inv_fin i. { by move=>/= ??[=?]. } by move=>/= ???[=/IH ?]. + inv_fin i. { by move=>/= ?? [=]. } by move=>/= ???[=/IH ?]. Qed. Fixpoint psum_map {Xl Yl} : @@ -570,7 +570,7 @@ Proof. - rewrite /equiv /hlist_equiv HForallTwo_forall. split=> H; induction H; constructor=>//; by apply equiv_dist. - apply _. - - rewrite /dist /hlist_dist. apply HForallTwo_impl=> >. apply dist_S. + - rewrite /dist /hlist_dist=>??. eapply HForallTwo_impl; [|done]=>/= ????. by eapply dist_lt. Qed. Canonical Structure hlistO := Ofe (hlist F Xl) hlist_ofe_mixin. @@ -671,6 +671,10 @@ Global Instance from_sep_big_sepHL_app {F: A → Type} {Xl Yl} (xl: hlist F Xl) (xl': hlist F Yl) (Φ: ∀C, F C → PROP) : FromSep (big_sepHL Φ (xl h++ xl')) (big_sepHL Φ xl) (big_sepHL Φ xl'). Proof. by rewrite /FromSep big_sepHL_app. Qed. +Global Instance combine_sep_as_big_sepHL_app {F: A → Type} {Xl Yl} + (xl: hlist F Xl) (xl': hlist F Yl) (Φ: ∀C, F C → PROP) : + CombineSepAs (big_sepHL Φ xl) (big_sepHL Φ xl') (big_sepHL Φ (xl h++ xl')). +Proof. by rewrite /CombineSepAs big_sepHL_app. Qed. Global Instance into_sep_big_sepHL_1_app {F: A → Type} {G Xl Yl} (xl: hlist F Xl) (xl': hlist F Yl) (yl: plist G Xl) (yl': plist G Yl) @@ -684,6 +688,13 @@ Global Instance from_sep_big_sepHL_1_app {F: A → Type} {G Xl Yl} FromSep (big_sepHL_1 Φ (xl h++ xl') (yl -++ yl')) (big_sepHL_1 Φ xl yl) (big_sepHL_1 Φ xl' yl'). Proof. by rewrite /FromSep big_sepHL_1_app. Qed. +Global Instance combine_sep_as_big_sepHL_1_app {F: A → Type} {G Xl Yl} + (xl: hlist F Xl) (xl': hlist F Yl) (yl: plist G Xl) (yl': plist G Yl) + (Φ: ∀C, F C → G C → PROP) : + CombineSepAs (big_sepHL_1 Φ xl yl) (big_sepHL_1 Φ xl' yl') + (big_sepHL_1 Φ (xl h++ xl') (yl -++ yl')). +Proof. by rewrite /CombineSepAs big_sepHL_1_app. Qed. + Global Instance frame_big_sepHL_app {F G: A → Type} {Xl Yl} p R Q (xl: hlist F Xl) (xl': hlist F Yl) (Φ: ∀C, F C → PROP) : @@ -709,7 +720,7 @@ Qed. Lemma big_sepHL_2_lookup {F G H: A → Type} {Xl} i (Φ: ∀X, F X → G X → H X → PROP) (xl: hlist F Xl) yl zl : - big_sepHL_2 Φ xl yl zl -∗ Φ _ (xl +!! i) (yl -!! i) (zl -!! i). + big_sepHL_2 Φ xl yl zl ⊢ Φ _ (xl +!! i) (yl -!! i) (zl -!! i). Proof. iIntros "All". iInduction xl as [|] "IH" forall (i); [by inv_fin i|]. move: yl zl=> [??][??]/=. iDestruct "All" as "[Φ All]". diff --git a/theories/util/update.v b/theories/util/update.v index 6b4e1911..4b334384 100644 --- a/theories/util/update.v +++ b/theories/util/update.v @@ -41,6 +41,9 @@ Global Instance step_fupdN_from_sep n P Q E : Proof. rewrite /FromSep. iIntros "[P Q]". iApply (step_fupdN_sep with "P Q"). Qed. +Global Instance step_fupdN_combine_sep_as n P Q E : + CombineSepAs (|={E}▷=>^n P) (|={E}▷=>^n Q) (|={E}▷=>^n (P ∗ Q)). +Proof. apply step_fupdN_from_sep. Qed. Global Instance frame_step_fupdN p R E n P Q : Frame p R P Q → Frame p R (|={E}▷=>^n P) (|={E}▷=>^n Q). @@ -52,8 +55,8 @@ Qed. Lemma step_fupdN_sep_max m n P Q E : (|={E}▷=>^m P) -∗ (|={E}▷=>^n Q) -∗ (|={E}▷=>^(m `max` n) (P ∗ Q)). Proof. - set l := m `max` n. rewrite !(step_fupdN_nmono _ l _); [|lia|lia]. - apply step_fupdN_sep. + set l := m `max` n. iIntros "HP HQ". + iApply (step_fupdN_sep with "[HP] [HQ]"); iApply (step_fupdN_nmono with "[$]"); lia. Qed. Global Instance step_fupdN_from_sep_max m n P Q E : @@ -61,6 +64,9 @@ Global Instance step_fupdN_from_sep_max m n P Q E : Proof. rewrite /FromSep. iIntros "[P Q]". iApply (step_fupdN_sep_max with "P Q"). Qed. +Global Instance step_fupdN_combine_sep_as_max m n P Q E : + CombineSepAs (|={E}▷=>^m P) (|={E}▷=>^n Q) (|={E}▷=>^(m `max` n) (P ∗ Q)) | 10. +Proof. apply step_fupdN_from_sep_max. Qed. Lemma step_fupdN_with_emp n P E F : (|={E, F}=> |={F}▷=>^n |={F, E}=> P) -∗ (|={E, ∅}=> |={∅}▷=>^n |={∅, E}=> P). -- GitLab