diff --git a/opam.pins b/opam.pins
index b0471b4c998052c8460afa0002879fff0185bdab..0939fcaa75a090619c35abf9f81001ce94be4068 100644
--- a/opam.pins
+++ b/opam.pins
@@ -1 +1 @@
-coq-iris https://gitlab.mpi-sws.org/FP/iris-coq 4a154f085ab5be64d30eec3967f1ab46b5467061
+coq-iris https://gitlab.mpi-sws.org/FP/iris-coq a378b8288ebab8d6d16afe0864453731bd753a1d
diff --git a/theories/lang/notation.v b/theories/lang/notation.v
index bb7cadf6bbb1dac3e786f0714c6a6aa58975946b..425839b1c919152cf4cca944fa236ad72cadda8b 100644
--- a/theories/lang/notation.v
+++ b/theories/lang/notation.v
@@ -86,7 +86,7 @@ Notation "'withcont:' k1 ':' e1 'cont:' k2 xl := e2" :=
   ((Lam (@cons binder k1%RustB nil) e1%E) [Rec k2%RustB ((fun _ : eq k1%RustB k2%RustB => xl%RustB) eq_refl) e2%E])
   (only parsing, at level 151, k1, k2, xl at level 1, e2 at level 150) : expr_scope.
 
-Notation "'call:' f args → k" := (f (@cons expr k args))%E
+Notation "'call:' f args → k" := (f (@cons expr (λ: ["_r"], Endlft ;; k ["_r"]) args))%E
   (only parsing, at level 102, f, args, k at level 1) : expr_scope.
 Notation "'letcall:' x := f args 'in' e" :=
   (letcont: "_k" [ x ] := e in call: f args → "_k")%E
diff --git a/theories/lang/tactics.v b/theories/lang/tactics.v
index 73e0cd6f73cb5e33d8b9ab8ff40992117d4203c0..58955702511792aa470d5930ab1328c86eb0ce68 100644
--- a/theories/lang/tactics.v
+++ b/theories/lang/tactics.v
@@ -181,7 +181,7 @@ Ltac solve_to_val :=
   match goal with
   | |- to_val ?e = Some ?v =>
      let e' := W.of_expr e in change (to_val (W.to_expr e') = Some v);
-     apply W.to_val_Some; simpl; unfold W.to_expr; reflexivity
+     apply W.to_val_Some; simpl; unfold W.to_expr; unlock; reflexivity
   | |- is_Some (to_val ?e) =>
      let e' := W.of_expr e in change (is_Some (to_val (W.to_expr e')));
      apply W.to_val_is_Some, (bool_decide_unpack _); vm_compute; exact I
diff --git a/theories/lifetime/lifetime.v b/theories/lifetime/lifetime.v
index 8feb89a6cc3325df7688c904bbb90d0274b2abe1..f1e0c77ef0cb5968d9478c39573cc32cea405b88 100644
--- a/theories/lifetime/lifetime.v
+++ b/theories/lifetime/lifetime.v
@@ -15,6 +15,8 @@ Module Export lifetime : lifetime_sig.
   Include creation.
 End lifetime.
 
+Notation lft_intersect_list l := (foldr lft_intersect static l).
+
 Instance lft_inhabited : Inhabited lft := populate static.
 
 Canonical lftC := leibnizC lft.
@@ -111,4 +113,14 @@ Proof.
   - iIntros (q) "?". iExists 1%Qp. iSplitR. by iApply lft_tok_static. auto.
   - iIntros "Hst". by iDestruct (lft_dead_static with "Hst") as "[]".
 Qed.
+
+Lemma lft_intersect_list_elem_of_incl κs κ :
+  κ ∈ κs → (lft_intersect_list κs ⊑ κ)%I.
+Proof.
+  induction 1 as [|???? IH].
+  - iApply lft_intersect_incl_l.
+  - iApply lft_incl_trans; last iApply IH. (* FIXME: Why does "done" not do this? Looks like "assumption" to me. *)
+    iApply lft_intersect_incl_r.
+Qed.
+  
 End derived.
diff --git a/theories/typing/bool.v b/theories/typing/bool.v
index cb2c8deff443538d4aed5bca4637d81b73e33535..db1825ad8540ed93b416432e38195bc98e22b88a 100644
--- a/theories/typing/bool.v
+++ b/theories/typing/bool.v
@@ -16,6 +16,8 @@ Section bool.
   Next Obligation. intros ? [|[[| |[|[]|]]|] []]; auto. Qed.
   Next Obligation. intros ? [|[[| |[|[]|]]|] []]; apply _. Qed.
 
+  Global Instance bool_wf : TyWf bool := { ty_lfts := []; ty_wf_E := [] }.
+
   Global Instance bool_send : Send bool.
   Proof. done. Qed.
 End bool.
@@ -25,7 +27,7 @@ Section typing.
 
   Lemma type_bool_instr (b : Datatypes.bool) : typed_val #b bool.
   Proof.
-    iIntros (E L tid qE) "_ $ $ $ _". wp_value.
+    iIntros (E L tid) "_ _ $ $ _". wp_value.
     rewrite tctx_interp_singleton tctx_hasty_val. by destruct b.
   Qed.
 
@@ -40,11 +42,11 @@ Section typing.
     typed_body E L C T e1 -∗ typed_body E L C T e2 -∗
     typed_body E L C T (if: p then e1 else e2).
   Proof.
-    iIntros (Hp) "He1 He2". iIntros (tid qE) "#LFT Htl HE HL HC HT".
+    iIntros (Hp) "He1 He2". iIntros (tid) "#LFT #HE Htl HL HC HT".
     iDestruct (big_sepL_elem_of _ _ _ Hp with "HT") as "#Hp".
     wp_bind p. iApply (wp_hasty with "Hp").
     iIntros ([[| |[|[]|]]|]) "_ H1"; try done; (iApply wp_case; [done..|iNext]).
-    - iApply ("He2" with "LFT Htl HE HL HC HT").
-    - iApply ("He1" with "LFT Htl HE HL HC HT").
+    - iApply ("He2" with "LFT HE Htl HL HC HT").
+    - iApply ("He1" with "LFT HE Htl HL HC HT").
   Qed.
 End typing.
diff --git a/theories/typing/borrow.v b/theories/typing/borrow.v
index 4d075647a35339279cd78c07daba29219d4ff5b9..e97f244779f2e1e94b2a2b7abfc2a8b22d48188e 100644
--- a/theories/typing/borrow.v
+++ b/theories/typing/borrow.v
@@ -15,7 +15,7 @@ Section borrow.
   Lemma tctx_borrow E L p n ty κ :
     tctx_incl E L [p ◁ own_ptr n ty] [p ◁ &uniq{κ}ty; p ◁{κ} own_ptr n ty].
   Proof.
-    iIntros (tid ??) "#LFT $ $ H".
+    iIntros (tid ?)  "#LFT _ $ H".
     rewrite /tctx_interp big_sepL_singleton big_sepL_cons big_sepL_singleton.
     iDestruct "H" as ([[]|]) "[% Hown]"; try done. iDestruct "Hown" as "[Hmt ?]".
     iMod (bor_create with "LFT Hmt") as "[Hbor Hext]". done.
@@ -47,7 +47,7 @@ Section borrow.
     lctx_lft_alive E L κ →
     typed_instruction_ty E L [p ◁ &uniq{κ} own_ptr n ty] (!p) (&uniq{κ} ty).
   Proof.
-    iIntros (Hκ tid qE) "#LFT $ HE HL Hp".
+    iIntros (Hκ tid) "#LFT HE $ HL Hp".
     rewrite tctx_interp_singleton.
     iMod (Hκ with "HE HL") as (q) "[Htok Hclose]"; first set_solver.
     wp_apply (wp_hasty with "Hp"). iIntros ([[]|]) "_ Hown"; try done.
@@ -77,7 +77,7 @@ Section borrow.
     lctx_lft_alive E L κ →
     typed_instruction_ty E L [p ◁ &shr{κ} own_ptr n ty] (!p) (&shr{κ} ty).
   Proof.
-    iIntros (Hκ tid qE) "#LFT $ HE HL Hp".
+    iIntros (Hκ tid) "#LFT HE $ HL Hp".
     rewrite tctx_interp_singleton.
     iMod (Hκ with "HE HL") as (q) "[[Htok1 Htok2] Hclose]"; first set_solver.
     wp_apply (wp_hasty with "Hp"). iIntros ([[]|]) "_ Hown"; try done.
@@ -103,10 +103,9 @@ Section borrow.
     lctx_lft_alive E L κ → lctx_lft_incl E L κ κ' →
     typed_instruction_ty E L [p ◁ &uniq{κ} &uniq{κ'} ty] (!p) (&uniq{κ} ty).
   Proof.
-    iIntros (Hκ Hincl tid qE) "#LFT $ HE HL Hp".
+    iIntros (Hκ Hincl tid) "#LFT #HE $ HL Hp".
     rewrite tctx_interp_singleton.
-    iPoseProof (Hincl with "[#] [#]") as "Hincl".
-    { by iApply elctx_interp_persist. } { by iApply llctx_interp_persist. }
+    iDestruct (Hincl with "HL HE") as "#Hincl".
     iMod (Hκ with "HE HL") as (q) "[Htok Hclose]"; first set_solver.
     wp_apply (wp_hasty with "Hp"). iIntros ([[]|]) "_ Hown"; try done.
     iMod (bor_exists with "LFT Hown") as (vl) "Hbor". done.
@@ -137,9 +136,8 @@ Section borrow.
     lctx_lft_alive E L κ → lctx_lft_incl E L κ κ' →
     typed_instruction_ty E L [p ◁ &shr{κ} &uniq{κ'} ty] (!p) (&shr{κ} ty).
   Proof.
-    iIntros (Hκ Hincl tid qE) "#LFT $ HE HL Hp". rewrite tctx_interp_singleton.
-    iPoseProof (Hincl with "[#] [#]") as "Hincl".
-    { by iApply elctx_interp_persist. } { by iApply llctx_interp_persist. }
+    iIntros (Hκ Hincl tid) "#LFT HE $ HL Hp". rewrite tctx_interp_singleton.
+    iDestruct (Hincl with "HL HE") as "#Hincl".
     iMod (Hκ with "HE HL") as (q) "[[Htok1 Htok2] Hclose]"; first set_solver.
     wp_apply (wp_hasty with "Hp"). iIntros ([[]|]) "_ Hshr"; try done.
     iDestruct "Hshr" as (l') "[H↦ Hown]".
diff --git a/theories/typing/cont.v b/theories/typing/cont.v
index 020f883e2ca17db120560445d04b36bb185c964a..916652106e8c9c8dc97ae3eb41581d6c21cac368 100644
--- a/theories/typing/cont.v
+++ b/theories/typing/cont.v
@@ -17,9 +17,9 @@ Section typing.
     tctx_incl E L T (T' (list_to_vec argsv)) →
     typed_body E L C T (k args).
   Proof.
-    iIntros (Hargs HC Hincl tid qE) "#LFT Hna HE HL HC HT".
-    iMod (Hincl with "LFT HE HL HT") as "(HE & HL & HT)".
-    iSpecialize ("HC" with "HE []"); first done.
+    iIntros (Hargs HC Hincl tid) "#LFT #HE Hna HL HC HT".
+    iMod (Hincl with "LFT HE HL HT") as "(HL & HT)".
+    iSpecialize ("HC" with "[]"); first done.
     assert (args = of_val <$> argsv) as ->.
     { clear -Hargs. induction Hargs as [|a av ?? [<-%of_to_val| ->] _ ->]=>//=. }
     rewrite -{3}(vec_to_list_of_list argsv). iApply ("HC" with "Hna HL HT").
@@ -33,15 +33,15 @@ Section typing.
                      (subst_v (kb::argsb) (k:::args) econt)) -∗
     typed_body E L2 C T (letcont: kb argsb := econt in e2).
   Proof.
-    iIntros (Hc1 Hc2) "He2 #Hecont". iIntros (tid qE) "#LFT Htl HE HL HC HT".
+    iIntros (Hc1 Hc2) "He2 #Hecont". iIntros (tid) "#LFT #HE Htl HL HC HT".
     iApply wp_let'; first by rewrite /= decide_left.
-    iModIntro. iApply ("He2" with "LFT Htl HE HL [HC] HT").
-    iIntros "HE". iLöb as "IH". iIntros (x) "H".
-    iDestruct "H" as %[->|?]%elem_of_cons; last by iApply ("HC" with "HE").
+    iModIntro. iApply ("He2" with "LFT HE Htl HL [HC] HT").
+    iLöb as "IH". iIntros (x) "H".
+    iDestruct "H" as %[->|?]%elem_of_cons; last by iApply "HC".
     iIntros (args) "Htl HL HT". iApply wp_rec; try done.
     { rewrite Forall_fmap Forall_forall=>? _. rewrite /= to_of_val. eauto. }
     { by rewrite -(subst_v_eq (_ :: _) (RecV _ _ _ ::: _)). }
-    iNext. iApply ("Hecont" with "LFT Htl HE HL [HC] HT"). by iApply "IH".
+    iNext. iApply ("Hecont" with "LFT HE Htl HL [HC] HT"). by iApply "IH".
   Qed.
 
   Lemma type_cont_norec argsb L1 (T' : vec val (length argsb) → _) E L2 C T econt e2 kb :
@@ -51,14 +51,14 @@ Section typing.
           typed_body E L1 C (T' args) (subst_v (kb :: argsb) (k:::args) econt)) -∗
     typed_body E L2 C T (letcont: kb argsb := econt in e2).
   Proof.
-    iIntros (Hc1 Hc2) "He2 Hecont". iIntros (tid qE) "#LFT Htl HE HL HC HT".
+    iIntros (Hc1 Hc2) "He2 Hecont". iIntros (tid) "#LFT #HE Htl HL HC HT".
     iApply wp_let'; first by rewrite /= decide_left.
-    iModIntro. iApply ("He2" with "LFT Htl HE HL [HC Hecont] HT").
-    iIntros "HE". iIntros (x) "H".
-    iDestruct "H" as %[->|?]%elem_of_cons; last by iApply ("HC" with "HE").
+    iModIntro. iApply ("He2" with "LFT HE Htl HL [HC Hecont] HT").
+    iIntros (x) "H".
+    iDestruct "H" as %[->|?]%elem_of_cons; last by iApply "HC".
     iIntros (args) "Htl HL HT". iApply wp_rec; try done.
     { rewrite Forall_fmap Forall_forall=>? _. rewrite /= to_of_val. eauto. }
     { by rewrite -(subst_v_eq (_ :: _) (RecV _ _ _ ::: _)). }
-    iNext. iApply ("Hecont" with "LFT Htl HE HL HC HT").
+    iNext. iApply ("Hecont" with "LFT HE Htl HL HC HT").
   Qed.
 End typing.
diff --git a/theories/typing/cont_context.v b/theories/typing/cont_context.v
index e09992b70b849675c48af348bcbb2a1229eaa793..5d90ab31139a0fd4e6508c76c90ec08380739c71 100644
--- a/theories/typing/cont_context.v
+++ b/theories/typing/cont_context.v
@@ -20,7 +20,7 @@ Notation cctx := (list cctx_elt).
 Delimit Scope lrust_cctx_scope with CC.
 Bind Scope lrust_cctx_scope with cctx_elt.
 
-Notation "k ◁cont( L , T )" := (CCtx_iscont k L _ T%TC)
+Notation "k ◁cont( L , T )" := (CCtx_iscont k L%LL _ T%TC)
   (at level 70, format "k  ◁cont( L ,  T )") : lrust_cctx_scope.
 Notation "a :: b" := (@cons cctx_elt a%CC b%CC)
   (at level 60, right associativity) : lrust_cctx_scope.
@@ -77,38 +77,35 @@ Section cont_context.
   Qed.
 
   Definition cctx_incl (E : elctx) (C1 C2 : cctx): Prop :=
-    ∀ tid q, lft_ctx -∗ (elctx_interp E q -∗ cctx_interp tid C1) -∗
-                       (elctx_interp E q -∗ cctx_interp tid C2).
+    ∀ tid, lft_ctx -∗ elctx_interp E -∗ cctx_interp tid C1 -∗ cctx_interp tid C2.
   Global Arguments cctx_incl _%EL _%CC _%CC.
 
   Global Instance cctx_incl_preorder E : PreOrder (cctx_incl E).
   Proof.
     split.
-    - iIntros (???) "_ $".
-    - iIntros (??? H1 H2 ??) "#LFT HE".
-      iApply (H2 with "LFT"). by iApply (H1 with "LFT").
+    - iIntros (??) "_ _ $".
+    - iIntros (??? H1 H2 ?) "#LFT #HE HC".
+      iApply (H2 with "LFT HE"). by iApply (H1 with "LFT HE").
   Qed.
 
   Lemma incl_cctx_incl E C1 C2 : C1 ⊆ C2 → cctx_incl E C2 C1.
   Proof.
-    rewrite /cctx_incl /cctx_interp. iIntros (HC1C2 tid ?) "_ H HE * %".
-    iApply ("H" with "HE [%]"). by apply HC1C2.
+    rewrite /cctx_incl /cctx_interp. iIntros (HC1C2 tid) "_ #HE H * %".
+    iApply ("H" with "[%]"). by apply HC1C2.
   Qed.
 
   Lemma cctx_incl_cons_match E k L n (T1 T2 : vec val n → tctx) C1 C2 :
     cctx_incl E C1 C2 → (∀ args, tctx_incl E L (T2 args) (T1 args)) →
     cctx_incl E (k ◁cont(L, T1) :: C1) (k ◁cont(L, T2) :: C2).
   Proof.
-    iIntros (HC1C2 HT1T2 ??) "#LFT H HE". rewrite /cctx_interp.
+    iIntros (HC1C2 HT1T2 ?) "#LFT #HE H". rewrite /cctx_interp.
     iIntros (x) "Hin". iDestruct "Hin" as %[->|Hin]%elem_of_cons.
     - iIntros (args) "Htl HL HT".
-      iMod (HT1T2 with "LFT HE HL HT") as "(HE & HL & HT)".
-      iSpecialize ("H" with "HE").
+      iMod (HT1T2 with "LFT HE HL HT") as "(HL & HT)".
       iApply ("H" $! (_ ◁cont(_, _))%CC with "[%] Htl HL HT").
       constructor.
-    - iApply (HC1C2 with "LFT [H] HE [%]"); last done.
-      iIntros "HE". iIntros (x') "%".
-      iApply ("H" with "HE [%]"). by apply elem_of_cons; auto.
+    - iApply (HC1C2 with "LFT HE [H] [%]"); last done.
+      iIntros (x') "%". iApply ("H" with "[%]"). by apply elem_of_cons; auto.
   Qed.
 
   Lemma cctx_incl_nil E C : cctx_incl E C [].
@@ -121,7 +118,7 @@ Section cont_context.
     cctx_incl E C1 (k ◁cont(L, T2) :: C2).
   Proof.
     intros Hin ??. rewrite <-cctx_incl_cons_match; try done.
-    iIntros (??) "_ HC HE". iSpecialize ("HC" with "HE").
+    iIntros (?) "_ #HE HC".
     rewrite cctx_interp_cons. iSplit; last done. clear -Hin.
     iInduction Hin as [] "IH"; rewrite cctx_interp_cons;
       [iDestruct "HC" as "[$ _]" | iApply "IH"; iDestruct "HC" as "[_ $]"].
diff --git a/theories/typing/examples/get_x.v b/theories/typing/examples/get_x.v
index 91b7f70955d349bc6dbc4b8ac1aa594bf713b95e..0c0d96edc86b3123cd613ed419cf68c9759ba101 100644
--- a/theories/typing/examples/get_x.v
+++ b/theories/typing/examples/get_x.v
@@ -12,12 +12,12 @@ Section get_x.
        delete [ #1; "p"] ;; "return" ["r"].
 
   Lemma get_x_type :
-    typed_val get_x (fn(∀ α, [α]; &uniq{α} Π[int; int]) → &shr{α} int).
+    typed_val get_x (fn(∀ α, ∅; &uniq{α} Π[int; int]) → &shr{α} int).
   (* FIXME: The above is pretty-printed with some explicit scope annotations,
      and without using 'typed_instruction_ty'.  I think that's related to
      the list notation that we added to %TC. *)
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ret p).
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ϝ ret p).
     inv_vec p=>p. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (p'); simpl_subst.
     iApply (type_letalloc_1 (&shr{α}int)); [solve_typing..|]. iIntros (r). simpl_subst.
diff --git a/theories/typing/examples/init_prod.v b/theories/typing/examples/init_prod.v
index 76c97b277491a62475411ea0a84f457410965312..b13fe440196806f1c38bfedbcf66829ee1d02589 100644
--- a/theories/typing/examples/init_prod.v
+++ b/theories/typing/examples/init_prod.v
@@ -12,11 +12,10 @@ Section init_prod.
        "r" +â‚— #0 <- "x'";; "r" +â‚— #1 <- "y'";;
        delete [ #1; "x"] ;; delete [ #1; "y"] ;; return: ["r"].
 
-  Lemma init_prod_type :
-    typed_val init_prod (fn([]; int, int) → Π[int;int]).
+  Lemma init_prod_type : typed_val init_prod (fn(∅; int, int) → Π[int;int]).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros ([] ret p).
-      inv_vec p=>x y. simpl_subst.
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
+      iIntros ([] ϝ ret p). inv_vec p=>x y. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (y'). simpl_subst.
     iApply (type_new_subtype (Π[uninit 1; uninit 1])); [solve_typing..|].
diff --git a/theories/typing/examples/lazy_lft.v b/theories/typing/examples/lazy_lft.v
index c498e9a97ff6ec2be206b08989df14043cdb52a3..12d8c2956bdb76cb7b855ca3091aa544028a62cb 100644
--- a/theories/typing/examples/lazy_lft.v
+++ b/theories/typing/examples/lazy_lft.v
@@ -17,11 +17,10 @@ Section lazy_lft.
       let: "r" := new [ #0] in
       Endlft;; delete [ #2; "t"];; delete [ #1; "f"];; delete [ #1; "g"];; "return" ["r"].
 
-  Lemma lazy_lft_type :
-    typed_val lazy_lft (fn([]) → unit).
+  Lemma lazy_lft_type : typed_val lazy_lft (fn(∅) → unit).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros ([] ret p).
-      inv_vec p. simpl_subst.
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
+      iIntros ([] ϝ ret p). inv_vec p. simpl_subst.
     iApply (type_newlft []). iIntros (α).
     iApply (type_new_subtype (Π[uninit 1;uninit 1])); [solve_typing..|].
       iIntros (t). simpl_subst.
diff --git a/theories/typing/examples/rebor.v b/theories/typing/examples/rebor.v
index f3a30336d872d678ac795f7eac5752dcc2bb4165..a434296c11933105dcd57600436714d10dee3752 100644
--- a/theories/typing/examples/rebor.v
+++ b/theories/typing/examples/rebor.v
@@ -17,10 +17,10 @@ Section rebor.
                  delete [ #1; "x"] ;; "return" ["r"].
 
   Lemma rebor_type :
-    typed_val rebor (fn([]; Π[int; int], Π[int; int]) → int).
+    typed_val rebor (fn(∅; Π[int; int], Π[int; int]) → int).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros ([] ret p).
-       inv_vec p=>t1 t2. simpl_subst.
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
+      iIntros ([] ϝ ret p). inv_vec p=>t1 t2. simpl_subst.
     iApply (type_newlft []). iIntros (α).
     iApply (type_letalloc_1 (&uniq{α}Π[int; int])); [solve_typing..|]. iIntros (x). simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
diff --git a/theories/typing/examples/unbox.v b/theories/typing/examples/unbox.v
index 85fb12ea03f23d528c2ffa38b3e0a55eb2a3d230..586e2aa81acde2265196dfdbda8e103df23dead6 100644
--- a/theories/typing/examples/unbox.v
+++ b/theories/typing/examples/unbox.v
@@ -12,10 +12,10 @@ Section unbox.
        delete [ #1; "b"] ;; "return" ["r"].
 
   Lemma ubox_type :
-    typed_val unbox (fn(∀ α, [α]; &uniq{α}box (Π[int; int])) → &uniq{α} int).
+    typed_val unbox (fn(∀ α, ∅; &uniq{α}box (Π[int; int])) → &uniq{α} int).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ret b).
-      inv_vec b=>b. simpl_subst.
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
+      iIntros (α ϝ ret b). inv_vec b=>b. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (b'); simpl_subst.
     iApply type_deref_uniq_own; [solve_typing..|]. iIntros (bx); simpl_subst.
     iApply type_letalloc_1; [solve_typing..|]. iIntros (r). simpl_subst.
diff --git a/theories/typing/fixpoint.v b/theories/typing/fixpoint.v
index 59bb90bf09df9b9b1e65c53e9352aeae71412d82..847cac84d96b4b11dd024a849054751ac6d5e654 100644
--- a/theories/typing/fixpoint.v
+++ b/theories/typing/fixpoint.v
@@ -16,6 +16,25 @@ Section fixpoint_def.
   Qed.
 
   Definition type_fixpoint : type := fixpointK 2 T.
+
+  (* The procedure for computing ty_lfts and ty_wf_E for fixpoints is
+     the following:
+       - We do a first pass for computing [ty_lfts].
+       - In a second pass, we compute [ty_wf_E], by using the result of the
+         first pass.
+     I believe this gives the right sets for all types that we can define in
+     Rust, but I do not have any proof of this.
+     TODO : investigate this in more detail. *)
+  Global Instance type_fixpoint_wf `{∀ `{!TyWf ty}, TyWf (T ty)} : TyWf type_fixpoint :=
+    let lfts :=
+      let _ : TyWf type_fixpoint := {| ty_lfts := []; ty_wf_E := [] |} in
+      (T type_fixpoint).(ty_lfts)
+    in
+    let wf_E :=
+      let _ : TyWf type_fixpoint := {| ty_lfts := lfts; ty_wf_E := [] |} in
+      (T type_fixpoint).(ty_wf_E)
+    in
+    {| ty_lfts := lfts; ty_wf_E := wf_E |}.
 End fixpoint_def.
 
 Lemma type_fixpoint_ne `{typeG Σ} (T1 T2 : type → type)
@@ -74,7 +93,7 @@ Section fixpoint.
   Proof.
     unfold eqtype, subtype, type_incl.
     setoid_rewrite <-fixpointK_unfold; [| by apply type_contractive_ne, _..].
-    split; iIntros "_ _"; (iSplit; first done); iSplit; iIntros "!#*$".
+    split; iIntros (qL) "_ !# _"; (iSplit; first done); iSplit; iIntros "!#*$".
   Qed.
 End fixpoint.
 
diff --git a/theories/typing/function.v b/theories/typing/function.v
index ca51e24e35fe73c8252265a1893e28b9bc900593..b788a6002b5718dda5c2507d02d3f26030c54008 100644
--- a/theories/typing/function.v
+++ b/theories/typing/function.v
@@ -8,14 +8,19 @@ Set Default Proof Using "Type".
 Section fn.
   Context `{typeG Σ} {A : Type} {n : nat}.
 
-  Record fn_params := FP' { fp_tys : vec type n; fp_ty : type; fp_E : elctx }.
+  Record fn_params := FP { fp_E : lft → elctx; fp_tys : vec type n; fp_ty : type }.
+
+  Definition FP_wf E (tys : vec type n) `{!TyWfLst tys} ty `{!TyWf ty} :=
+    FP (λ ϝ, E ϝ ++ tys.(tyl_wf_E) ++ tys.(tyl_outlives_E) ϝ ++
+                    ty.(ty_wf_E) ++ ty.(ty_outlives_E) ϝ)
+       tys ty.
 
   Program Definition fn (fp : A → fn_params) : type :=
     {| st_own tid vl := (∃ fb kb xb e H,
          ⌜vl = [@RecV fb (kb::xb) e H]⌝ ∗ ⌜length xb = n⌝ ∗
-         ▷ ∀ (x : A) (k : val) (xl : vec val (length xb)),
-            â–¡ typed_body (fp x).(fp_E) []
-                         [k◁cont([], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])]
+         ▷ ∀ (x : A) (ϝ : lft) (k : val) (xl : vec val (length xb)),
+            □ typed_body ((fp x).(fp_E)  ϝ) [ϝ ⊑ []]
+                         [k◁cont([ϝ ⊑ []], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])]
                          (zip_with (TCtx_hasty ∘ of_val) xl
                                    (box <$> (vec_to_list (fp x).(fp_tys))))
                          (subst_v (fb::kb::xb) (RecV fb (kb::xb) e:::k:::xl) e))%I |}.
@@ -23,13 +28,21 @@ Section fn.
     iIntros (fp tid vl) "H". iDestruct "H" as (fb kb xb e ?) "[% _]". by subst.
   Qed.
 
+  (* FIXME : This definition is less restrictive than the one used in
+     Rust. In Rust, the type of parameters are taken into account for
+     well-formedness, and all the liftime constrains relating a
+     generalized liftime are ignored. For simplicity, we ignore all of
+     them, but this is not very faithful. *)
+  Global Instance fn_wf fp : TyWf (fn fp) :=
+    { ty_lfts := []; ty_wf_E := [] }.
+
   Global Instance fn_send fp : Send (fn fp).
   Proof. iIntros (tid1 tid2 vl). done. Qed.
 
   Definition fn_params_rel (ty_rel : relation type) : relation fn_params :=
     λ fp1 fp2,
       Forall2 ty_rel fp2.(fp_tys) fp1.(fp_tys) ∧ ty_rel fp1.(fp_ty) fp2.(fp_ty) ∧
-      fp1.(fp_E) = fp2.(fp_E).
+      pointwise_relation lft eq fp1.(fp_E) fp2.(fp_E).
 
   Global Instance fp_tys_proper R :
     Proper (flip (fn_params_rel R) ==> (Forall2 R : relation (vec _ _))) fp_tys.
@@ -43,11 +56,13 @@ Section fn.
   Proof. intros ?? HR. apply HR. Qed.
 
   Global Instance fp_E_proper R :
-    Proper (fn_params_rel R ==> eq) fp_E.
-  Proof. intros ?? HR. apply HR. Qed.
+    Proper (fn_params_rel R ==> eq ==> eq) fp_E.
+  Proof. intros ?? HR ??->. apply HR. Qed.
 
   Global Instance FP_proper R :
-    Proper (flip (Forall2 R : relation (vec _ _)) ==> R ==> eq ==> fn_params_rel R) FP'.
+    Proper (pointwise_relation lft eq ==>
+            flip (Forall2 R : relation (vec _ _)) ==> R ==>
+            fn_params_rel R) FP.
   Proof. by split; [|split]. Qed.
 
   Global Instance fn_type_contractive n' :
@@ -59,11 +74,10 @@ Section fn.
     (* TODO: 'f_equiv' is slow here because reflexivity is slow. *)
     (* The clean way to do this would be to have a metric on type contexts. Oh well. *)
     intros tid vl. unfold typed_body.
-    do 12 f_equiv. f_contractive. do 17 (apply Hfp || f_equiv).
-    - f_equiv. apply Hfp.
+    do 12 f_equiv. f_contractive. do 16 ((eapply fp_E_proper; try reflexivity) || exact: Hfp || f_equiv).
     - rewrite !cctx_interp_singleton /=. do 5 f_equiv.
       rewrite !tctx_interp_singleton /tctx_elt_interp /=. repeat (apply Hfp || f_equiv).
-    - rewrite /tctx_interp !big_sepL_zip_with /=. do 3 f_equiv.
+    - rewrite /tctx_interp !big_sepL_zip_with /=. do 4 f_equiv.
       cut (∀ n tid p i, Proper (dist n ==> dist n)
         (λ (l : list _), ∀ ty, ⌜l !! i = Some ty⌝ → tctx_elt_interp tid (p ◁ ty))%I).
       { intros Hprop. apply Hprop, list_fmap_ne; last first.
@@ -98,12 +112,6 @@ End fn.
 
 Arguments fn_params {_ _} _.
 
-(* The parameter of [FP'] are in the wrong order in order to make sure
-   that type-checking is done in that order, so that the [ELCtx_Alive]
-   is taken as a coercion. We reestablish the intuitive order with
-   [FP] *)
-Notation FP E tys ty := (FP' tys ty E).
-
 (* We use recursive notation for binders as well, to allow patterns
    like '(a, b) to be used. In practice, only one binder is ever used,
    but using recursive binders is the only way to make Coq accept
@@ -112,69 +120,78 @@ Notation FP E tys ty := (FP' tys ty E).
    printing. Once on 8.6pl1, this should work.  *)
 Notation "'fn(∀' x .. x' ',' E ';' T1 ',' .. ',' TN ')' '→' R" :=
   (fn (λ x, (.. (λ x',
-      FP E%EL (Vector.cons T1 .. (Vector.cons TN Vector.nil) ..)%T R%T)..)))
+      FP_wf E%EL (Vector.cons T1 .. (Vector.cons TN Vector.nil) ..)%T R%T)..)))
   (at level 99, R at level 200, x binder, x' binder,
    format "'fn(∀'  x .. x' ','  E ';'  T1 ','  .. ','  TN ')'  '→'  R") : lrust_type_scope.
 Notation "'fn(∀' x .. x' ',' E ')' '→' R" :=
-  (fn (λ x, (.. (λ x', FP E%EL Vector.nil R%T)..)))
+  (fn (λ x, (.. (λ x', FP_wf E%EL Vector.nil R%T)..)))
   (at level 99, R at level 200, x binder, x' binder,
    format "'fn(∀'  x .. x' ','  E ')'  '→'  R") : lrust_type_scope.
 Notation "'fn(' E ';' T1 ',' .. ',' TN ')' '→' R" :=
-  (fn (λ _:(), FP E%EL (Vector.cons T1 .. (Vector.cons TN Vector.nil) ..) R%T))
+  (fn (λ _:(), FP_wf E%EL (Vector.cons T1 .. (Vector.cons TN Vector.nil) ..) R%T))
   (at level 99, R at level 200,
    format "'fn(' E ';'  T1 ','  .. ','  TN ')'  '→'  R") : lrust_type_scope.
 Notation "'fn(' E ')' '→' R" :=
-  (fn (λ _:(), FP E%EL Vector.nil R%T))
+  (fn (λ _:(), FP_wf E%EL Vector.nil R%T))
   (at level 99, R at level 200,
    format "'fn(' E ')'  '→'  R") : lrust_type_scope.
 
+Instance elctx_empty : Empty (lft → elctx) := λ ϝ, [].
+
 Section typing.
   Context `{typeG Σ}.
 
   Lemma fn_subtype {A n} E0 L0 (fp fp' : A → fn_params n) :
-    (∀ x, elctx_incl E0 L0 (fp' x).(fp_E) (fp x).(fp_E)) →
-    (∀ x, Forall2 (subtype (E0 ++ (fp' x).(fp_E)) L0)
-                  (fp' x).(fp_tys) (fp x).(fp_tys)) →
-    (∀ x, subtype (E0 ++ (fp' x).(fp_E)) L0 (fp x).(fp_ty) (fp' x).(fp_ty)) →
+    (∀ x ϝ, let EE := E0 ++ (fp' x).(fp_E) ϝ in
+            elctx_sat EE L0 ((fp x).(fp_E) ϝ) ∧
+            Forall2 (subtype EE L0) (fp' x).(fp_tys) (fp x).(fp_tys) ∧
+            subtype EE L0 (fp x).(fp_ty) (fp' x).(fp_ty)) →
     subtype E0 L0 (fn fp) (fn fp').
   Proof.
-    intros HE Htys Hty. apply subtype_simple_type=>//= _ vl.
-    iIntros "#HE0 #HL0 Hf". iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst.
+    intros Hcons. apply subtype_simple_type=>//= qL. iIntros "HL0".
+    (* We massage things so that we can throw away HL0 before going under the box. *)
+    iAssert (∀ x ϝ, let EE := E0 ++ (fp' x).(fp_E) ϝ in □ (elctx_interp EE -∗
+                 elctx_interp ((fp x).(fp_E) ϝ) ∗
+                 ([∗ list] tys ∈ (zip (fp' x).(fp_tys) (fp x).(fp_tys)), type_incl (tys.1) (tys.2)) ∗
+                 type_incl (fp x).(fp_ty) (fp' x).(fp_ty)))%I as "#Hcons".
+    { iIntros (x ϝ). destruct (Hcons x ϝ) as (HE &Htys &Hty). clear Hcons.
+      iDestruct (HE with "HL0") as "#HE".
+      iDestruct (subtype_Forall2_llctx with "HL0") as "#Htys"; first done.
+      iDestruct (Hty with "HL0") as "#Hty".
+      iClear "∗". iIntros "!# #HEE".
+      iSplit; last iSplit.
+      - by iApply "HE".
+      - by iApply "Htys".
+      - by iApply "Hty". }
+    iClear "∗". clear Hcons. iIntros "!# #HE0 * Hf".
+    iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst.
     iExists fb, kb, xb, e, _. iSplit. done. iSplit. done. iNext.
-    rewrite /typed_body. iIntros (x k xl) "!# * #LFT Htl HE HL HC HT".
-    iDestruct (elctx_interp_persist with "HE") as "#HEp".
-    iMod (HE with "HE0 HL0 HE") as (qE') "[HE' Hclose]". done.
-    iApply ("Hf" with "LFT Htl HE' HL [HC Hclose] [HT]").
-    - iIntros "HE'". unfold cctx_interp. iIntros (elt) "Helt".
+    rewrite /typed_body. iIntros (x ϝ k xl) "!# * #LFT #HE' Htl HL HC HT".
+    iDestruct ("Hcons" with "[$]") as "#(HE & Htys & Hty)".
+    iApply ("Hf" with "LFT HE Htl HL [HC] [HT]").
+    - unfold cctx_interp. iIntros (elt) "Helt".
       iDestruct "Helt" as %->%elem_of_list_singleton. iIntros (ret) "Htl HL HT".
-      iMod ("Hclose" with "HE'") as "HE".
-      iSpecialize ("HC" with "HE"). unfold cctx_elt_interp.
+      unfold cctx_elt_interp.
       iApply ("HC" $! (_ ◁cont(_, _)%CC) with "[%] Htl HL [> -]").
       { by apply elem_of_list_singleton. }
       rewrite /tctx_interp !big_sepL_singleton /=.
       iDestruct "HT" as (v) "[HP Hown]". iExists v. iFrame "HP".
-      assert (Hst : subtype (E0 ++ (fp' x).(fp_E)) L0
-                            (box (fp x).(fp_ty)) (box (fp' x).(fp_ty)))
-        by by rewrite ->Hty.
-      iDestruct (Hst with "[HE0 HEp] HL0") as "(_ & Hty & _)".
-      { rewrite /elctx_interp_0 big_sepL_app. by iSplit. }
-      by iApply "Hty".
-    - rewrite /tctx_interp
-         -(fst_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?vec_to_list_length //
-         -{1}(snd_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?vec_to_list_length //
+      iDestruct (box_type_incl with "[$Hty]") as "(_ & #Hincl & _)".
+      by iApply "Hincl".
+    - iClear "Hf". rewrite /tctx_interp
+         -{2}(fst_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?vec_to_list_length //
+         -{2}(snd_zip (fp x).(fp_tys) (fp' x).(fp_tys)) ?vec_to_list_length //
          !zip_with_fmap_r !(zip_with_zip (λ _ _, (_ ∘ _) _ _)) !big_sepL_fmap.
-      iApply big_sepL_impl. iSplit; last done. iIntros "{HT}!#".
+      iApply big_sepL_impl. iSplit; last done. iIntros "{HT Hty}!#".
       iIntros (i [p [ty1' ty2']]) "#Hzip H /=".
       iDestruct "H" as (v) "[? Hown]". iExists v. iFrame.
       rewrite !lookup_zip_with.
       iDestruct "Hzip" as %(? & ? & ([? ?] & (? & Hty'1 &
         (? & Hty'2 & [=->->])%bind_Some)%bind_Some & [=->->->])%bind_Some)%bind_Some.
-      eapply Forall2_lookup_lr in Htys; try done.
-      assert (Hst : subtype (E0 ++ (fp' x).(fp_E)) L0 (box ty2') (box ty1'))
-        by by rewrite ->Htys.
-      iDestruct (Hst with "[] []") as "(_ & #Ho & _)"; [ |done|].
-      { rewrite /elctx_interp_0 big_sepL_app. by iSplit. }
-      by iApply "Ho".
+      iDestruct (big_sepL_lookup with "Htys") as "#Hty".
+      { rewrite lookup_zip_with /=. erewrite Hty'2. simpl. by erewrite Hty'1. }
+      iDestruct (box_type_incl with "[$Hty]") as "(_ & #Hincl & _)".
+      by iApply "Hincl".
   Qed.
 
   (* This proper and the next can probably not be inferred, but oh well. *)
@@ -182,8 +199,9 @@ Section typing.
     Proper (pointwise_relation A (fn_params_rel (n:=n) (subtype E0 L0)) ==>
             subtype E0 L0) fn.
   Proof.
-    intros fp1 fp2 Hfp. apply fn_subtype=>x; destruct (Hfp x) as (Htys & Hty & HE).
-    - by rewrite HE.
+    intros fp1 fp2 Hfp. apply fn_subtype=>x ϝ. destruct (Hfp x) as (Htys & Hty & HE).
+    split; last split.
+    - rewrite (HE ϝ). solve_typing.
     - eapply Forall2_impl; first eapply Htys. intros ??.
       eapply subtype_weaken; last done. by apply submseteq_inserts_r.
     - eapply subtype_weaken, Hty; last done. by apply submseteq_inserts_r.
@@ -193,12 +211,12 @@ Section typing.
     Proper (pointwise_relation A (fn_params_rel (n:=n) (eqtype E0 L0)) ==>
             eqtype E0 L0) fn.
   Proof.
-    intros fp1 fp2 Hfp. split; eapply fn_subtype=>x; destruct (Hfp x) as (Htys & Hty & HE).
-    - by rewrite HE.
+    intros fp1 fp2 Hfp. split; eapply fn_subtype=>x ϝ; destruct (Hfp x) as (Htys & Hty & HE); (split; last split).
+    - rewrite (HE ϝ). solve_typing.
     - eapply Forall2_impl; first eapply Htys. intros t1 t2 Ht.
       eapply subtype_weaken; last apply Ht; last done. by apply submseteq_inserts_r.
     - eapply subtype_weaken; last apply Hty; last done. by apply submseteq_inserts_r.
-    - by rewrite HE.
+    - rewrite (HE ϝ). solve_typing.
     - symmetry in Htys. eapply Forall2_impl; first eapply Htys. intros t1 t2 Ht.
       eapply subtype_weaken; last apply Ht; last done. by apply submseteq_inserts_r.
     - eapply subtype_weaken; last apply Hty; last done. by apply submseteq_inserts_r.
@@ -207,28 +225,29 @@ Section typing.
   Lemma fn_subtype_specialize {A B n} (σ : A → B) E0 L0 fp :
     subtype E0 L0 (fn (n:=n) fp) (fn (fp ∘ σ)).
   Proof.
-    apply subtype_simple_type=>//= _ vl.
-    iIntros "_ _ Hf". iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst.
+    apply subtype_simple_type=>//= qL.
+    iIntros "_ !# _ * Hf". iDestruct "Hf" as (fb kb xb e ?) "[% [% #Hf]]". subst.
     iExists fb, kb, xb, e, _. iSplit. done. iSplit. done.
     rewrite /typed_body. iNext. iIntros "*". iApply "Hf".
   Qed.
 
-  Lemma type_call' {A} E L T p (ps : list path)
-                         (fp : A → fn_params (length ps)) k x :
-    elctx_sat E L (fp x).(fp_E) →
+  Lemma type_call' {A} E L T p (κs : list lft) (ps : list path)
+                         (fp : A → fn_params (length ps)) (k : val) x :
+    Forall (lctx_lft_alive E L) κs →
+    (∀ ϝ, elctx_sat (((λ κ, ϝ ⊑ κ) <$> κs)%EL ++ E) L ((fp x).(fp_E) ϝ)) →
     typed_body E L [k ◁cont(L, λ v : vec _ 1, (v!!!0 ◁ box (fp x).(fp_ty)) :: T)]
                ((p ◁ fn fp) ::
                 zip_with TCtx_hasty ps (box <$> (vec_to_list (fp x).(fp_tys))) ++
                 T)
                (call: p ps → k).
   Proof.
-    iIntros (HE tid qE) "#LFT Htl HE HL HC".
-    rewrite tctx_interp_cons tctx_interp_app. iIntros "(Hf & Hargs & HT)".
+    iIntros (Hκs HE tid) "#LFT #HE Htl HL HC (Hf & Hargs & HT)".
     wp_apply (wp_hasty with "Hf"). iIntros (v) "% Hf".
-    iApply (wp_app_vec _ _ (_::_) ((λ v, ⌜v = k⌝):::
+    iApply (wp_app_vec _ _ (_::_) ((λ v, ⌜v = (λ: ["_r"], (#() ;; #()) ;; k ["_r"])%V⌝):::
                vmap (λ ty (v : val), tctx_elt_interp tid (v ◁ box ty)) (fp x).(fp_tys))%I
             with "[Hargs]"); first wp_done.
-    - rewrite /= big_sepL_cons. iSplitR "Hargs"; first by iApply wp_value'.
+    - rewrite /= big_sepL_cons. iSplitR "Hargs".
+      { simpl. iApply wp_value; last done. solve_to_val. }
       clear dependent k p.
       rewrite /tctx_interp vec_to_list_map !zip_with_fmap_r
               (zip_with_zip (λ e ty, (e, _))) zip_with_zip !big_sepL_fmap.
@@ -240,16 +259,25 @@ Section typing.
       iDestruct "EQ" as %[=->]. iDestruct "EQl" as %EQl. revert vl fp HE.
       rewrite <-EQl=>vl fp HE. iApply wp_rec; try done.
       { rewrite -fmap_cons Forall_fmap Forall_forall=>? _. rewrite /= to_of_val. eauto. }
-      { rewrite -fmap_cons -(subst_v_eq (fb::kb::xb) (_:::k:::vl)) //. }
-      iNext. iSpecialize ("Hf" $! x k vl).
-      iMod (HE with "HE HL") as (q') "[HE' Hclose]"; first done.
-      iApply ("Hf" with "LFT Htl HE' [] [HC Hclose HT]").
-      + by rewrite /llctx_interp big_sepL_nil.
-      + iIntros "HE'". iApply fupd_cctx_interp. iMod ("Hclose" with "HE'") as "[HE HL]".
-        iSpecialize ("HC" with "HE"). iModIntro. iIntros (y) "IN".
-        iDestruct "IN" as %->%elem_of_list_singleton. iIntros (args) "Htl _ Hret".
+      { rewrite -fmap_cons -(subst_v_eq (fb::kb::xb) (_:::_:::vl)) //. }
+      iNext. iMod (lft_create with "LFT") as (ϝ) "[Htk #Hinh]"; first done.
+      iSpecialize ("Hf" $! x ϝ _ vl).
+      iDestruct (HE ϝ with "HL") as "#HE'".
+      iMod (lctx_lft_alive_list with "LFT [# $HE //] HL") as "[#HEE Hclose]"; [done..|].
+      iApply ("Hf" with "LFT [>] Htl [Htk] [HC HT Hclose]").
+      + iApply "HE'". iFrame "#". auto.
+      + iSplitL; last done. iExists ϝ. iSplit; first by rewrite /= left_id.
+        iFrame "#∗".
+      + iIntros (y) "IN". iDestruct "IN" as %->%elem_of_list_singleton.
+        iIntros (args) "Htl [HL _] Hret". inv_vec args=>r.
+        iDestruct "HL" as  (κ') "(EQ & Htk & _)". iDestruct "EQ" as %EQ.
+        rewrite /= left_id in EQ. subst κ'. simpl. wp_rec. wp_bind Endlft.
+        iSpecialize ("Hinh" with "Htk").
+        iApply (wp_mask_mono (↑lftN)); first done.
+        iApply (wp_step_fupd with "Hinh"); [set_solver+..|]. wp_seq.
+        iIntros "#Htok !>". wp_seq. iMod ("Hclose" with "Htok") as "HL".
         iSpecialize ("HC" with "[]"); first by (iPureIntro; apply elem_of_list_singleton).
-        iApply ("HC" $! args with "Htl HL").
+        iApply ("HC" $! [#r] with "Htl HL").
         rewrite tctx_interp_singleton tctx_interp_cons. iFrame.
       + rewrite /tctx_interp vec_to_list_map !zip_with_fmap_r
                 (zip_with_zip (λ v ty, (v, _))) zip_with_zip !big_sepL_fmap.
@@ -259,14 +287,15 @@ Section typing.
   Lemma type_call {A} x E L C T T' T'' p (ps : list path)
                         (fp : A → fn_params (length ps)) k :
     (p ◁ fn fp)%TC ∈ T →
-    elctx_sat E L (fp x).(fp_E) →
+    Forall (lctx_lft_alive E L) (L.*1) →
+    (∀ ϝ, elctx_sat (((λ κ, ϝ ⊑ κ) <$> (L.*1))%EL ++ E) L ((fp x).(fp_E) ϝ)) →
     tctx_extract_ctx E L (zip_with TCtx_hasty ps
                                    (box <$> vec_to_list (fp x).(fp_tys))) T T' →
     (k ◁cont(L, T''))%CC ∈ C →
     (∀ ret : val, tctx_incl E L ((ret ◁ box (fp x).(fp_ty))::T') (T'' [# ret])) →
     typed_body E L C T (call: p ps → k).
   Proof.
-    intros Hfn HE HTT' HC HT'T''.
+    intros Hfn HL HE HTT' HC HT'T''.
     rewrite -typed_body_mono /flip; last done; first by eapply type_call'.
     - etrans. eapply (incl_cctx_incl _ [_]); first by intros ? ->%elem_of_list_singleton.
       apply cctx_incl_cons_match; first done. intros args. by inv_vec args.
@@ -278,13 +307,14 @@ Section typing.
                         (fp : A → fn_params (length ps)) b e :
     Closed (b :b: []) e → Closed [] p → Forall (Closed []) ps →
     (p ◁ fn fp)%TC ∈ T →
-    elctx_sat E L (fp x).(fp_E) →
+    Forall (lctx_lft_alive E L) (L.*1) →
+    (∀ ϝ, elctx_sat (((λ κ, ϝ ⊑ κ) <$> (L.*1))%EL ++ E) L ((fp x).(fp_E) ϝ)) →
     tctx_extract_ctx E L (zip_with TCtx_hasty ps
                                    (box <$> vec_to_list (fp x).(fp_tys))) T T' →
     (∀ ret : val, typed_body E L C ((ret ◁ box (fp x).(fp_ty))::T') (subst' b ret e)) -∗
     typed_body E L C T (letcall: b := p ps in e).
   Proof.
-    iIntros (?? Hpsc ???) "He".
+    iIntros (?? Hpsc ????) "He".
     iApply (type_cont_norec [_] _ (λ r, (r!!!0 ◁ box (fp x).(fp_ty)) :: T')%TC).
     - (* TODO : make [solve_closed] work here. *)
       eapply is_closed_weaken; first done. set_solver+.
@@ -295,8 +325,8 @@ Section typing.
         intros. eapply Is_true_eq_true, is_closed_weaken=>//. set_solver+.
     - iIntros (k).
       (* TODO : make [simpl_subst] work here. *)
-      change (subst' "_k" k (p (Var "_k" :: ps))) with
-             ((subst "_k" k p) (of_val k :: map (subst "_k" k) ps)).
+      change (subst' "_k" k (p ((λ: ["_r"], (#() ;; #()) ;; "_k" ["_r"])%E :: ps))) with
+             ((subst "_k" k p) ((λ: ["_r"], (#() ;; #()) ;; k ["_r"])%E :: map (subst "_k" k) ps)).
       rewrite is_closed_nil_subst //.
       assert (map (subst "_k" k) ps = ps) as ->.
       { clear -Hpsc. induction Hpsc=>//=. rewrite is_closed_nil_subst //. congruence. }
@@ -313,22 +343,22 @@ Section typing.
     ef = (funrec: fb argsb := e)%E →
     n = length argsb →
     Closed (fb :b: "return" :b: argsb +b+ []) e →
-    □ (∀ x (f : val) k (args : vec val (length argsb)),
-          typed_body (fp x).(fp_E) []
-                     [k ◁cont([], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])]
+    □ (∀ x ϝ (f : val) k (args : vec val (length argsb)),
+          typed_body ((fp x).(fp_E) ϝ) [ϝ ⊑ []]
+                     [k ◁cont([ϝ ⊑ []], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])]
                      ((f ◁ fn fp) ::
                         zip_with (TCtx_hasty ∘ of_val) args
                                  (box <$> vec_to_list (fp x).(fp_tys)) ++ T)
                      (subst_v (fb :: BNamed "return" :: argsb) (f ::: k ::: args) e)) -∗
     typed_instruction_ty E L T ef (fn fp).
   Proof.
-    iIntros (-> -> Hc) "#Hbody". iIntros (tid qE) " #LFT $ $ $ #HT". iApply wp_value.
+    iIntros (-> -> Hc) "#Hbody". iIntros (tid) "#LFT _ $ $ #HT". iApply wp_value.
     { simpl. rewrite ->(decide_left Hc). done. }
     rewrite tctx_interp_singleton. iLöb as "IH". iExists _. iSplit.
     { simpl. rewrite decide_left. done. }
-    iExists fb, _, argsb, e, _. iSplit. done. iSplit. done. iNext. clear qE.
-    iIntros (x k args) "!#". iIntros (tid' qE) "_ Htl HE HL HC HT'".
-    iApply ("Hbody" with "LFT Htl HE HL HC").
+    iExists fb, _, argsb, e, _. iSplit. done. iSplit. done. iNext.
+    iIntros (x ϝ k args) "!#". iIntros (tid') "_ HE Htl HL HC HT'".
+    iApply ("Hbody" with "LFT HE Htl HL HC").
     rewrite tctx_interp_cons tctx_interp_app. iFrame "HT' IH".
     by iApply sendc_change_tid.
   Qed.
@@ -338,9 +368,9 @@ Section typing.
     ef = (funrec: <> argsb := e)%E →
     n = length argsb →
     Closed ("return" :b: argsb +b+ []) e →
-    □ (∀ x k (args : vec val (length argsb)),
-        typed_body (fp x).(fp_E) []
-                   [k ◁cont([], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])]
+    □ (∀ x ϝ k (args : vec val (length argsb)),
+        typed_body ((fp x).(fp_E) ϝ) [ϝ ⊑ []]
+                   [k ◁cont([ϝ ⊑ []], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])]
                    (zip_with (TCtx_hasty ∘ of_val) args
                              (box <$> vec_to_list (fp x).(fp_tys)) ++ T)
                    (subst_v (BNamed "return" :: argsb) (k ::: args) e)) -∗
diff --git a/theories/typing/int.v b/theories/typing/int.v
index f87c9ac00fcc9ca0cfcc3f8c16bb94ecc88d9429..93e879cea9ae6f6e8aba338e644cc2aadb895f68 100644
--- a/theories/typing/int.v
+++ b/theories/typing/int.v
@@ -15,6 +15,8 @@ Section int.
   Next Obligation. intros ? [|[[]|] []]; auto. Qed.
   Next Obligation. intros ? [|[[]|] []]; apply _. Qed.
 
+  Global Instance int_wf : TyWf int := { ty_lfts := []; ty_wf_E := [] }.
+
   Global Instance int_send : Send int.
   Proof. done. Qed.
 End int.
@@ -24,7 +26,7 @@ Section typing.
 
   Lemma type_int_instr (z : Z) : typed_val #z int.
   Proof.
-    iIntros (E L tid qE) "_ $ $ $ _". wp_value.
+    iIntros (E L tid) "_ _ $ $ _". wp_value.
     by rewrite tctx_interp_singleton tctx_hasty_val.
   Qed.
 
@@ -37,7 +39,7 @@ Section typing.
   Lemma type_plus_instr E L p1 p2 :
     typed_instruction_ty E L [p1 ◁ int; p2 ◁ int] (p1 + p2) int.
   Proof.
-    iIntros (tid qE) "_ $ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
+    iIntros (tid) "_ _ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
     iIntros "[Hp1 Hp2]".
     wp_apply (wp_hasty with "Hp1"). iIntros ([[]|]) "_ H1"; try done.
     wp_apply (wp_hasty with "Hp2"). iIntros ([[]|]) "_ H2"; try done.
@@ -54,7 +56,7 @@ Section typing.
   Lemma type_minus_instr E L p1 p2 :
     typed_instruction_ty E L [p1 ◁ int; p2 ◁ int] (p1 - p2) int.
   Proof.
-    iIntros (tid qE) "_ $ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
+    iIntros (tid) "_ _ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
     iIntros "[Hp1 Hp2]".
     wp_apply (wp_hasty with "Hp1"). iIntros ([[]|]) "_ H1"; try done.
     wp_apply (wp_hasty with "Hp2"). iIntros ([[]|]) "_ H2"; try done.
@@ -71,7 +73,7 @@ Section typing.
   Lemma type_le_instr E L p1 p2 :
     typed_instruction_ty E L [p1 ◁ int; p2 ◁ int] (p1 ≤ p2) bool.
   Proof.
-    iIntros (tid qE) "_ $ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
+    iIntros (tid) "_ _ $ $". rewrite tctx_interp_cons tctx_interp_singleton.
     iIntros "[Hp1 Hp2]".
     wp_apply (wp_hasty with "Hp1"). iIntros ([[]|]) "_ H1"; try done.
     wp_apply (wp_hasty with "Hp2"). iIntros ([[]|]) "_ H2"; try done.
diff --git a/theories/typing/lft_contexts.v b/theories/typing/lft_contexts.v
index 43042ef78ac1ba6185a21b0c590ef6ce9a92439e..5c16ddf563a8ffb4d201b6f695b279951daf6612 100644
--- a/theories/typing/lft_contexts.v
+++ b/theories/typing/lft_contexts.v
@@ -6,9 +6,7 @@ From lrust.typing Require Export base.
 From lrust.lifetime Require Import frac_borrow.
 Set Default Proof Using "Type".
 
-Inductive elctx_elt : Type :=
-| ELCtx_Alive (κ : lft)
-| ELCtx_Incl (κ κ' : lft).
+Definition elctx_elt : Type := lft * lft.
 Notation elctx := (list elctx_elt).
 
 Delimit Scope lrust_elctx_scope with EL.
@@ -17,8 +15,7 @@ Delimit Scope lrust_elctx_scope with EL.
    notations, so we have to use [Arguments] everywhere. *)
 Bind Scope lrust_elctx_scope with elctx_elt.
 
-Coercion ELCtx_Alive : lft >-> elctx_elt.
-Infix "⊑" := ELCtx_Incl (at level 70) : lrust_elctx_scope.
+Notation "κ1 ⊑ κ2" := (@pair lft lft κ1 κ2) (at level 70) : lrust_elctx_scope.
 
 Notation "a :: b" := (@cons elctx_elt a%EL b%EL)
   (at level 60, right associativity) : lrust_elctx_scope.
@@ -47,61 +44,22 @@ Section lft_contexts.
   Implicit Type (κ : lft).
 
   (* External lifetime contexts. *)
-  Definition elctx_elt_interp (x : elctx_elt) (q : Qp) : iProp Σ :=
-    match x with
-    | ELCtx_Alive κ => q.[κ]%I
-    | κ ⊑ κ' => (κ ⊑ κ')%I
-    end%EL.
-  Global Instance elctx_elt_interp_fractional x:
-    Fractional (elctx_elt_interp x).
-  Proof. destruct x; unfold elctx_elt_interp; apply _. Qed.
-  Definition elctx_elt_interp_0 (x : elctx_elt) : iProp Σ :=
-    match x with
-    | ELCtx_Alive κ => True%I
-    | κ ⊑ κ' => (κ ⊑ κ')%I
-    end%EL.
-  Global Instance elctx_elt_interp_0_persistent x:
-    PersistentP (elctx_elt_interp_0 x).
-  Proof. destruct x; apply _. Qed.
-
-  Lemma elctx_elt_interp_persist x q :
-    elctx_elt_interp x q -∗ elctx_elt_interp_0 x.
-  Proof. destruct x; by iIntros "?/=". Qed.
-
-  Definition elctx_interp (E : elctx) (q : Qp) : iProp Σ :=
-    ([∗ list] x ∈ E, elctx_elt_interp x q)%I.
-  Global Arguments elctx_interp _%EL _%Qp.
-  Global Instance elctx_interp_fractional E:
-    Fractional (elctx_interp E).
-  Proof. intros ??. rewrite -big_sepL_sepL. by setoid_rewrite <-fractional. Qed.
-  Global Instance elctx_interp_as_fractional E q:
-    AsFractional (elctx_interp E q) (elctx_interp E) q.
-  Proof. split. done. apply _. Qed.
-  Global Instance elctx_interp_permut:
-    Proper ((≡ₚ) ==> eq ==> (⊣⊢)) elctx_interp.
-  Proof. intros ????? ->. by apply big_opL_permutation. Qed.
-
-  Definition elctx_interp_0 (E : elctx) : iProp Σ :=
-    ([∗ list] x ∈ E, elctx_elt_interp_0 x)%I.
-  Global Arguments elctx_interp_0 _%EL.
-  Global Instance elctx_interp_0_persistent E:
-    PersistentP (elctx_interp_0 E).
-  Proof. apply _. Qed.
-  Global Instance elctx_interp_0_permut:
-    Proper ((≡ₚ) ==> (⊣⊢)) elctx_interp_0.
+  Definition elctx_elt_interp (x : elctx_elt) : iProp Σ :=
+    (x.1 ⊑ x.2)%I.
+
+  Definition elctx_interp (E : elctx) : iProp Σ :=
+    ([∗ list] x ∈ E, elctx_elt_interp x)%I.
+  Global Arguments elctx_interp _%EL.
+  Global Instance elctx_interp_permut :
+    Proper ((≡ₚ) ==> (⊣⊢)) elctx_interp.
   Proof. intros ???. by apply big_opL_permutation. Qed.
-
-  Lemma elctx_interp_persist x q :
-    elctx_interp x q -∗ elctx_interp_0 x.
-  Proof.
-    unfold elctx_interp, elctx_interp_0. f_equiv. intros _ ?.
-    apply elctx_elt_interp_persist.
-  Qed.
+  Global Instance elctx_interp_persistent E :
+    PersistentP (elctx_interp E).
+  Proof. apply _. Qed.
 
   (* Local lifetime contexts. *)
-
   Definition llctx_elt_interp (x : llctx_elt) (q : Qp) : iProp Σ :=
-    let κ' := foldr lft_intersect static (x.2) in
+    let κ' := lft_intersect_list (x.2) in
     (∃ κ0, ⌜x.1 = κ' ⊓ κ0⌝ ∗ q.[κ0] ∗ □ (1.[κ0] ={↑lftN,∅}▷=∗ [†κ0]))%I.
   Global Instance llctx_elt_interp_fractional x :
     Fractional (llctx_elt_interp x).
@@ -116,41 +74,22 @@ Section lft_contexts.
       iExists κ0. by iFrame "∗%".
   Qed.
 
-  Definition llctx_elt_interp_0 (x : llctx_elt) : Prop :=
-    let κ' := foldr lft_intersect static (x.2) in (∃ κ0, x.1 = κ' ⊓ κ0).
-  Lemma llctx_elt_interp_persist x q :
-    llctx_elt_interp x q -∗ ⌜llctx_elt_interp_0 x⌝.
-  Proof.
-    iIntros "H". unfold llctx_elt_interp.
-    iDestruct "H" as (κ0) "[% _]". by iExists κ0.
-  Qed.
-
   Definition llctx_interp (L : llctx) (q : Qp) : iProp Σ :=
     ([∗ list] x ∈ L, llctx_elt_interp x q)%I.
   Global Arguments llctx_interp _%LL _%Qp.
-  Global Instance llctx_interp_fractional L:
+  Global Instance llctx_interp_fractional L :
     Fractional (llctx_interp L).
   Proof. intros ??. rewrite -big_sepL_sepL. by setoid_rewrite <-fractional. Qed.
-  Global Instance llctx_interp_as_fractional L q:
+  Global Instance llctx_interp_as_fractional L q :
     AsFractional (llctx_interp L q) (llctx_interp L) q.
   Proof. split. done. apply _. Qed.
-  Global Instance llctx_interp_permut:
+  Global Instance llctx_interp_permut :
     Proper ((≡ₚ) ==> eq ==> (⊣⊢)) llctx_interp.
   Proof. intros ????? ->. by apply big_opL_permutation. Qed.
 
-  Definition llctx_interp_0 (L : llctx) : Prop :=
-    ∀ x, x ∈ L → llctx_elt_interp_0 x.
-  Global Arguments llctx_interp_0 _%LL.
-  Lemma llctx_interp_persist L q :
-    llctx_interp L q -∗ ⌜llctx_interp_0 L⌝.
-  Proof.
-    iIntros "H". iIntros (x ?). iApply llctx_elt_interp_persist.
-    unfold llctx_interp. by iApply (big_sepL_elem_of with "H").
-  Qed.
-
-  Lemma lctx_equalize_lft qE qL κ1 κ2 `{!frac_borG Σ} :
+  Lemma lctx_equalize_lft qL κ1 κ2 `{!frac_borG Σ} :
     lft_ctx -∗ llctx_elt_interp (κ1 ⊑ [κ2]%list) qL ={⊤}=∗
-      elctx_elt_interp (κ1 ⊑ κ2) qE ∗ elctx_elt_interp (κ2 ⊑ κ1) qE.
+      elctx_elt_interp (κ1 ⊑ κ2) ∗ elctx_elt_interp (κ2 ⊑ κ1).
   Proof.
     iIntros "#LFT Hκ". rewrite /llctx_elt_interp /=. (* TODO: Why is this unfold necessary? *)
     iDestruct "Hκ" as (κ) "(% & Hκ & _)".
@@ -169,21 +108,26 @@ Section lft_contexts.
   Context (E : elctx) (L : llctx).
 
   (* Lifetime inclusion *)
-
   Definition lctx_lft_incl κ κ' : Prop :=
-    elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗ κ ⊑ κ'.
+    ∀ qL, llctx_interp L qL -∗ □ (elctx_interp E -∗ κ ⊑ κ').
 
   Definition lctx_lft_eq κ κ' : Prop :=
     lctx_lft_incl κ κ' ∧ lctx_lft_incl κ' κ.
 
+  Lemma lctx_lft_incl_incl κ κ' : lctx_lft_incl κ κ' → lctx_lft_incl κ κ'.
+  Proof. done. Qed.
+
   Lemma lctx_lft_incl_relf κ : lctx_lft_incl κ κ.
-  Proof. iIntros "_ _". iApply lft_incl_refl. Qed.
+  Proof. iIntros (qL) "_ !# _". iApply lft_incl_refl. Qed.
 
   Global Instance lctx_lft_incl_preorder : PreOrder lctx_lft_incl.
   Proof.
     split; first by intros ?; apply lctx_lft_incl_relf.
-    iIntros (??? H1 H2) "#HE #HL". iApply (lft_incl_trans with "[#] [#]").
-    iApply (H1 with "HE HL"). iApply (H2 with "HE HL").
+    iIntros (??? H1 H2 ?) "HL".
+    iDestruct (H1 with "HL") as "#H1".
+    iDestruct (H2 with "HL") as "#H2".
+    iClear "∗". iIntros "!# #HE". iApply lft_incl_trans.
+    by iApply "H1". by iApply "H2".
   Qed.
 
   Global Instance lctx_lft_incl_proper :
@@ -199,17 +143,17 @@ Section lft_contexts.
   Qed.
 
   Lemma lctx_lft_incl_static κ : lctx_lft_incl κ static.
-  Proof. iIntros "_ _". iApply lft_incl_static. Qed.
+  Proof. iIntros (qL) "_ !# _". iApply lft_incl_static. Qed.
 
   Lemma lctx_lft_incl_local κ κ' κs :
     (κ ⊑ κs)%LL ∈ L → κ' ∈ κs → lctx_lft_incl κ κ'.
   Proof.
-    iIntros (? Hκ'κs) "_ H". iDestruct "H" as %HL.
-    edestruct HL as [κ0 EQ]. done. simpl in EQ; subst.
+    iIntros (? Hκ'κs qL) "H".
+    iDestruct (big_sepL_elem_of with "H") as "H"; first done.
+    iDestruct "H" as (κ'') "[EQ _]". iDestruct "EQ" as %EQ.
+    simpl in EQ; subst. iIntros "!# #HE".
     iApply lft_incl_trans; first iApply lft_intersect_incl_l.
-    clear -Hκ'κs. induction Hκ'κs.
-    - iApply lft_intersect_incl_l.
-    - iApply lft_incl_trans; last done. iApply lft_intersect_incl_r.
+    by iApply lft_intersect_list_elem_of_incl.
   Qed.
 
   Lemma lctx_lft_incl_local' κ κ' κ'' κs :
@@ -218,8 +162,8 @@ Section lft_contexts.
 
   Lemma lctx_lft_incl_external κ κ' : (κ ⊑ κ')%EL ∈ E → lctx_lft_incl κ κ'.
   Proof.
-    iIntros (?) "H _".
-    rewrite /elctx_interp_0 /elctx_elt_interp_0 big_sepL_elem_of //. done.
+    iIntros (??) "_ !# #HE".
+    rewrite /elctx_interp /elctx_elt_interp big_sepL_elem_of //. done.
   Qed.
 
   Lemma lctx_lft_incl_external' κ κ' κ'' :
@@ -229,241 +173,142 @@ Section lft_contexts.
   (* Lifetime aliveness *)
 
   Definition lctx_lft_alive (κ : lft) : Prop :=
-    ∀ F qE qL, ↑lftN ⊆ F → elctx_interp E qE -∗ llctx_interp L qL ={F}=∗
-          ∃ q', q'.[κ] ∗ (q'.[κ] ={F}=∗ elctx_interp E qE ∗ llctx_interp L qL).
+    ∀ F qL, ↑lftN ⊆ F → elctx_interp E -∗ llctx_interp L qL ={F}=∗
+          ∃ q', q'.[κ] ∗ (q'.[κ] ={F}=∗ llctx_interp L qL).
+
+  Lemma lctx_lft_alive_tok κ F q :
+    ↑lftN ⊆ F →
+    lctx_lft_alive κ → elctx_interp E -∗ llctx_interp L q ={F}=∗
+      ∃ q', q'.[κ] ∗ llctx_interp L q' ∗
+                   (q'.[κ] -∗ llctx_interp L q' ={F}=∗ llctx_interp L q).
+  Proof.
+    iIntros (? Hal) "#HE [HL1 HL2]".
+    iMod (Hal with "HE HL1") as (q') "[Htok Hclose]"; first done.
+    destruct (Qp_lower_bound (q/2) q') as (qq & q0  & q'0 & Hq & ->). rewrite Hq.
+    iExists qq. iDestruct "HL2" as "[$ HL]". iDestruct "Htok" as "[$ Htok]".
+    iIntros "!> Htok' HL'". iCombine "HL'" "HL" as "HL". rewrite -Hq. iFrame.
+    iApply "Hclose". iFrame.
+  Qed.
 
   Lemma lctx_lft_alive_static : lctx_lft_alive static.
   Proof.
-    iIntros (F qE qL ?) "$$". iExists 1%Qp. iSplitL. by iApply lft_tok_static. auto.
+    iIntros (F qL ?) "_ $". iExists 1%Qp. iSplitL. by iApply lft_tok_static. auto.
   Qed.
 
   Lemma lctx_lft_alive_local κ κs:
     (κ ⊑ κs)%LL ∈ L → Forall lctx_lft_alive κs → lctx_lft_alive κ.
   Proof.
-    iIntros ([i HL]%elem_of_list_lookup_1 Hκs F qE qL ?) "HE HL".
+    iIntros ([i HL]%elem_of_list_lookup_1 Hκs F qL ?) "#HE HL".
     iDestruct "HL" as "[HL1 HL2]". rewrite {2}/llctx_interp /llctx_elt_interp.
     iDestruct (big_sepL_lookup_acc with "HL2") as "[Hκ Hclose]". done.
     iDestruct "Hκ" as (κ0) "(EQ & Htok & #Hend)". simpl. iDestruct "EQ" as %->.
     iAssert (∃ q', q'.[foldr lft_intersect static κs] ∗
-      (q'.[foldr lft_intersect static κs] ={F}=∗ elctx_interp E qE ∗ llctx_interp L (qL / 2)))%I
+      (q'.[foldr lft_intersect static κs] ={F}=∗ llctx_interp L (qL / 2)))%I
       with "[> HE HL1]" as "H".
     { move:(qL/2)%Qp=>qL'. clear HL. iClear "Hend".
-      iInduction Hκs as [|κ κs Hκ ?] "IH" forall (qE qL').
+      iInduction Hκs as [|κ κs Hκ ?] "IH" forall (qL').
       - iExists 1%Qp. iFrame. iSplitR; last by auto. iApply lft_tok_static.
-      - iDestruct "HL1" as "[HL1 HL2]". iDestruct "HE" as "[HE1 HE2]".
-        iMod (Hκ with "HE1 HL1") as (q') "[Htok' Hclose]". done.
-        iMod ("IH" with "HE2 HL2") as (q'') "[Htok'' Hclose']".
+      - 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 & -> & ->).
         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. }
+        iMod ("Hclose" with "[$Hκ $Hr']") as "$". iApply "Hclose'". iFrame. }
     iDestruct "H" as (q') "[Htok' Hclose']". rewrite -{5}(Qp_div_2 qL).
     destruct (Qp_lower_bound q' (qL/2)) as (q0 & q'2 & q''2 & -> & ->).
     iExists q0. rewrite -(lft_tok_sep q0). iDestruct "Htok" as "[$ Htok]".
     iDestruct "Htok'" as "[$ Htok']". iIntros "!>[Hfold Hκ0]".
-    iMod ("Hclose'" with "[$Hfold $Htok']") as "[$$]".
+    iMod ("Hclose'" with "[$Hfold $Htok']") as "$".
     rewrite /llctx_interp /llctx_elt_interp. iApply "Hclose". iExists κ0. iFrame. auto.
   Qed.
 
-  Lemma lctx_lft_alive_external κ: (ELCtx_Alive κ) ∈ E → lctx_lft_alive κ.
-  Proof.
-    iIntros ([i HE]%elem_of_list_lookup_1 F qE qL ?) "HE $ !>".
-    iDestruct (big_sepL_lookup_acc with "HE") as "[Hκ Hclose]". done.
-    iExists qE. iFrame. iIntros "?!>". by iApply "Hclose".
-  Qed.
-
   Lemma lctx_lft_alive_incl κ κ':
     lctx_lft_alive κ → lctx_lft_incl κ κ' → lctx_lft_alive κ'.
   Proof.
-    iIntros (Hal Hinc F qE qL ?) "HE HL".
-    iAssert (κ ⊑ κ')%I with "[#]" as "#Hincl". iApply (Hinc with "[HE] [HL]").
-      by iApply elctx_interp_persist. by iApply llctx_interp_persist.
+    iIntros (Hal Hinc F qL ?) "#HE HL".
+    iAssert (κ ⊑ κ')%I with "[#]" as "#Hincl". iApply (Hinc with "HL HE").
     iMod (Hal with "HE HL") as (q') "[Htok Hclose]". done.
     iMod (lft_incl_acc with "Hincl Htok") as (q'') "[Htok Hclose']". done.
     iExists q''. iIntros "{$Htok}!>Htok". iApply ("Hclose" with "[> -]").
     by iApply "Hclose'".
   Qed.
 
-  Lemma lctx_lft_alive_external' κ κ':
-    (ELCtx_Alive κ) ∈ E → (κ ⊑ κ')%EL ∈ E → lctx_lft_alive κ'.
+  Lemma lctx_lft_alive_external κ κ':
+    (κ ⊑ κ')%EL ∈ E → lctx_lft_alive κ → lctx_lft_alive κ'.
+  Proof.
+    intros. by eapply lctx_lft_alive_incl, lctx_lft_incl_external.
+  Qed.
+
+  Lemma lctx_lft_alive_list κs ϝ `{!frac_borG Σ} :
+    Forall lctx_lft_alive κs →
+    ∀ (F : coPset) (qL : Qp),
+      ↑lftN ⊆ F → lft_ctx -∗ elctx_interp E -∗
+         llctx_interp L qL ={F}=∗ elctx_interp ((λ κ, ϝ ⊑ κ) <$> κs)%EL ∗
+                                   ([†ϝ] ={F}=∗ llctx_interp L qL).
   Proof.
-    intros. eapply lctx_lft_alive_incl, lctx_lft_incl_external; last done.
-    by apply lctx_lft_alive_external.
+    iIntros (Hκs F qL ?) "#LFT #HE". iInduction κs as [|κ κs] "IH" forall (qL Hκs).
+    { iIntros "$ !>". rewrite /elctx_interp big_sepL_nil. auto. }
+    iIntros "[HL1 HL2]". inversion Hκs as [|?? Hκ Hκs']. clear Hκs. subst.
+    iMod ("IH" with "[% //] HL2") as "[#Hκs Hclose1] {IH}".
+    iMod (Hκ with "HE HL1") as (q) "[Hκ Hclose2]"; first done.
+    iMod (bor_create with "LFT [Hκ]") as "[Hκ H†]"; first done. iApply "Hκ".
+    iDestruct (frac_bor_lft_incl _ _ q with "LFT [> Hκ]") as "#Hincl".
+    { iApply (bor_fracture with "LFT [>-]"); first done. simpl.
+      rewrite Qp_mult_1_r. done. (* FIXME: right_id? *) }
+    iModIntro. iFrame "#". iIntros "#Hϝ†".
+    iMod ("H†" with "Hϝ†") as ">Hκ". iMod ("Hclose2" with "Hκ") as "$".
+    by iApply "Hclose1".
   Qed.
 
   (* External lifetime context satisfiability *)
 
   Definition elctx_sat E' : Prop :=
-    ∀ qE qL F, ↑lftN ⊆ F → elctx_interp E qE -∗ llctx_interp L qL ={F}=∗
-      ∃ qE', elctx_interp E' qE' ∗
-       (elctx_interp E' qE' ={F}=∗ elctx_interp E qE ∗ llctx_interp L qL).
+    ∀ qL, llctx_interp L qL -∗ □ (elctx_interp E -∗ elctx_interp E').
 
   Lemma elctx_sat_nil : elctx_sat [].
   Proof.
-    iIntros (qE qL F ?) "$$". iExists 1%Qp. rewrite /elctx_interp big_sepL_nil. auto.
-  Qed.
-
-  Lemma elctx_sat_alive E' κ :
-    lctx_lft_alive κ → elctx_sat E' → elctx_sat (κ :: E')%EL.
-  Proof.
-    iIntros (Hκ HE' qE qL F ?) "[HE1 HE2] [HL1 HL2]".
-    iMod (Hκ with "HE1 HL1") as (q) "[Htok Hclose]". done.
-    iMod (HE' with "HE2 HL2") as (q') "[HE' Hclose']". done.
-    destruct (Qp_lower_bound q q') as (q0 & q2 & q'2 & -> & ->). iExists q0.
-    rewrite {5 6}/elctx_interp big_sepL_cons /=.
-    iDestruct "Htok" as "[$ Htok]". iDestruct "HE'" as "[Hf HE']".
-    iSplitL "Hf". by rewrite /elctx_interp.
-    iIntros "!>[Htok' ?]". iMod ("Hclose" with "[$Htok $Htok']") as "[$$]".
-    iApply "Hclose'". iFrame.
+    iIntros (?) "_ !# _". rewrite /elctx_interp big_sepL_nil. auto.
   Qed.
 
   Lemma elctx_sat_lft_incl E' κ κ' :
     lctx_lft_incl κ κ' → elctx_sat E' → elctx_sat ((κ ⊑ κ') :: E')%EL.
   Proof.
-    iIntros (Hκκ' HE' qE qL F ?) "HE HL".
-    iAssert (κ ⊑ κ')%I with "[#]" as "#Hincl". iApply (Hκκ' with "[HE] [HL]").
-      by iApply elctx_interp_persist. by iApply llctx_interp_persist.
-    iMod (HE' with "HE HL") as (q) "[HE' Hclose']". done.
-    iExists q. iFrame "HE'". iIntros "{$Hincl}!>[_ ?]". by iApply "Hclose'".
+    iIntros (Hκκ' HE' qL) "HL".
+    iDestruct (Hκκ' with "HL") as "#Hincl".
+    iDestruct (HE' with "HL") as "#HE'".
+    iClear "∗". iIntros "!# #HE".
+    (* FIXME: Why does iSplit fail here?  The goal is persistent. *)
+    iSplitL.
+    - by iApply "Hincl".
+    - by iApply "HE'".
   Qed.
 
   Lemma elctx_sat_app E1 E2 :
     elctx_sat E1 → elctx_sat E2 → elctx_sat (E1 ++ E2).
   Proof.
-    iIntros (HE1 HE2 ????) "[HE1 HE2] [HL1 HL2]".
-    iMod (HE1 with "HE1 HL1") as (qE1) "[HE1 Hclose1]". done.
-    iMod (HE2 with "HE2 HL2") as (qE2) "[HE2 Hclose2]". done.
-    destruct (Qp_lower_bound qE1 qE2) as (q & qE1' & qE2' & -> & ->). iExists q.
-    rewrite !fractional [elctx_interp E qE]fractional_half /elctx_interp big_sepL_app.
-    iDestruct "HE1" as "[$ HE1]". iDestruct "HE2" as "[$ HE2]". iIntros "!> [Hq1 Hq2]".
-    iMod ("Hclose1" with "[$HE1 $Hq1]") as "[$ $]". iApply "Hclose2". iFrame.
+    iIntros (HE1 HE2 ?) "HL".
+    iDestruct (HE1 with "HL") as "#HE1".
+    iDestruct (HE2 with "HL") as "#HE2".
+    iClear "∗". iIntros "!# #HE".
+    iDestruct ("HE1" with "HE") as "#$".
+    iApply ("HE2" with "HE").
   Qed.
 
   Lemma elctx_sat_refl : elctx_sat E.
-  Proof. iIntros (????) "??". iExists _. eauto with iFrame. Qed.
+  Proof. iIntros (?) "_ !# ?". done. Qed.
 End lft_contexts.
 
 Arguments lctx_lft_incl {_ _ _} _%EL _%LL _ _.
 Arguments lctx_lft_eq {_ _ _} _%EL _%LL _ _.
 Arguments lctx_lft_alive {_ _ _} _%EL _%LL _.
 Arguments elctx_sat {_ _ _} _%EL _%LL _%EL.
+Arguments lctx_lft_incl_incl {_ _ _ _%EL _%LL} _ _.
+Arguments lctx_lft_alive_tok {_ _ _ _%EL _%LL} _ _ _.
 
-Lemma elctx_sat_cons_weaken `{invG Σ, lftG Σ} e0 E E' L :
-  elctx_sat E L E' → elctx_sat (e0 :: E) L E'.
-Proof.
-  iIntros (HE' ????). rewrite /elctx_interp. setoid_rewrite big_sepL_cons.
-  iIntros "[? HE] HL". iMod (HE' with "HE HL") as (?) "[H Hclose]". done.
-  auto 10 with iFrame.
-Qed.
-
-Lemma elctx_sat_app_weaken_l `{invG Σ, lftG Σ} E0 E E' L :
-  elctx_sat E L E' → elctx_sat (E0 ++ E) L E'.
-Proof. induction E0=>//= ?. auto using elctx_sat_cons_weaken. Qed.
-
-Lemma elctx_sat_app_weaken_r `{invG Σ, lftG Σ} E0 E E' L :
-  elctx_sat E L E' → elctx_sat (E ++ E0) L E'.
-Proof.
-  intros ?????. rewrite /elctx_sat /elctx_interp.
-  setoid_rewrite (big_opL_permutation _ (E++E0)); try apply Permutation_app_comm.
-  by apply elctx_sat_app_weaken_l.
-Qed.
-
-Section elctx_incl.
-  (* External lifetime context inclusion, in a persistent context.
-     (Used for function types subtyping).
-     On paper, this corresponds to using only the inclusions from E, L
-     to show E1; [] |- E2.
-  *)
-
-  Context `{invG Σ, lftG Σ} (E : elctx) (L : llctx).
-
-  Definition elctx_incl E1 E2 : Prop :=
-    ∀ F, ↑lftN ⊆ F → elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗
-    ∀ qE1, elctx_interp E1 qE1 ={F}=∗ ∃ qE2, elctx_interp E2 qE2 ∗
-       (elctx_interp E2 qE2 ={F}=∗ elctx_interp E1 qE1).
-  Global Instance : RewriteRelation elctx_incl.
-  Arguments elctx_incl _%EL _%EL.
-
-  Global Instance elctx_incl_preorder : PreOrder elctx_incl.
-  Proof.
-    split.
-    - iIntros (???) "_ _ * ?". iExists _. iFrame. eauto.
-    - iIntros (x y z Hxy Hyz ??) "#HE #HL * HE1".
-      iMod (Hxy with "HE HL HE1") as (qy) "[HE1 Hclose1]"; first done.
-      iMod (Hyz with "HE HL HE1") as (qz) "[HE1 Hclose2]"; first done.
-      iModIntro. iExists qz. iFrame "HE1". iIntros "HE1".
-      iApply ("Hclose1" with "[> -]"). by iApply "Hclose2".
-  Qed.
-
-  Lemma elctx_incl_refl E' : elctx_incl E' E'.
-  Proof. reflexivity. Qed.
-
-  Lemma elctx_incl_nil E' : elctx_incl E' [].
-  Proof.
-    iIntros (??) "_ _ * $". iExists 1%Qp.
-    rewrite /elctx_interp big_sepL_nil. auto.
-  Qed.
-
-  Global Instance elctx_incl_app_proper :
-    Proper (elctx_incl ==> elctx_incl ==> elctx_incl) app.
-  Proof.
-    iIntros (?? HE1 ?? HE2 ??) "#HE #HL *". rewrite {1}/elctx_interp big_sepL_app.
-    iIntros "[HE1 HE2]".
-    iMod (HE1 with "HE HL HE1") as (q1) "[HE1 Hclose1]"; first done.
-    iMod (HE2 with "HE HL HE2") as (q2) "[HE2 Hclose2]"; first done.
-    destruct (Qp_lower_bound q1 q2) as (q & ? & ? & -> & ->).
-    iModIntro. iExists q. rewrite /elctx_interp !big_sepL_app.
-    iDestruct "HE1" as "[$ HE1]". iDestruct "HE2" as "[$ HE2]".
-    iIntros "[HE1' HE2']".
-    iMod ("Hclose1" with "[$HE1 $HE1']") as "$".
-    iMod ("Hclose2" with "[$HE2 $HE2']") as "$".
-    done.
-  Qed.
-  Global Instance elctx_incl_cons_proper x :
-    Proper (elctx_incl ==> elctx_incl) (cons x).
-  Proof. by apply (elctx_incl_app_proper [_] [_]). Qed.
-
-  Lemma elctx_incl_dup E1 :
-    elctx_incl E1 (E1 ++ E1).
-  Proof.
-    iIntros (??) "#HE #HL * [HE1 HE1']".
-    iModIntro. iExists (qE1 / 2)%Qp.
-    rewrite /elctx_interp !big_sepL_app. iFrame.
-    iIntros "[HE1 HE1']". by iFrame.
-  Qed.
-
-  Lemma elctx_sat_incl E1 E2 :
-    elctx_sat E1 [] E2 → elctx_incl E1 E2.
-   Proof.
-    iIntros (H12 ??) "#HE #HL * HE1".
-    iMod (H12 _ 1%Qp with "HE1 []") as (qE2) "[HE2 Hclose]". done.
-    { by rewrite /llctx_interp big_sepL_nil. }
-    iExists qE2. iFrame. iIntros "!> HE2".
-    by iMod ("Hclose" with "HE2") as "[$ _]".
-   Qed.
-
-  Lemma elctx_incl_lft_alive E1 E2 κ :
-    lctx_lft_alive E1 [] κ → elctx_incl E1 E2 → elctx_incl E1 (κ :: E2).
-  Proof.
-    intros Hκ HE2. rewrite (elctx_incl_dup E1).
-    apply (elctx_incl_app_proper _ [_]); last done.
-    apply elctx_sat_incl. apply elctx_sat_alive, elctx_sat_nil; first done.
-  Qed.
-
-  Lemma elctx_incl_lft_incl E1 E2 κ κ' :
-    lctx_lft_incl (E ++ E1) L κ κ' → elctx_incl E1 E2 →
-    elctx_incl E1 ((κ ⊑ κ') :: E2).
-  Proof.
-    iIntros (Hκκ' HE2 ??) "#HE #HL * HE1".
-    iDestruct (elctx_interp_persist with "HE1") as "#HE1'".
-    iDestruct (Hκκ' with "[HE HE1'] HL") as "#Hκκ'".
-    { rewrite /elctx_interp_0 big_sepL_app. auto. }
-    iMod (HE2 with "HE HL HE1") as (qE2) "[HE2 Hclose']". done.
-    iExists qE2. iFrame "∗#". iIntros "!> [_ HE2']". by iApply "Hclose'".
-  Qed.
-End elctx_incl.
-
-Arguments elctx_incl {_ _ _} _%EL _%LL _%EL _%EL.
+Lemma elctx_sat_submseteq `{invG Σ, lftG Σ} E E' L :
+  E' ⊆+ E → elctx_sat E L E'.
+Proof. iIntros (HE' ?) "_ !# H". by iApply big_sepL_submseteq. Qed.
 
 Hint Constructors Forall Forall2 elem_of_list : lrust_typing.
 Hint Resolve of_val_unlock : lrust_typing.
@@ -471,18 +316,16 @@ Hint Resolve
      lctx_lft_incl_relf lctx_lft_incl_static lctx_lft_incl_local'
      lctx_lft_incl_external'
      lctx_lft_alive_static lctx_lft_alive_local lctx_lft_alive_external
-     elctx_sat_nil elctx_sat_alive elctx_sat_lft_incl elctx_sat_app elctx_sat_refl
-     elctx_incl_refl elctx_incl_nil elctx_incl_lft_alive elctx_incl_lft_incl
+     elctx_sat_nil elctx_sat_lft_incl elctx_sat_app elctx_sat_refl
+     submseteq_cons submseteq_inserts_l submseteq_inserts_r
   : lrust_typing.
-Hint Resolve elctx_sat_cons_weaken elctx_sat_app_weaken_l elctx_sat_app_weaken_r
-             lctx_lft_alive_external' | 100 : lrust_typing.
-
-Hint Opaque elctx_incl elctx_sat lctx_lft_alive lctx_lft_incl : lrust_typing.
-
-Lemma elctx_incl_lft_incl_alive `{invG Σ, lftG Σ} E L E1 E2 κ κ' :
-  lctx_lft_incl (E ++ E1) L κ κ' → elctx_incl E L E1 E2 →
-  elctx_incl E L (κ :: E1) (κ' :: E2).
-Proof.
-  move=> ? /elctx_incl_lft_incl -> //.
-  apply (elctx_incl_app_proper _ _ [_; _] [_]); solve_typing.
-Qed.
+
+Hint Resolve elctx_sat_submseteq | 100 : lrust_typing.
+
+(* FIXME : I would prefer using a [Hint Resolve <-] for this, but
+   unfortunately, this is not preserved across modules. See:
+     https://coq.inria.fr/bugs/show_bug.cgi?id=5189
+   This should be fixed in the next version of Coq. *)
+Hint Extern 1 (_ ∈ _ ++ _) => apply <- @elem_of_app : lrust_typing.
+
+Hint Opaque elctx_sat lctx_lft_alive lctx_lft_incl : lrust_typing.
diff --git a/theories/typing/lib/cell.v b/theories/typing/lib/cell.v
index 469d46d960d848ea844212df40adb83bbbfd2eea..1de6c057ffc02905d1993707608fe1f4b843c84f 100644
--- a/theories/typing/lib/cell.v
+++ b/theories/typing/lib/cell.v
@@ -21,6 +21,9 @@ Section cell.
     iIntros (ty ?? tid l) "#H⊑ H". by iApply (na_bor_shorten with "H⊑").
   Qed.
 
+  Global Instance cell_wf ty `{!TyWf ty} : TyWf (cell ty) :=
+    { ty_lfts := ty.(ty_lfts); ty_wf_E := ty.(ty_wf_E) }.
+
   Global Instance cell_type_ne : TypeNonExpansive cell.
   Proof. solve_type_proper. Qed.
 
@@ -32,8 +35,9 @@ Section cell.
 
   Global Instance cell_mono E L : Proper (eqtype E L ==> subtype E L) cell.
   Proof.
-    iIntros (?? EQ%eqtype_unfold) "#HE #HL".
-    iDestruct (EQ with "HE HL") as "(% & #Hown & #Hshr)".
+    move=>?? /eqtype_unfold EQ. iIntros (?) "HL".
+    iDestruct (EQ with "HL") as "#EQ". iIntros "!# #HE".
+    iDestruct ("EQ" with "HE") as "(% & #Hown & #Hshr)".
     iSplit; [done|iSplit; iIntros "!# * H"].
     - iApply ("Hown" with "H").
     - iApply na_bor_iff; last done. iSplit; iIntros "!>!# H";
@@ -84,37 +88,36 @@ Section typing.
   (* Constructing a cell. *)
   Definition cell_new : val := funrec: <> ["x"] := return: ["x"].
 
-  Lemma cell_new_type ty :
-    typed_val cell_new (fn([]; ty) → cell ty).
+  Lemma cell_new_type ty `{!TyWf ty} : typed_val cell_new (fn(∅; ty) → cell ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_jump; [solve_typing..|].
-    iIntros (???) "#LFT $ $ Hty". rewrite !tctx_interp_singleton /=. done.
+    iIntros (??) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
   Qed.
 
   (* The other direction: getting ownership out of a cell. *)
   Definition cell_into_inner : val := funrec: <> ["x"] := return: ["x"].
 
-  Lemma cell_into_inner_type ty :
-    typed_val cell_into_inner (fn([]; cell ty) → ty).
+  Lemma cell_into_inner_type ty `{!TyWf ty} :
+    typed_val cell_into_inner (fn(∅; cell ty) → ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_jump; [solve_typing..|].
-    iIntros (???) "#LFT $ $ Hty". rewrite !tctx_interp_singleton /=. done.
+    iIntros (??) "#LFT _ $ Hty". rewrite !tctx_interp_singleton /=. done.
   Qed.
 
   Definition cell_get_mut : val :=
     funrec: <> ["x"] := return: ["x"].
 
-  Lemma cell_get_mut_type ty :
-    typed_val cell_get_mut (fn(∀ α, [α]; &uniq{α} (cell ty)) → &uniq{α} ty).
+  Lemma cell_get_mut_type ty `{!TyWf ty} :
+    typed_val cell_get_mut (fn(∀ α, ∅; &uniq{α} (cell ty)) → &uniq{α} ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_jump; [solve_typing..|]. rewrite /tctx_incl /=.
-    iIntros (???) "_ $$". rewrite !tctx_interp_singleton /tctx_elt_interp /=.
+    iIntros (??) "_ _ $". rewrite !tctx_interp_singleton /tctx_elt_interp /=.
     by iIntros "$".
   Qed.
 
@@ -125,11 +128,11 @@ Section typing.
       letalloc: "r" <-{ty.(ty_size)} !"x'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma cell_get_type `(!Copy ty) :
-    typed_val (cell_get ty) (fn(∀ α, [α]; &shr{α} (cell ty)) → ty).
+  Lemma cell_get_type ty `{!TyWf ty} `(!Copy ty) :
+    typed_val (cell_get ty) (fn(∀ α, ∅; &shr{α} (cell ty)) → ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
     iApply type_letalloc_n; [solve_typing| |iIntros (r); simpl_subst].
     { apply (read_shr _ _ _ (cell ty)); solve_typing. }
@@ -145,22 +148,22 @@ Section typing.
       "c'" <-{ty.(ty_size)} !"x";;
       delete [ #1; "c"] ;; delete [ #ty.(ty_size); "x"] ;; return: ["r"].
 
-  Lemma cell_replace_type ty :
-    typed_val (cell_replace ty) (fn(∀ α, [α]; &shr{α} cell ty, ty) → ty).
+  Lemma cell_replace_type ty `{!TyWf ty} :
+    typed_val (cell_replace ty) (fn(∀ α, ∅; &shr{α} cell ty, ty) → ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>c x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>c x. simpl_subst.
     iApply type_deref; [solve_typing..|].
     iIntros (c'); simpl_subst.
     iApply type_new; [solve_typing..|]. iIntros (r); simpl_subst.
     (* Drop to Iris level. *)
-    iIntros (tid qE) "#LFT Htl HE HL HC".
+    iIntros (tid) "#LFT #HE Htl HL HC".
     rewrite 3!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
     iIntros "(Hr & Hc & #Hc' & Hx)".
-    rewrite {1}/elctx_interp big_opL_singleton /=.
     destruct c' as [[|c'|]|]; try done. destruct x as [[|x|]|]; try done.
     destruct r as [[|r|]|]; try done.
-    iMod (na_bor_acc with "LFT Hc' HE Htl") as "(Hc'↦ & Htl & Hclose)"; [solve_ndisj..|].
+    iMod (lctx_lft_alive_tok α with "HE HL") as (q') "(Htok & HL & Hclose1)"; [solve_typing..|].
+    iMod (na_bor_acc with "LFT Hc' Htok Htl") as "(Hc'↦ & Htl & Hclose2)"; [solve_ndisj..|].
     iDestruct "Hc'↦" as (vc') "[>Hc'↦ Hc'own]".
     iDestruct (ty_size_eq with "Hc'own") as "#>%".
     iDestruct "Hr" as "[Hr↦ Hr†]". iDestruct "Hr↦" as (vr) "[>Hr↦ Hrown]".
@@ -174,16 +177,16 @@ Section typing.
     rewrite Nat2Z.id. iDestruct (ty_size_eq with "Hxown") as "#%".
     wp_apply (wp_memcpy with "[$Hc'↦ $Hx↦]"); try by f_equal.
     iIntros "[Hc'↦ Hx↦]". wp_seq.
-    iMod ("Hclose" with "[Hc'↦ Hxown] Htl") as "[HE Htl]"; first by auto with iFrame.
+    iMod ("Hclose2" with "[Hc'↦ Hxown] Htl") as "[Htok Htl]"; first by auto with iFrame.
+    iMod ("Hclose1" with "Htok HL") as "HL".
     (* Now go back to typing level. *)
     iApply (type_type _ _ _
            [c ◁ box (&shr{α} cell ty); #x ◁ box (uninit ty.(ty_size)); #r ◁ box ty]%TC
-    with "[] LFT Htl [HE] HL HC"); last first.
+    with "[] LFT HE Htl HL HC"); last first.
     { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
       iFrame "Hc". rewrite !tctx_hasty_val' //. iSplitL "Hx↦ Hx†".
       - iFrame. iExists _. iFrame. iNext. iApply uninit_own. done.
       - iFrame. iExists _. iFrame. }
-    { rewrite /elctx_interp big_opL_singleton. done. }
     iApply type_delete; [solve_typing..|].
     iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
diff --git a/theories/typing/lib/fake_shared_box.v b/theories/typing/lib/fake_shared_box.v
index 819cbba274ef456518620cfbcc91bf0559ee018b..f25234813054d7fa35f008316fd2bb07b67de976 100644
--- a/theories/typing/lib/fake_shared_box.v
+++ b/theories/typing/lib/fake_shared_box.v
@@ -9,17 +9,15 @@ Section fake_shared_box.
   Definition fake_shared_box : val :=
     funrec: <> ["x"] := Skip ;; return: ["x"].
 
-  Lemma cell_replace_type ty :
+  Lemma fake_shared_box_type ty `{!TyWf ty} :
     typed_val fake_shared_box
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL
-                              [# &shr{α}(&shr{β} ty)] (&shr{α}box ty))).
+      (fn (fun '(α, β) => FP_wf ∅ [# &shr{α}(&shr{β} ty)] (&shr{α}box ty))).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT".
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
+    iIntros (tid) "#LFT #HE Hna HL Hk HT".
     rewrite tctx_interp_singleton tctx_hasty_val.
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "(Hα & Hβ & #Hαβ)".
+    iDestruct (lctx_lft_incl_incl α β with "HL HE") as "#Hαβ"; [solve_typing..|].
     iAssert (▷ ty_own (own_ptr 1 (&shr{α} box ty)) tid [x])%I with "[HT]" as "HT".
     { destruct x as [[|l|]|]=>//=. iDestruct "HT" as "[H $]".
       iNext. iDestruct "H" as ([|[[]|][]]) "[H↦ H]"; try done.
@@ -30,10 +28,10 @@ Section fake_shared_box.
       iDestruct "H" as "#H". iIntros "!# * % $". iApply step_fupd_intro. set_solver.
       by iApply ty_shr_mono. }
     wp_seq.
-    iApply (type_type _ _ _ [ x ◁ box (&shr{α}box ty) ]%TC
-            with "[] LFT Hna [Hα Hβ Hαβ] HL Hk [HT]"); last first.
+    iApply (type_type [] _ _ [ x ◁ box (&shr{α}box ty) ]%TC
+            with "[] LFT [] Hna HL Hk [HT]"); last first.
     { by rewrite tctx_interp_singleton tctx_hasty_val. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
+    { by rewrite /elctx_interp. }
     iApply type_jump; simpl; solve_typing.
   Qed.
 End fake_shared_box.
diff --git a/theories/typing/lib/option.v b/theories/typing/lib/option.v
index 9b7df4c5e8911a3cbb02b72bf73db6322d65c942..9b6594b0a530baed23f97f47f3246f46996156e2 100644
--- a/theories/typing/lib/option.v
+++ b/theories/typing/lib/option.v
@@ -22,21 +22,21 @@ Section option.
     cont: "k" [] :=
       delete [ #1; "o"];; return: ["r"].
 
-  Lemma option_as_mut_type Ï„ :
+  Lemma option_as_mut_type Ï„ `{!TyWf Ï„} :
     typed_val
-      option_as_mut (fn(∀ α, [α]; &uniq{α} (option τ)) → option (&uniq{α}τ)).
+      option_as_mut (fn(∀ α, ∅; &uniq{α} (option τ)) → option (&uniq{α}τ)).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ret p).
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ϝ ret p).
       inv_vec p=>o. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (o'). simpl_subst.
     iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
-    iApply (type_cont [] [] (λ _, [o ◁ _; r ◁ _])%TC) ; [solve_typing..| |].
+    iApply (type_cont [] [ϝ ⊑ []]%LL (λ _, [o ◁ _; r ◁ _])%TC) ; [solve_typing..| |].
     - iIntros (k). simpl_subst.
       iApply type_case_uniq; [solve_typing..|].
         constructor; last constructor; last constructor; left.
       + iApply (type_sum_unit (option $ &uniq{α}τ)%T); [solve_typing..|].
         iApply type_jump; solve_typing.
-      + iApply (type_sum_assign (option $ &uniq{α}τ)%T); [solve_typing..|].
+      + iApply (type_sum_assign (option $ &uniq{α}τ)%T); [try solve_typing..|].
         iApply type_jump; solve_typing.
     - iIntros "/= !#". iIntros (k args). inv_vec args. simpl_subst.
       iApply type_delete; [solve_typing..|].
@@ -53,11 +53,11 @@ Section option.
         delete [ #(S Ï„.(ty_size)); "o"];; delete [ #Ï„.(ty_size); "def"];;
         return: ["r"]].
 
-  Lemma option_unwrap_or_type Ï„ :
-    typed_val (option_unwrap_or τ) (fn([]; option τ, τ) → τ).
+  Lemma option_unwrap_or_type Ï„ `{!TyWf Ï„} :
+    typed_val (option_unwrap_or τ) (fn(∅; option τ, τ) → τ).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros ([] ret p).
-      inv_vec p=>o def. simpl_subst.
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
+      iIntros ([] ϝ ret p). inv_vec p=>o def. simpl_subst.
     iApply type_case_own; [solve_typing|]. constructor; last constructor; last constructor.
     + right. iApply type_delete; [solve_typing..|].
       iApply type_jump; solve_typing.
diff --git a/theories/typing/lib/rc.v b/theories/typing/lib/rc.v
index b546de5bc4ee3e18ec16640b48070af23604fb76..632ef0a308dd7c6296fa2c011706de5e9fb8097f 100644
--- a/theories/typing/lib/rc.v
+++ b/theories/typing/lib/rc.v
@@ -36,13 +36,21 @@ Section rc.
       | None => True
       end)%I.
 
+  Global Instance rc_type_ne n :
+    Proper (type_dist2 n ==> dist n) (rc_inv tid ν γ l).
+  Proof.
+    solve_proper_core
+      ltac:(fun _ => exact: type_dist2_S || f_type_equiv || f_contractive || f_equiv).
+  Qed.
+
   Program Definition rc (ty : type) :=
     {| ty_size := 1;
        ty_own tid vl :=
          match vl return _ with
          | [ #(LitLoc l) ] =>
-           (l ↦ #1 ∗ †{1%Qp}l…(S ty.(ty_size)) ∗ ▷ shift_loc l 1 ↦∗: ty.(ty_own) tid) ∨
-           (∃ γ ν q, na_inv tid rc_invN (rc_inv tid ν γ l ty) ∗
+           (l ↦ #1 ∗ ▷ †{1%Qp}l…(S ty.(ty_size)) ∗ ▷ shift_loc l 1 ↦∗: ty.(ty_own) tid) ∨
+           (∃ γ ν q ty', ▷ type_incl ty' ty ∗
+                     na_inv tid rc_invN (rc_inv tid ν γ l ty') ∗
                      own γ (◯ Some (q, 1%positive)) ∗
                      ty.(ty_shr) ν tid (shift_loc l 1) ∗
                      q.[ν] ∗ □ (1.[ν] ={↑lftN,∅}▷=∗ [†ν]))
@@ -50,8 +58,9 @@ Section rc.
        ty_shr κ tid l :=
          ∃ (l' : loc), &frac{κ} (λ q, l ↦∗{q} [ #l']) ∗
            □ ∀ F q, ⌜↑shrN ∪ lftE ⊆ F⌝ -∗ q.[κ]
-             ={F, F∖↑shrN}▷=∗ q.[κ] ∗ ∃ γ ν q', κ ⊑ ν ∗
-                na_inv tid rc_invN (rc_inv tid ν γ l' ty) ∗
+             ={F, F∖↑shrN}▷=∗ q.[κ] ∗ ∃ γ ν q' ty', κ ⊑ ν ∗
+                ▷ type_incl ty' ty ∗
+                na_inv tid rc_invN (rc_inv tid ν γ l' ty') ∗
                 &na{κ, tid, rc_shrN}(own γ (◯ Some (q', 1%positive))) ∗
                 ty.(ty_shr) ν tid (shift_loc l' 1)
     |}%I.
@@ -67,7 +76,7 @@ Section rc.
       try by iMod (bor_persistent_tok with "LFT Hb Htok") as "[>[] _]".
     iFrame "Htok". iExists _. iFrame "#". rewrite bor_unfold_idx.
     iDestruct "Hb" as (i) "(#Hpb&Hpbown)".
-    set (C := (∃ _ _ _, _ ∗ _ ∗ &na{_,_,_} _ ∗ _)%I).
+    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_open with "Hinv") as "[INV Hclose1]"; first set_solver.
@@ -80,23 +89,28 @@ Section rc.
       iModIntro. iNext. iAssert X with "[>HP]" as "HX".
       { iDestruct "HP" as "[[Hl' [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. *)
+        (* 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 ((1/2)%Qp, 1%positive) ⋅ ◯ Some ((1/2)%Qp, 1%positive))) as (γ) "[Ha Hf]".
+        iMod (own_alloc (● Some ((1/2)%Qp, 1%positive) ⋅
+                         ◯ Some ((1/2)%Qp, 1%positive))) as (γ) "[Ha Hf]".
         { apply auth_valid_discrete_2. done. }
         iMod (na_inv_alloc tid _ _ (rc_inv tid ν γ l' ty) with "[Ha Hν2 Hl' H† HPend]").
         { rewrite /rc_inv. iExists (Some (_, _)). iFrame. iExists _. iFrame "#∗".
           rewrite Qp_div_2; auto. }
         iMod (ty_share with "LFT HP Hν1") as "[??]"; first solve_ndisj.
-        iExists _, _, _. iFrame. done. }
-      iDestruct "HX" as (γ ν q') "(#Hinv & Hrc & #Hshr & Htok & #Hν†)".
+        iExists _, _, _, _. iFrame. iModIntro. iSplit; last done.
+        by iApply type_incl_refl. }
+      iDestruct "HX" as (γ ν q' ty') "(#Hincl & #Hinv & Hrc & #Hshr & Htok & #Hν†)".
       iMod ("Hclose2" with "[] [Hrc Htok]") as "[HX Htok]".
       { iNext. iIntros "HX". iModIntro. iNext.
-        iRight. iExists γ, ν, q'. iExact "HX". }
-      { iNext. by iFrame "#∗". }
+        iRight. iExists γ, ν, q', _. iExact "HX". }
+      { iNext.  by iFrame "#∗". }
       iAssert C with "[>HX]" as "#HC".
-      { iExists _, _, _. iFrame "Hinv".
+      { iExists _, _, _, _. iFrame "Hinv".
+        iMod (bor_sep with "LFT HX") as "[_ HX]"; first solve_ndisj.
         iMod (bor_sep with "LFT HX") as "[_ HX]"; first solve_ndisj.
         iMod (bor_sep with "LFT HX") as "[Hrc HX]"; first solve_ndisj.
         iMod (bor_sep with "LFT HX") as "[_ HX]"; first solve_ndisj.
@@ -117,11 +131,61 @@ Section rc.
     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.
+    iMod ("Hclose" with "Htok") as "$". iDestruct "HX" as (? ν ??) "(? & ? & ? & ? & ?)".
+    iExists _, _, _, _. iModIntro. iFrame. iSplit.
     - by iApply lft_incl_trans.
     - by iApply na_bor_shorten.
   Qed.
+
+  Global Instance rc_wf ty `{!TyWf ty} : TyWf (rc ty) :=
+    { ty_lfts := ty.(ty_lfts); ty_wf_E := ty.(ty_wf_E) }.
+
+  Global Instance rc_type_contractive : TypeContractive rc.
+  Proof.
+    constructor;
+      solve_proper_core ltac:(fun _ => exact: type_dist2_S || exact: type_dist2_dist ||
+                                       (eapply refcell_inv_type_ne; try reflexivity) ||
+                                       f_type_equiv || f_contractive || f_equiv).
+  Qed.
+
+  Global Instance rc_ne : NonExpansive rc.
+  Proof. apply type_contractive_ne, _. Qed.
+
+  Lemma rc_subtype ty1 ty2 :
+    type_incl ty1 ty2 -∗ type_incl (rc ty1) (rc ty2).
+  Proof.
+    iIntros "#Hincl". iPoseProof "Hincl" as "Hincl2".
+    iDestruct "Hincl2" as "(#Hsz & #Hoincl & #Hsincl)".
+    (* FIXME: Would be nice to have an easier way to duplicate
+       destructed things.  Maybe iPoseProof with a destruct pattern? *)
+    iSplit; first done. iSplit; iAlways.
+    - iIntros "* Hvl". destruct vl as [|[[|vl|]|] [|]]; try done.
+      iDestruct "Hvl" as "[(Hl & H† & Hc) | Hvl]".
+      { iLeft. iFrame "Hl". iNext. iDestruct "Hsz" as %->.
+        iFrame. iApply (heap_mapsto_pred_wand with "Hc").
+        iApply "Hoincl". }
+      iDestruct "Hvl" as (γ ν q ty') "(#Hincl' & #Hinc & Htk & #Hsr & Hν)".
+      iRight. iExists _, _, _, _. iFrame "#∗". iSplit.
+      + iNext. iApply type_incl_trans; done.
+      + iApply "Hsincl". done.
+    - iIntros "* #Hshr". iDestruct "Hshr" as (l') "[Hfrac Hshr]". iExists l'.
+      iFrame "Hfrac". iIntros "!# * % Htok". iMod ("Hshr" with "[% //] Htok") as "{Hshr} H".
+      iModIntro. iNext. iMod "H" as "[Htok H]". iModIntro. iFrame "Htok".
+      iDestruct "H" as (γ ν q' ty') "(Hlft & #Hincl' & Hinv & Hna & Hshr)". iExists _, _, _, _.
+      iFrame. iSplit.
+      + iNext. iApply type_incl_trans; done.
+      + iApply "Hsincl". done.
+  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. iApply "Hsub"; done.
+  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.
@@ -134,7 +198,7 @@ Section code.
     {{{ c, RET #(Zpos c); l ↦ #(Zpos c) ∗
         ((⌜c = 1%positive⌝ ∗ †{1%Qp}l…(S ty.(ty_size)) ∗ na_own tid F ∗ ▷ shift_loc l 1 ↦∗: ty.(ty_own) tid) ∨
          (⌜(1 < c)%positive⌝ ∗ na_own tid (F ∖ ↑rc_invN) ∗
-          ∃ γ ν q q', na_inv tid rc_invN (rc_inv tid ν γ l ty) ∗
+          ∃ γ ν q q' ty', ▷ type_incl ty' ty ∗ na_inv tid rc_invN (rc_inv tid ν γ l ty') ∗
             own γ (◯ Some (q, 1%positive)) ∗ own γ (● Some ((q + q')%Qp, c)) ∗
             q.[ν] ∗ □ ((1).[ν] ={↑lftN,∅}▷=∗ [†ν]) ∗
             ty.(ty_shr) ν tid (shift_loc l 1) ∗
@@ -145,10 +209,12 @@ Section code.
   Proof.
     iIntros (? Φ) "[Hna [(Hl & H† & Hown)|Hown]] HΦ".
     - wp_read. iApply "HΦ". auto with iFrame.
-    - iDestruct "Hown" as (γ ν q) "(#Hinv & Htok & #Hshr & Hν1 & #Hνend)".
+    - iDestruct "Hown" as (γ ν q ty') "(#Hincl & #Hinv & Htok & #Hshr & Hν1 & #Hνend)".
       iMod (na_inv_open with "Hinv Hna") as "(Hproto & Hna & Hclose)"; [solve_ndisj..|].
       iDestruct "Hproto" as (st) "[>Hst Hproto]".
-      iDestruct (own_valid_2 with "Hst Htok") as %[[Hval|(? & (qa, c) & [=<-] & -> & Hst)]%option_included _]%auth_valid_discrete_2; first done. (* Oh my, what a pattern... *)
+      iDestruct (own_valid_2 with "Hst Htok")
+        as %[[Hval|(? &(qa,c) & [=<-] & -> & Hst)]%option_included _]%auth_valid_discrete_2;
+        first done. (* Oh my, what a pattern... *)
       iDestruct "Hproto" as (q') "(>Hl & H† & >% & >Hν2 & Hν† & _)". iApply wp_fupd.
       destruct (decide (c ≤ 1)%positive) as [Hle|Hnle].
       + (* Tear down the protocol. *)
@@ -165,12 +231,14 @@ Section code.
         { rewrite -H0. iFrame. }
         iApply wp_mask_mono; last iApply (wp_step_fupd with "Hνend"); [done..|solve_ndisj|].
         wp_read. iIntros "#Hν". iMod ("Hν†" with "Hν") as "Hown". iModIntro.
-        iApply "HΦ". iFrame. iLeft. auto with iFrame.
+        iApply "HΦ". iFrame. iLeft. iModIntro. iSplit; first done.
+        iDestruct "Hincl" as "(#Hsz & #Hincl & _)". iDestruct "Hsz" as %->. iFrame.
+        iApply (heap_mapsto_pred_wand with "Hown"). iApply "Hincl".
       + destruct Hst as [[??%leibniz_equiv]|[[q'' Hqa%leibniz_equiv] ?%pos_included]%prod_included].
         { exfalso. simpl in *. subst c. apply Hnle. done. }
         simpl in *. subst qa. wp_read. iApply "HΦ". iFrame. iRight. iModIntro. iSplit.
         { iPureIntro. apply Pos.lt_nle. done. }
-        iFrame "Hna". iExists _, _, q, q''. iFrame "#∗%".
+        iFrame "Hna". iExists _, _, q, q'', _. iFrame "#∗%".
         iIntros (c' qx) "(Hst & Hl & Hq'' & Hna)". iApply "Hclose". iFrame "Hna".
         iExists _. iFrame "Hst". iDestruct "Hq''" as "[%|Hq'']".
         * iExists _. subst qx. iFrame "∗%". by iIntros "!> !#".
@@ -188,14 +256,14 @@ Section code.
       "rc" +â‚— #0 <- "rcbox";;
       delete [ #ty.(ty_size); "x"];; return: ["rc"].
 
-  Lemma rc_new_type ty :
-    typed_val (rc_new ty) (fn([]; ty) → rc ty).
+  Lemma rc_new_type ty `{!TyWf ty} :
+    typed_val (rc_new ty) (fn(∅; ty) → rc ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_new; [solve_typing..|]; iIntros (rcbox); simpl_subst.
     iApply type_new; [solve_typing..|]; iIntros (rc); simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL Hk [Hrc [Hrcbox [Hx _]]]".
+    iIntros (tid) "#LFT #HE Hna HL Hk [Hrc [Hrcbox [Hx _]]]".
     rewrite (Nat2Z.id (S ty.(ty_size))) !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.
@@ -207,7 +275,7 @@ Section code.
     wp_apply (wp_memcpy with "[$Hrcbox↦1 $Hx↦]"); [by auto using vec_to_list_length..|].
     iIntros "[Hrcbox↦1 Hx↦]". wp_seq. wp_op. rewrite shift_loc_0. wp_write.
     iApply (type_type _ _ _ [ #lx ◁ box (uninit (ty_size ty)); #lrc ◁ box (rc ty)]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        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. }
@@ -227,32 +295,35 @@ Section code.
       "rc2" <- "rc''";;
       delete [ #1; "rc" ];; return: ["rc2"].
 
-  Lemma rc_clone_type ty :
-    typed_val rc_clone (fn(∀ α, [α]; &shr{α} rc ty) → rc ty).
+  Lemma rc_clone_type ty `{!TyWf ty} :
+    typed_val rc_clone (fn(∀ α, ∅; &shr{α} rc ty) → rc ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_new; [solve_typing..|]; iIntros (rc2); simpl_subst.
     rewrite (Nat2Z.id 1). (* Having to do this is rather annoying... *) 
     iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL Hk [Hx [Hrc' [Hrc2 _]]]".
+    iIntros (tid) "#LFT #HE Hna HL Hk [Hx [Hrc' [Hrc2 _]]]".
     rewrite !tctx_hasty_val [[x]]lock.
     destruct rc' as [[|lrc|]|]; try done. iDestruct "Hrc'" as (l') "[#Hlrc #Hshr]".
     iDestruct (ownptr_uninit_own with "Hrc2") as (lrc2 vlrc2) "(% & Hrc2 & Hrc2†)".
     subst rc2. inv_vec vlrc2=>rc2. rewrite heap_mapsto_vec_singleton.
     (* All right, we are done preparing our context. Let's get going. *)
-    iDestruct "HE" as "[[Hα1 Hα2] _]". wp_bind (!_)%E.
+    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.
     rewrite heap_mapsto_vec_singleton. iApply wp_fupd. wp_read.
     iMod ("Hclose" with "[$Hlrc↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
-    iDestruct "Hproto" as (γ ν q'') "(#Hincl & #Hinv & #Hrctokb & #Hshr)". iModIntro.
+    iDestruct "Hproto" as (γ ν q'' ty') "(#Hincl & #Hty & #Hinv & #Hrctokb & #Hshr)". iModIntro.
     wp_let. wp_op. rewrite shift_loc_0.
     (* Finally, finally... opening the thread-local Rc protocol. *)
-    iMod (na_inv_open with "Hinv Hna") as "(Hrcproto & Hna & Hclose1)"; [solve_ndisj..|].
-    iMod (na_bor_acc with "LFT Hrctokb Hα1 Hna") as "(>Hrctok & Hna & Hclose2)"; [solve_ndisj..|].
+    iMod (na_inv_open 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) "[>Hrc● Hrcst]".
-    iDestruct (own_valid_2 with "Hrc● Hrctok") as %[[Hval|(_ & (qa, c) & _ & -> & _)]%option_included _]%auth_valid_discrete_2; first done. (* Oh my, what a pattern... *)
+    iDestruct (own_valid_2 with "Hrc● Hrctok") as
+        %[[Hval|(_ & (qa, c) & _ & -> & _)]%option_included _]%auth_valid_discrete_2;
+      first done. (* Oh my, what a pattern... *)
     iDestruct "Hrcst" as (qb) "(>Hl' & >Hl'† & >% & Hνtok & Hν† & #Hνend)".
     wp_read. wp_let. wp_op. rewrite shift_loc_0. wp_op. wp_write. wp_write.
     (* And closing it again. *)
@@ -263,17 +334,17 @@ Section code.
       apply Qcplus_le_mono_l. rewrite -{1}(Qcplus_0_l (_ / _)%Qp).
       apply Qcplus_le_mono_r, Qp_ge_0. }
     rewrite right_id -Some_op pair_op. iDestruct "Hνtok" as "[Hνtok1 Hνtok2]".
-    iMod ("Hclose2" with "[$Hrctok] Hna") as "[Hα1 Hna]".
-    iMod ("Hclose1" with "[Hrc● Hl' Hl'† Hνtok2 Hν† $Hna]") as "Hna".
+    iMod ("Hclose3" with "[$Hrctok] Hna") as "[Hα1 Hna]".
+    iMod ("Hclose2" with "[Hrc● Hl' Hl'† Hνtok2 Hν† $Hna]") as "Hna".
     { iExists _. iFrame "Hrc●". iExists _. rewrite Z.add_comm. iFrame "∗ Hνend".
-      rewrite [_ â‹… _]comm frac_op' -assoc Qp_div_2. auto. }
+      rewrite [_ â‹… _]comm frac_op' -[(_ + _)%Qp]assoc. rewrite 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); #lrc2 ◁ box (rc ty)]%TC
-        with "[] LFT Hna [Hα1 Hα2] HL Hk [-]"); last first.
+        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 "Hrc2†". iExists [_]. rewrite heap_mapsto_vec_singleton.
-      iFrame "Hrc2". iNext. iRight. iExists _, ν, _. iFrame "#∗". }
-    { by iFrame. }
+      iFrame "Hrc2". iNext. iRight. iExists _, ν, _, _. iFrame "#∗". }
     iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
   Qed.
@@ -285,34 +356,35 @@ Section code.
       "x" <- (!"rc'" +â‚— #1);;
       delete [ #1; "rc" ];; return: ["x"].
 
-  Lemma rc_deref_type ty :
-    typed_val rc_deref (fn(∀ α, [α]; &shr{α} rc ty) → &shr{α} ty).
+  Lemma rc_deref_type ty `{!TyWf ty} :
+    typed_val rc_deref (fn(∀ α, ∅; &shr{α} rc ty) → &shr{α} ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>rcx. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
     iApply type_new; [solve_typing..|]; iIntros (x); simpl_subst. rewrite (Nat2Z.id 1).
     iApply type_deref; [solve_typing..|]; iIntros (rc'); simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL Hk [Hrcx [Hrc' [Hx _]]]".
+    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. *)
-    iDestruct "HE" as "[[Hα1 Hα2] _]". wp_bind (!_)%E.
+    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'↦ Hclose]"; first solve_ndisj.
+    iMod (frac_bor_acc with "LFT Hrc' Hα2") as (q') "[Hrc'↦ Hclose2]"; first solve_ndisj.
     rewrite heap_mapsto_vec_singleton. iApply wp_fupd. wp_read.
-    iMod ("Hclose" with "[$Hrc'↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
-    iDestruct "Hproto" as (γ ν q'') "(#Hincl & #Hinv & #Hrctokb & #Hshr)". iModIntro.
+    iMod ("Hclose2" with "[$Hrc'↦]") as "Hα2". iIntros "!> [Hα1 Hproto]".
+    iDestruct "Hproto" as (γ ν q'' ty') "(#? & #? & #Hinv & #Hrctokb & #Hshr)". iModIntro.
     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)]%TC
-        with "[] LFT Hna [Hα1 Hα2] HL Hk [-]"); last first.
+        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". iNext. iApply ty_shr_mono; done. }
-    { by iFrame. }
     iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
   Qed.
@@ -334,20 +406,23 @@ Section code.
     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 ]).
+  Lemma rc_try_unwrap_type ty `{!TyWf ty} :
+    typed_val (rc_try_unwrap ty) (fn(∅; rc ty) → Σ[ ty; rc ty ]).
   Proof.
     (* TODO: There is a *lot* of duplication between this proof and the one for drop. *)
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>rcx. simpl_subst.
-    iApply type_new; [solve_typing..|]; iIntros (x); simpl_subst. rewrite (Nat2Z.id (Σ[ ty; rc ty ]).(ty_size)).
-    iApply (type_cont [] [] (λ _, [rcx ◁ box (uninit 1); x ◁ box (Σ[ ty; rc ty ])])%TC) ; [solve_typing..| |]; last first.
+    iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
+    iApply type_new; [solve_typing..|]; iIntros (x); simpl_subst.
+    rewrite (Nat2Z.id (Σ[ ty; rc ty ]).(ty_size)).
+    iApply (type_cont [] [ϝ ⊑ []]%LL
+                      (λ _, [rcx ◁ box (uninit 1); x ◁ box (Σ[ ty; rc ty ])])%TC) ;
+      [solve_typing..| |]; last first.
     { simpl. iAlways. 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 qE) "#LFT Hna HE HL Hk [Hrcx [Hrc' [Hx _]]]".
+    iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hx _]]]".
     rewrite !tctx_hasty_val [[rcx]]lock [[x]]lock.
     destruct rc' as [[|rc'|]|]; try done.
     rewrite [[ #rc' ]]lock. wp_op. rewrite shift_loc_0 -(lock [ #rc' ]).
@@ -359,27 +434,28 @@ Section code.
                                 rcx ◁ box (uninit 1);
                                 #rc' +ₗ #1 ◁ own_ptr (S ty.(ty_size)) ty;
                                 #rc' +ₗ #0 ◁ own_ptr (S ty.(ty_size)) (uninit 1%nat)]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        with "[] LFT HE Hna HL Hk [-]"); last first.
       { unlock. rewrite 3!tctx_interp_cons tctx_interp_singleton.
         rewrite 2!tctx_hasty_val tctx_hasty_val' // tctx_hasty_val' //.
         rewrite -freeable_sz_full_S -(freeable_sz_split _ 1%nat).
         iDestruct "Hrc†" as "[Hrc†1 Hrc†2]". iFrame.
-        rewrite shift_loc_0. iFrame. iExists [_]. rewrite uninit_own heap_mapsto_vec_singleton.
-        auto with iFrame. }
+        rewrite shift_loc_0. iFrame. iExists [_].
+        rewrite uninit_own heap_mapsto_vec_singleton. auto with iFrame. }
       iApply (type_sum_memcpy Σ[ ty; rc ty ]); [solve_typing..|].
       iApply (type_delete (Π[uninit _; uninit _])); [solve_typing..|].
       iApply type_jump; solve_typing.
-    - iDestruct "Hproto" as (γ ν q q') "(#Hinv & Hrctok & Hrc● & Hν & #Hν† & #Hshr & Hclose)".
+    - iDestruct "Hproto" as (γ ν q q' ty')
+          "(#Hincl & #Hinv & Hrctok & Hrc● & Hν & #Hν† & #Hshr & Hclose)".
       wp_op; intros Hc.
       { exfalso. move:Hc. move/Zpos_eq_iff. intros ->. exact: Pos.lt_irrefl. }
       wp_if. iMod ("Hclose" with "[> $Hrc● $Hrc $Hna]") as "Hna"; first by eauto.
       iApply (type_type _ _ _ [ x ◁ own_ptr (ty_size Σ[ ty; rc ty ]) (uninit _);
                                 rcx ◁ box (uninit 1);
                                 #rc' ◁ rc ty ]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        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.
-        iRight. iExists _, _, _. iFrame "∗#". }
+        iRight. iExists _, _, _, _. iFrame "∗#". }
       iApply (type_sum_assign Σ[ ty; rc ty ]); [solve_typing..|].
       iApply type_jump; solve_typing.
   Qed.
@@ -402,19 +478,22 @@ Section code.
     cont: "k" [] :=
       delete [ #1; "rc"];; return: ["x"].
 
-  Lemma rc_drop_type ty :
-    typed_val (rc_drop ty) (fn([]; rc ty) → option ty).
+  Lemma rc_drop_type ty `{!TyWf ty} :
+    typed_val (rc_drop ty) (fn(∅; rc ty) → option ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>rcx. simpl_subst.
-    iApply type_new; [solve_typing..|]; iIntros (x); simpl_subst. rewrite (Nat2Z.id (option ty).(ty_size)).
-    iApply (type_cont [] [] (λ _, [rcx ◁ box (uninit 1); x ◁ box (option ty)])%TC) ; [solve_typing..| |]; last first.
+    iIntros (_ ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
+    iApply type_new; [solve_typing..|]; iIntros (x); simpl_subst.
+    rewrite (Nat2Z.id (option ty).(ty_size)).
+    iApply (type_cont [] [ϝ ⊑ []]%LL
+                      (λ _, [rcx ◁ box (uninit 1); x ◁ box (option ty)])%TC);
+      [solve_typing..| |]; last first.
     { simpl. iAlways. 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 qE) "#LFT Hna HE HL Hk [Hrcx [Hrc' [Hx _]]]".
+    iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hx _]]]".
     rewrite !tctx_hasty_val [[rcx]]lock [[x]]lock.
     destruct rc' as [[|rc'|]|]; try done. rewrite [[ #rc' ]]lock.
     wp_op. rewrite shift_loc_0. rewrite -(lock [ #rc' ]).
@@ -427,17 +506,18 @@ Section code.
                                 rcx ◁ box (uninit 1);
                                 #rc' +ₗ #1 ◁ own_ptr (S ty.(ty_size)) ty;
                                 #rc' +ₗ #0 ◁ own_ptr (S ty.(ty_size)) (uninit 1%nat)]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        with "[] LFT HE Hna HL Hk [-]"); last first.
       { unlock. rewrite 3!tctx_interp_cons tctx_interp_singleton.
         rewrite 2!tctx_hasty_val tctx_hasty_val' // tctx_hasty_val' //.
         rewrite -freeable_sz_full_S -(freeable_sz_split _ 1%nat).
         iDestruct "Hrc†" as "[Hrc†1 Hrc†2]". iFrame.
-        rewrite shift_loc_0. iFrame. iExists [_]. rewrite uninit_own heap_mapsto_vec_singleton.
-        auto with iFrame. }
+        rewrite shift_loc_0. iFrame. iExists [_].
+        rewrite uninit_own heap_mapsto_vec_singleton. auto with iFrame. }
       iApply (type_sum_memcpy (option _)); [solve_typing..|].
       iApply (type_delete (Π[uninit _; uninit _])); [solve_typing..|].
       iApply type_jump; solve_typing.
-    - iDestruct "Hproto" as (γ ν q q') "(#Hinv & Hrctok & Hrc● & Hν & _ & _ & Hclose)".
+    - iDestruct "Hproto" as (γ ν q q' ty')
+                              "(#Hincl & #Hinv & Hrctok & Hrc● & Hν & _ & _ & Hclose)".
       wp_write. wp_op; intros Hc.
       { exfalso. move:Hc. move/Zpos_eq_iff. intros ->. exact: Pos.lt_irrefl. }
       wp_if. iMod ("Hclose" $! (c-1)%positive q' with "[> Hrc● Hrctok Hrc Hν $Hna]") as "Hna".
@@ -450,7 +530,7 @@ Section code.
         auto with iFrame. }
       iApply (type_type _ _ _ [ x ◁ own_ptr (ty_size (option ty)) (uninit _);
                                 rcx ◁ box (uninit 1)]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        with "[] LFT HE Hna HL Hk [-]"); last first.
       { rewrite tctx_interp_cons tctx_interp_singleton.
         rewrite !tctx_hasty_val. unlock. iFrame. }
       iApply (type_sum_unit (option ty)); [solve_typing..|].
@@ -473,23 +553,25 @@ Section code.
     cont: "k" [] :=
       delete [ #1; "rc"];; return: ["x"].
 
-  Lemma rc_get_mut_type ty :
-    typed_val rc_get_mut (fn(∀ α, [α]; &uniq{α} rc ty) → option (&uniq{α} ty)).
+  Lemma rc_get_mut_type ty `{!TyWf ty} :
+    typed_val rc_get_mut (fn(∀ α, ∅; &uniq{α} rc ty) → option (&uniq{α} ty)).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>rcx. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>rcx. simpl_subst.
     iApply type_new; [solve_typing..|]; iIntros (x); simpl_subst. rewrite (Nat2Z.id 2%nat).
-    iApply (type_cont [] [] (λ r, [rcx ◁ box (uninit 1); x ◁ box (option $ &uniq{α} ty)])%TC) ; [solve_typing..| |]; last first.
+    iApply (type_cont [] [ϝ ⊑ []]%LL
+                      (λ r, [rcx ◁ box (uninit 1); x ◁ box (option $ &uniq{α} ty)])%TC);
+      [solve_typing..| |]; last first.
     { simpl. iAlways. 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 qE) "#LFT Hna HE HL Hk [Hrcx [Hrc' [Hx _]]]".
+    iIntros (tid) "#LFT #HE Hna HL Hk [Hrcx [Hrc' [Hx _]]]".
     rewrite !tctx_hasty_val  [[rcx]]lock [[x]]lock.
     destruct rc' as [[|rc'|]|]; try done.
-    iDestruct "HE" as "[Hα _]".
-    iMod (bor_acc_cons with "LFT Hrc' Hα") as "[Hrc' Hclose1]"; first solve_ndisj.
+    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.
@@ -500,32 +582,33 @@ Section code.
     iIntros (c) "(Hrc & Hc)". wp_let. wp_op; intros Hc.
     - wp_if. iDestruct "Hc" as "[(% & Hl† & Hna & Hown)|(% & _)]"; last first.
       { exfalso. move:Hc. move/Zpos_eq_iff. intros->. exact: Pos.lt_irrefl. }
-      subst c. iMod ("Hclose1" with "[Hrc Hrc' Hl†] [Hown]") as "[Hown Hα]"; [|iNext; iExact "Hown"|].
+      subst c. iMod ("Hclose2" with "[Hrc Hrc' Hl†] [Hown]") as "[Hown Hα]"; [|iNext; iExact "Hown"|].
       { iIntros "!> Hown". iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame "Hrc'".
         iLeft. by iFrame. }
+      iMod ("Hclose1" with "Hα HL") as "HL".
       iApply (type_type _ _ _ [ x ◁ box (uninit 2);
                                  #l +ₗ #1 ◁ &uniq{α} ty;
                                 rcx ◁ box (uninit 1)]%TC
-        with "[] LFT Hna [Hα] HL Hk [-]"); last first.
+        with "[] LFT HE Hna HL Hk [-]"); last first.
       { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val tctx_hasty_val' //.
         unlock. iFrame. }
-      { by iFrame. }
       iApply (type_sum_assign (option (&uniq{_} _))); [solve_typing..|].
       iApply type_jump; solve_typing.
     - wp_if. iDestruct "Hc" as "[(% & _)|(% & Hna & Hproto)]".
       { exfalso. subst c. done. }
-      iDestruct "Hproto" as (γ ν q q') "(#Hinv & Hrctok & Hrc● & Hν & #Hshr & #Hν† & Hclose2)".
-      iMod ("Hclose2" with "[$Hrc● $Hrc $Hna]") as "Hna"; first by auto.
-      iMod ("Hclose1" with "[] [Hrc' Hrctok Hν]") as "[Hrc' Hα]".
+      iDestruct "Hproto" as (γ ν q' q'' ty')
+              "(#Hincl & #Hinv & Hrctok & Hrc● & Hν & #Hshr & #Hν† & Hclose3)".
+      iMod ("Hclose3" with "[$Hrc● $Hrc $Hna]") as "Hna"; first by auto.
+      iMod ("Hclose2" with "[] [Hrc' Hrctok Hν]") as "[Hrc' Hα]".
       { iIntros "!> HX". iModIntro. iExact "HX". }
       { iExists [_]. rewrite heap_mapsto_vec_singleton. iFrame "Hrc'".
-        iRight. iExists _, _, _. iFrame "#∗". }
+        iRight. iExists _, _, _, _. iFrame "#∗". }
+      iMod ("Hclose1" with "Hα HL") as "HL".
       iApply (type_type _ _ _ [ x ◁ box (uninit 2);
                                 rcx ◁ box (uninit 1)]%TC
-        with "[] LFT Hna [Hα] HL Hk [-]"); last first.
+        with "[] LFT HE Hna HL Hk [-]"); last first.
       { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
         unlock. iFrame. }
-      { by iFrame. }
       iApply (type_sum_unit (option (&uniq{_} _))); [solve_typing..|].
       iApply type_jump; solve_typing.
   Qed.
diff --git a/theories/typing/lib/refcell/ref.v b/theories/typing/lib/refcell/ref.v
index e2f9c77105ac4456dd43381a15ebdbbca1c63282..dfa8aec05845124e3622da049d188f0d1284ee46 100644
--- a/theories/typing/lib/refcell/ref.v
+++ b/theories/typing/lib/refcell/ref.v
@@ -68,6 +68,9 @@ Section ref.
     - by iApply na_bor_shorten.
   Qed.
 
+  Global Instance ref_wf α ty `{!TyWf ty} : TyWf (ref α ty) :=
+    { ty_lfts := [α]; ty_wf_E := ty.(ty_wf_E) ++ ty.(ty_outlives_E) α }.
+
   Global Instance ref_type_contractive α : TypeContractive (ref α).
   Proof. solve_type_proper. Qed.
   Global Instance ref_ne α : NonExpansive (ref α).
@@ -76,19 +79,19 @@ Section ref.
   Global Instance ref_mono E L :
     Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) ref.
   Proof.
-    iIntros (α1 α2 Hα ty1 ty2 Hty) "#HE #HL".
-    iDestruct (Hty with "HE HL") as "(%&#Ho&#Hs)".
-    iDestruct (Hα with "HE HL") as "Hα1α2".
-    iSplit; [|iSplit; iAlways].
+    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 "(%&#Ho&#Hs)". iSplit; [|iSplit; iAlways].
     - done.
     - iIntros (tid [|[[]|][|[[]|][]]]) "H"; try done.
-      iDestruct "H" as (ν q γ β ty') "(#Hshr & #H⊑ & #Hinv & Htok & Hown)".
-      iExists ν, q, γ, β, ty'. iFrame "∗#". iSplit.
+      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.
+    - 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.
diff --git a/theories/typing/lib/refcell/ref_code.v b/theories/typing/lib/refcell/ref_code.v
index d612c25ec20c3f2249e7271c3b72c8d441ee1272..9f7c86c5f962b1c34d186195e7e9ccdd8793fcb4 100644
--- a/theories/typing/lib/refcell/ref_code.v
+++ b/theories/typing/lib/refcell/ref_code.v
@@ -54,20 +54,22 @@ Section ref_functions.
 
   (* FIXME : using λ instead of fun triggers an anomaly.
      See: https://coq.inria.fr/bugs/show_bug.cgi?id=5326 *)
-  Lemma ref_clone_type ty :
-    typed_val ref_clone
-      (fn (fun '(α, β) => FP [α; β]%EL [# &shr{α}(ref β ty)]%T (ref β ty)%T)).
+  Lemma ref_clone_type ty `{!TyWf ty} :
+    typed_val ref_clone (fn (fun '(α, β) =>
+                FP_wf ∅ [# &shr{α}(ref β ty)]%T (ref β ty)%T)).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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 (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hβδ & Hinv & H◯inv)".
-    wp_op. rewrite {1}/elctx_interp big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "[[Hα1 Hα2] Hβ]".
+    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..|].
@@ -88,20 +90,20 @@ Section ref_functions.
     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α2 Hna]".
+    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 _. rewrite Z.add_comm. iFrame. iExists _. iFrame. iSplitR.
       - rewrite /= agree_idemp. auto.
       - iExists _. iFrame.
         rewrite (comm Qp_plus) (assoc Qp_plus) Qp_div_2 (comm Qp_plus). auto. }
-    iMod ("Hcloseβ" with "Hδ") as "Hβ". iMod ("Hcloseα1" with "[$H↦]") as "Hα1".
+    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)]%TC
-        with "[] LFT Hna [Hα1 Hα2 Hβ] HL Hk"); first last.
+        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. iExists _, _, _, _, _.
       iFrame "∗#". }
-    { rewrite /elctx_interp big_sepL_cons big_sepL_singleton. iFrame. }
     iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
   Qed.
@@ -114,29 +116,28 @@ Section ref_functions.
       letalloc: "r" <- "r'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma ref_deref_type ty :
+  Lemma ref_deref_type ty `{!TyWf ty} :
     typed_val ref_deref
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL [# &shr{α}(ref β ty)]%T (&shr{α}ty)%T)).
+      (fn (fun '(α, β) => FP_wf ∅ [# &shr{α}(ref β ty)]%T (&shr{α}ty)%T)).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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 (ν q γ δ ty' lv lrc) "#(Hαν & Hfrac & Hshr & Hx')".
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "(Hα & Hβ & #Hαβ)".
+    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↦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]%TC
-        with "[] LFT Hna [Hα Hβ Hαβ] HL Hk"); first last.
+        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. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
     iApply (type_letalloc_1 (&shr{α}ty)); [solve_typing..|].
     iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
@@ -152,18 +153,18 @@ Section ref_functions.
       delete [ #2; "x"];;
       let: "r" := new [ #0] in return: ["r"].
 
-  Lemma ref_drop_type ty :
-    typed_val ref_drop (fn(∀ α, [α]; ref α ty) → unit).
+  Lemma ref_drop_type ty `{!TyWf ty} :
+    typed_val ref_drop (fn(∀ α, ∅; ref α ty) → unit).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
-    iIntros (tid qE) "#LFT Hna Hα HL Hk Hx".
-    rewrite {1}/elctx_interp big_sepL_singleton tctx_interp_singleton tctx_hasty_val.
+      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"; 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â—¯]")
@@ -192,12 +193,11 @@ Section ref_functions.
     wp_bind Endlft. iApply (wp_mask_mono (↑lftN)); 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α".
-    iApply (type_type _ _ _ [ #lx ◁ box (uninit 2)]%TC with "[] LFT Hna [Hα] HL Hk");
+    iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
+    iApply (type_type _ _ _ [ #lx ◁ box (uninit 2)]%TC 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. }
-    { by rewrite /elctx_interp big_sepL_singleton. }
     iApply type_delete; [solve_typing..|].
     iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
     iApply type_jump; solve_typing.
@@ -206,67 +206,73 @@ Section ref_functions.
   (* Apply a function within the ref, typically for accessing a component. *)
   Definition ref_map : val :=
     funrec: <> ["ref"; "f"; "env"] :=
-    withcont: "k":
-      let: "x'" := !"ref" in
       let: "f'" := !"f" in
+      Newlft;;
+      let: "x'" := !"ref" in
       letalloc: "x" <- "x'" in
       letcall: "r" := "f'" ["env"; "x"]%E in
-      let: "r'" := !"r" in
+      let: "r'" := !"r" in delete [ #1; "r"];;
+      Endlft;;
       "ref" <- "r'";;
-      delete [ #1; "f"];; "k" []
-    cont: "k" [] :=
-      return: ["ref"].
+      delete [ #1; "f"];; return: ["ref"].
 
-  Lemma ref_map_type ty1 ty2 envty E :
+  Lemma ref_map_type ty1 ty2 envty `{!TyWf ty1, !TyWf ty2, !TyWf envty} :
     typed_val ref_map
-      (fn(∀ β, [β]%EL ++ E;
-            ref β ty1, fn(∀ α, [α]%EL ++ E; envty, &shr{α}ty1) → &shr{α}ty2, envty)
+      (fn(∀ β, ∅; ref β ty1, fn(∀ α, ∅; envty, &shr{α}ty1) → &shr{α}ty2, envty)
        → ref β ty2).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ret arg).
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ϝ ret arg).
        inv_vec arg=>ref f env. simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT".
-    rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
-    iDestruct "HT" as "(Href & Hf & Henv)".
-    destruct ref as [[|lref|]|]; try done. iDestruct "Href" as "[Href Href†]".
+    iApply type_deref; [solve_typing..|]. iIntros (f'). simpl_subst.
+    iApply (type_newlft [ϝ]). iIntros (κ tid) "#LFT #HE Hna HL Hk (Hf & #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";
       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γ)".
-    rewrite -(freeable_sz_split _ 1 1). iDestruct "Href†" as "[Href†1 Href†2]".
-    destruct (Qp_lower_bound qE qν) as (q & qE' & qν' & -> & ->).
-    iDestruct (@fractional_split with "HE") as "[[Hα HE] HE']".
-    iDestruct "Hν" as "[Hν Hν']".
-    remember (RecV "k" [] (ret [ LitV lref])%E)%V as k eqn:EQk.
-    iApply (wp_let' _ _ _ _ k). { subst. solve_to_val. } simpl_subst.
-    iApply (type_type ((α ⊓ ν) :: E)%EL []
-        [k ◁cont([], λ _:vec val 0, [ #lref ◁ own_ptr 2 (&shr{α ⊓ ν}ty2)])]%CC
-        [ f ◁ box (fn(∀ α, [α]%EL ++ E; envty, &shr{α}ty1) → &shr{α}ty2);
-          #lref ◁ own_ptr 2 (&shr{α ⊓ ν}ty1); env ◁ box envty ]%TC
-       with "[] LFT Hna [$HE Hα Hν] HL [Hk HE' Hν' Href↦2 Hγ Href†2]"); first last.
-    { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
-      iFrame. iApply tctx_hasty_val'. done. iFrame. iExists [_].
-      rewrite heap_mapsto_vec_singleton. by iFrame. }
-    { rewrite !cctx_interp_singleton /=. iIntros "HE". iIntros (args) "Hna HL HT".
-      inv_vec args. subst. simpl. wp_rec. iDestruct "HE" as "[Hαν HE]".
-      iDestruct (lft_tok_sep with "Hαν") as "[Hα Hν]".
-      iSpecialize ("Hk" with "[$HE' $Hα $HE]").
-      iApply ("Hk" $! [# #lref] with "Hna HL").
-      rewrite !tctx_interp_singleton !tctx_hasty_val' //.
-      iDestruct "HT" as "[Href Ḥref†2]".
-      rewrite /= -(freeable_sz_split _ 1 1). iFrame.
-      iDestruct "Href" as ([|[[|lv'|]|] [|]]) "[H↦ Href]"; auto.
-      iExists [ #lv'; #lrc].
-      rewrite (heap_mapsto_vec_cons _ _ _ [_]) !heap_mapsto_vec_singleton. iFrame.
-      iExists ν, (q+qν')%Qp. eauto 10 with iFrame. }
-    { rewrite /= -lft_tok_sep. iFrame. }
-    iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
-    iApply type_deref; [solve_typing..|]. iIntros (f'). simpl_subst.
-    iApply type_letalloc_1; [solve_typing..|]. iIntros (x). simpl_subst.
-    iApply (type_letcall); [simpl; solve_typing..|]. iIntros (r). simpl_subst.
-    iApply type_deref; [solve_typing..|]. iIntros (r'). simpl_subst.
-    iApply type_assign; [solve_typing..|].
+    wp_read. wp_let. wp_apply wp_new; first done.
+    iIntros (lx [|? []]) "(% & H† & Hlx)"; try (simpl in *; lia).
+    rewrite heap_mapsto_vec_singleton. wp_let. wp_write.
+    match goal with | |- context [(WP (_ [?k']) {{_, _}})%I] =>
+      assert (∃ k, to_val k' = Some k) as [k EQk] by (eexists; solve_to_val) end.
+    iApply (wp_let' _ _ _ _ k _ EQk). simpl_subst. iNext.
+    iDestruct (lctx_lft_incl_incl κ α with "HL HE") as "#Hκα"; [solve_typing..|].
+    iMod (bor_create _ κ (qν).[ν] with "LFT [$Hν]") as "[Hb Hν]"; first done.
+    iAssert (κ ⊑ α ⊓ ν)%I with "[>Hb]" as "#Hκν".
+    { iApply (lft_incl_glb with "Hκα"). iApply (frac_bor_lft_incl with "LFT").
+      iApply (bor_fracture with "LFT [> -]"); first done. rewrite /= Qp_mult_1_r //. }
+    iApply (type_type ((κ ⊑ α ⊓ ν) :: (α ⊓ ν ⊑ α) :: _)%EL _
+        [k ◁cont(_, λ x:vec val 1, [ x!!!0 ◁ box (&shr{α ⊓ ν}ty2)])]%CC
+        [ f' ◁ fn(∀ α, ∅; envty, &shr{α}ty1) → &shr{α}ty2;
+          #lx ◁ box (&shr{α ⊓ ν}ty1); env ◁ box envty ]%TC
+       with "[] LFT [] Hna HL [-H† Hlx Henv]"); swap 1 2; swap 3 4.
+    { iSplitL; last iSplitL; [done|iApply lft_intersect_incl_l|iApply "HE"]. }
+    { iApply (type_call (α ⊓ ν)); solve_typing. }
+    { iFrame "∗#". iApply tctx_hasty_val'; first done. rewrite -freeable_sz_full.
+      iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. auto. }
+    iIntros (? ->%elem_of_list_singleton arg) "Hna HL Hr". inv_vec arg=>r.
+    apply of_to_val in EQk. rewrite EQk. iApply wp_rec; try (done || apply _).
+    { repeat econstructor. by rewrite to_of_val. } simpl_subst.
+    rewrite /tctx_interp big_sepL_singleton (tctx_hasty_val _ r) ownptr_own.
+    iDestruct "Hr" as (lr vr) "(% & Hlr & Hvr & H†)". subst. inv_vec vr=>r'. iNext.
+    rewrite heap_mapsto_vec_singleton. wp_read. wp_let.
+    wp_apply (wp_delete _ _ _ [_] with "[Hlr H†]"). done.
+    { rewrite heap_mapsto_vec_singleton freeable_sz_full. iFrame. }
+    iIntros "_". wp_seq. wp_bind Endlft. iDestruct "HL" as "[Hκ HL]".
+    iDestruct "Hκ" as (κ') "(% & Hκ' & #Hκ'†)". simpl in *. subst κ.
+    iSpecialize ("Hκ'†" with "Hκ'").
+    iApply wp_mask_mono; last iApply (wp_step_fupd with "Hκ'†"); auto with ndisj.
+    wp_seq. iIntros "Hκ'† !>". iMod ("Hν" with "[Hκ'†]") as "Hν";
+      first by rewrite -lft_dead_or; auto. wp_seq. wp_write.
+    iApply (type_type _ _ _
+        [ f ◁ box (fn(∀ α, ∅; envty, &shr{α}ty1) → &shr{α}ty2);
+          #lref ◁ box (ref α ty2) ]%TC
+       with "[] LFT HE Hna HL Hk"); first last.
+    { iFrame. rewrite big_sepL_singleton tctx_hasty_val. iExists _. iSplit. done.
+      iFrame. iExists [_;_]. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton.
+      iFrame. destruct r' as [[]|]=>//=. auto 10 with iFrame. }
     iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
   Qed.
diff --git a/theories/typing/lib/refcell/refcell.v b/theories/typing/lib/refcell/refcell.v
index 4d3ee9f1ee303cf89931e84eab66562f5161e98b..f34f6305b309983871d49fd916a7307be9773522 100644
--- a/theories/typing/lib/refcell/refcell.v
+++ b/theories/typing/lib/refcell/refcell.v
@@ -62,17 +62,18 @@ Section refcell_inv.
     eapply refcell_inv_type_ne, type_dist_dist2. done.
   Qed.
 
-  Lemma refcell_inv_proper E L tid l γ α ty1 ty2 :
+  Lemma refcell_inv_proper E L ty1 ty2 q :
     eqtype E L ty1 ty2 →
-    elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗
-    refcell_inv tid l γ α ty1 -∗ refcell_inv tid l γ α 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. *)
-    iIntros (Hty%eqtype_unfold) "#HE #HL H". unfold refcell_inv.
-    iDestruct (Hty with "HE HL") as "(% & Hown & Hshr)".
+    (* 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 "(% & #Hown & #Hshr)".
     iAssert (□ (&{α} shift_loc l 1 ↦∗: ty_own ty1 tid -∗
-                &{α} shift_loc l 1 ↦∗: ty_own ty2 tid))%I with "[]" as "#Hb".
+                &{α} shift_loc l 1 ↦∗: ty_own ty2 tid))%I as "#Hb".
     { iIntros "!# H". iApply bor_iff; last done.
       iSplit; iIntros "!>!#H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl;
       iFrame; by iApply "Hown". }
@@ -155,6 +156,9 @@ Section refcell.
     iExists _, _. iFrame. iApply lft_incl_trans; auto.
   Qed.
 
+  Global Instance refcell_wf ty `{!TyWf ty} : TyWf (refcell ty) :=
+    { ty_lfts := ty.(ty_lfts); ty_wf_E := ty.(ty_wf_E) }.
+
   Global Instance refcell_type_ne : TypeNonExpansive refcell.
   Proof.
     constructor;
@@ -171,14 +175,17 @@ Section refcell.
   Proof.
     (* TODO : this proof is essentially [solve_proper], but within the logic.
               It would easily gain from some automation. *)
-    iIntros (ty1 ty2 EQ) "#HE #HL". pose proof EQ as EQ'%eqtype_unfold.
-    iDestruct (EQ' with "HE HL") as "(% & #Hown & #Hshr)".
+    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; [|iSplit; iIntros "!# * H"].
     - iPureIntro. simpl. congruence.
     - destruct vl as [|[[]|]]; try done. iDestruct "H" as "[$ H]". by iApply "Hown".
     - iDestruct "H" as (a γ) "[Ha H]". iExists a, γ. iFrame.
-      iApply na_bor_iff; last done.
-      iSplit; iIntros "!>!# H"; by iApply refcell_inv_proper.
+      iApply na_bor_iff; last done. 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).
diff --git a/theories/typing/lib/refcell/refcell_code.v b/theories/typing/lib/refcell/refcell_code.v
index 5451b568d531047c288ae4de375477f0d6929ce3..6f98642452ef6bef6faad1c71f124c575c1b4aad 100644
--- a/theories/typing/lib/refcell/refcell_code.v
+++ b/theories/typing/lib/refcell/refcell_code.v
@@ -19,13 +19,13 @@ Section refcell_functions.
       "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).
+  Lemma refcell_new_type ty `{!TyWf ty} :
+    typed_val (refcell_new ty) (fn(∅; ty) → refcell ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_new; [solve_typing..|].
-    iIntros (r tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    iIntros (r tid) "#LFT #HE Hna HL Hk HT". simpl_subst.
     rewrite (Nat2Z.id (S ty.(ty_size))) tctx_interp_cons
             tctx_interp_singleton !tctx_hasty_val.
     iDestruct "HT" as "[Hr Hx]".
@@ -36,7 +36,7 @@ Section refcell_functions.
     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)]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        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.
@@ -54,13 +54,13 @@ Section refcell_functions.
              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).
+  Lemma refcell_into_inner_type ty `{!TyWf ty} :
+    typed_val (refcell_into_inner ty) (fn(∅; refcell ty) → ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_new; [solve_typing..|].
-    iIntros (r tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    iIntros (r tid) "#LFT #HE Hna HL Hk HT". simpl_subst.
     rewrite (Nat2Z.id (ty.(ty_size))) tctx_interp_cons
             tctx_interp_singleton !tctx_hasty_val.
     iDestruct "HT" as "[Hr Hx]".
@@ -72,7 +72,7 @@ Section refcell_functions.
     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]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        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.
@@ -88,13 +88,14 @@ Section refcell_functions.
       "x" <- "x'" +â‚— #1;;
       return: ["x"].
 
-  Lemma refcell_get_mut_type ty :
-    typed_val refcell_get_mut (fn(∀ α, [α]; &uniq{α} (refcell ty)) → &uniq{α} ty)%T.
+  Lemma refcell_get_mut_type ty `{!TyWf ty} :
+    typed_val refcell_get_mut
+              (fn(∀ α, ∅; &uniq{α} (refcell ty)) → &uniq{α} ty)%T.
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL HC HT".
+    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⌝) ∗
@@ -110,7 +111,7 @@ Section refcell_functions.
     iDestruct "Hx" as (vl) "[Hx↦ Hx]". rewrite uninit_own. wp_op.
     iApply (type_type _ _ _
             [ #lx ◁ box (uninit 1); #(shift_loc lx' 1) ◁ &uniq{α}ty]%TC
-            with "[] LFT Hna HE HL HC [-]"); last first.
+            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..|].
@@ -138,33 +139,35 @@ Section refcell_functions.
     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)).
+  Lemma refcell_try_borrow_type ty `{!TyWf ty} :
+    typed_val refcell_try_borrow
+      (fn(∀ α, ∅; &shr{α}(refcell ty)) → option (ref α ty)).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      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))])%TC);
+    iApply (type_cont [] [ϝ ⊑ []]%LL (λ _, [x ◁ box (&shr{α} refcell ty);
+                                         r ◁ box (option (ref α ty))])%TC);
       [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 qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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αβ Hinv]".
-    rewrite {1}/elctx_interp big_sepL_singleton.
-    iMod (lft_incl_acc with "Hαβ HE") as (qβ) "[[Hβtok1 Hβtok2] Hclose]". done.
-    iMod (na_bor_acc with "LFT Hinv Hβtok1 Hna") as "(INV & Hna & Hclose')"; try done.
+    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=>?; wp_if.
-    - iMod ("Hclose'" with "[Hlx Hownst Hst] Hna") as "[Hβtok1 Hna]";
+    - 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) ]%TC
-              with "[] LFT Hna [> Hclose Hβtok1 Hβtok2] HL Hk"); first last.
+              with "[] LFT HE Hna HL Hk"); first last.
       { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
-      { rewrite {1}/elctx_interp big_sepL_singleton /=. iApply "Hclose". by 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..|].
@@ -206,17 +209,17 @@ Section refcell_functions.
           + iIntros "!> Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro.
             iNext. iMod "Hν". iApply "Hh". rewrite -lft_dead_or. auto.
           + iExists _. iFrame. by rewrite Qp_div_2. }
-      iMod ("Hclose'" with "[$INV] Hna") as "[Hβtok1 Hna]".
+      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)]%TC
-              with "[] LFT Hna [> Hclose Hβtok1 Hβtok2] HL Hk");
-        first last.
+              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. iExists _, _, _, _, _. iFrame "∗#". iApply ty_shr_mono; try by auto.
         iApply lft_intersect_mono. done. iApply lft_incl_refl. }
-      { rewrite {1}/elctx_interp big_sepL_singleton /=. iApply "Hclose". by iFrame. }
       iApply (type_sum_memcpy (option $ ref α ty)); [solve_typing..|].
       simpl. iApply type_delete; [solve_typing..|].
       iApply type_jump; solve_typing.
@@ -243,27 +246,27 @@ Section refcell_functions.
     cont: "k" [] :=
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma refcell_try_borrow_mut_type ty :
+  Lemma refcell_try_borrow_mut_type ty `{!TyWf ty} :
     typed_val refcell_try_borrow_mut
-              (fn(∀ α, [α]; &shr{α}(refcell ty)) → option (refmut α ty))%T.
+              (fn(∀ α, ∅; &shr{α}(refcell ty)) → option (refmut α ty))%T.
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      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))]%TC));
+    iApply (type_cont [] [ϝ ⊑ []]%LL (λ _, [x ◁ box (&shr{α} refcell ty);
+                                            r ◁ box (option (refmut α ty))]%TC));
       [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 qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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αβ Hinv]".
-    rewrite {1}/elctx_interp big_sepL_singleton.
-    iMod (lft_incl_acc with "Hαβ HE") as (qβ) "[Hβtok Hclose]". done.
-    iMod (na_bor_acc with "LFT Hinv Hβtok Hna") as "(INV & Hna & Hclose')"; try done.
+    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=>?; wp_if.
     - wp_write. wp_apply wp_new; [done..|].
       iIntros (lref vl) "(EQ & H† & Hlref)". iDestruct "EQ" as %?%(inj Z.of_nat 2%nat).
@@ -277,29 +280,29 @@ Section refcell_functions.
       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] Hna") as "[Hβtok Hna]".
+      iModIntro. iMod ("Hclose''" with "[Hlx Hownst Hbh] Hna") as "[Hβtok Hna]".
       { iExists _. iFrame. iExists ν. iSplit; first by auto. iNext. iSplitL; last by auto.
         iIntros "Hν". iMod ("Hhν" with "Hν") as "Hν". iModIntro. iNext. iMod "Hν".
         iApply "Hbh". rewrite -lft_dead_or. auto. }
+      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)]%TC
-              with "[] LFT Hna [> Hclose Hβtok] HL Hk"); first last.
+              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. iExists _, _, _, _. iFrame "#∗". iApply (bor_shorten with "[] [$Hb]").
         iApply lft_intersect_mono; first done. iApply lft_incl_refl. }
-      { rewrite {1}/elctx_interp big_sepL_singleton /=. iApply "Hclose". by iFrame. }
       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]";
+    - 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) ]%TC
-              with "[] LFT Hna [> Hclose Hβtok] HL Hk"); first last.
+              with "[] LFT HE Hna HL Hk"); first last.
       { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
-      { rewrite {1}/elctx_interp big_sepL_singleton /=. iApply "Hclose". by iFrame. }
       iApply (type_sum_unit (option $ refmut α ty)); [solve_typing..|].
       simpl. iApply type_jump; solve_typing.
   Qed.
diff --git a/theories/typing/lib/refcell/refmut.v b/theories/typing/lib/refcell/refmut.v
index cfe1fbc3b6576017313a33613843443cb225d01a..98a2abe3304fcd4e7f1608725f3fd3a456d88616 100644
--- a/theories/typing/lib/refcell/refmut.v
+++ b/theories/typing/lib/refcell/refmut.v
@@ -72,6 +72,9 @@ Section refmut.
       iApply ty_shr_mono; try done. iApply lft_intersect_mono. iApply lft_incl_refl. done.
   Qed.
 
+  Global Instance refmut_wf α ty `{!TyWf ty} : TyWf (refmut α ty) :=
+    { ty_lfts := [α]; ty_wf_E := ty.(ty_wf_E) ++ ty.(ty_outlives_E) α }.
+
   Global Instance refmut_type_contractive α : TypeContractive (refmut α).
   Proof. solve_type_proper. Qed.
   Global Instance refmut_ne α : NonExpansive (refmut α).
@@ -80,11 +83,10 @@ Section refmut.
   Global Instance refmut_mono E L :
     Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) refmut.
   Proof.
-    iIntros (α1 α2 Hα ty1 ty2 Hty) "#HE #HL".
-    pose proof Hty as Hty'%eqtype_unfold.
-    iDestruct (Hty' with "HE HL") as "(%&#Ho&#Hs)".
-    iDestruct (Hα with "HE HL") as "Hα1α2".
-    iSplit; [|iSplit; iAlways].
+    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; [|iSplit; iAlways].
     - done.
     - iIntros (tid [|[[]|][|[[]|][]]]) "H"; try done.
       iDestruct "H" as (ν γ β ty') "(Hb & #H⊑ & #Hinv & Hν & Hown)".
diff --git a/theories/typing/lib/refcell/refmut_code.v b/theories/typing/lib/refcell/refmut_code.v
index a1d838cf20bd095db614e83f800d845c75349339..e2649e7c48431684d5a34e7f503711f5d0af1207 100644
--- a/theories/typing/lib/refcell/refmut_code.v
+++ b/theories/typing/lib/refcell/refmut_code.v
@@ -18,33 +18,36 @@ Section refmut_functions.
       letalloc: "r" <- "r'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma refmut_deref_type ty :
-    typed_val refmut_deref
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL [# &shr{α}(refmut β ty)]%T (&shr{α}ty))).
+  Lemma refmut_deref_type ty `{!TyWf ty} :
+    typed_val refmut_deref (fn (fun '(α, β) =>
+       FP_wf ∅ [# &shr{α}(refmut β ty)]%T (&shr{α}ty))).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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 (lv lrc) "#(Hfrac & #Hshr)".
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "([Hα1 Hα2] & Hβ & #Hαβ)".
+    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 (!(LitV lx'))%E. iApply (wp_step_fupd with "[Hα2β]");
          [done| |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]%TC
-            with "[] LFT Hna [Hα1 Hα2 Hβ Hαβ] HL Hk"); last first.
+            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.
-      iApply lft_incl_refl. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
+      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.
@@ -58,21 +61,20 @@ Section refmut_functions.
       letalloc: "r" <- "r'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma refmut_derefmut_type ty :
-    typed_val refmut_derefmut
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL [# &uniq{α}(refmut β ty)]%T
-                              (&uniq{α}ty)%T)).
+  Lemma refmut_derefmut_type ty `{!TyWf ty} :
+    typed_val refmut_derefmut (fn (fun '(α, β) =>
+      FP_wf ∅ [# &uniq{α}(refmut β ty)]%T (&uniq{α}ty)%T)).
    Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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']".
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "(Hα & Hβ & #Hαβ)". destruct x' as [[|lx'|]|]; try done.
+    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 [|[[|lv|]|][|[[|lrc|]|][]]];
       try by iMod (bor_persistent_tok with "LFT H Hα") as "[>[] _]".
     iMod (bor_exists with "LFT H") as (ν) "H". done.
@@ -92,13 +94,13 @@ Section refmut_functions.
     wp_bind (!(LitV lx'))%E. iApply (wp_step_fupd with "[Hb]");
       [done| |by iApply (bor_unnest with "LFT Hb")|]; first done.
     wp_read. iIntros "Hb !>". wp_let.
-    iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]".
+    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]%TC
-            with "[] LFT Hna [Hα Hβ Hαβ] HL Hk"); last first.
+            with "[] LFT HE Hna HL Hk"); last first.
     { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
       iFrame. iApply (bor_shorten with "[] Hb").
       by iApply lft_incl_glb; [iApply lft_incl_glb|iApply lft_incl_refl]. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
     iApply (type_letalloc_1 (&uniq{α}ty)); [solve_typing..|].
     iIntros (r). simpl_subst. iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
@@ -113,18 +115,19 @@ Section refmut_functions.
       delete [ #2; "x"];;
       let: "r" := new [ #0] in return: ["r"].
 
-  Lemma refmut_drop_type ty :
-    typed_val refmut_drop (fn(∀ α, [α]; refmut α ty) → unit).
+  Lemma refmut_drop_type ty `{!TyWf ty} :
+    typed_val refmut_drop (fn(∀ α, ∅; refmut α ty) → unit).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
-    iIntros (tid qE) "#LFT Hna Hα HL Hk Hx".
-    rewrite {1}/elctx_interp big_sepL_singleton tctx_interp_singleton tctx_hasty_val.
+      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"; 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 (ν γ β ty') "(Hb & #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 "INV" as (st) "(H↦lrc & H● & INV)". wp_write.
@@ -142,12 +145,11 @@ Section refmut_functions.
     { iExists None. iFrame. iMod (own_update_2 with "H● H◯") as "$"; last done.
       apply auth_update_dealloc. rewrite -(right_id None _ (Some _)).
       apply cancel_local_update_empty, _. }
-    iMod ("Hcloseα" with "Hβ") as "Hα". wp_seq.
+    iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL". wp_seq.
     iApply (type_type _ _ _ [ #lx ◁ box (uninit 2)]%TC
-            with "[] LFT Hna [Hα] HL Hk"); last first.
+            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. }
-    { by rewrite /elctx_interp big_sepL_singleton. }
     iApply type_delete; [solve_typing..|].
     iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
     iApply type_jump; solve_typing.
@@ -156,69 +158,74 @@ Section refmut_functions.
   (* Apply a function within the refmut, typically for accessing a component. *)
   Definition refmut_map : val :=
     funrec: <> ["ref"; "f"; "env"] :=
-    withcont: "k":
-      let: "x'" := !"ref" in
       let: "f'" := !"f" in
+      Newlft;;
+      let: "x'" := !"ref" in
       letalloc: "x" <- "x'" in
       letcall: "r" := "f'" ["env"; "x"]%E in
-      let: "r'" := !"r" in
+      let: "r'" := !"r" in delete [ #1; "r"];;
+      Endlft;;
       "ref" <- "r'";;
-      delete [ #1; "f"];; "k" []
-    cont: "k" [] :=
-      return: ["ref"].
+      delete [ #1; "f"];; return: ["ref"].
 
-  Lemma refmut_map_type ty1 ty2 envty E :
+  Lemma refmut_map_type ty1 ty2 envty `{!TyWf ty1, !TyWf ty2, !TyWf envty} :
     typed_val refmut_map
-      (fn(∀ β, [β]%EL ++ E; refmut β ty1,
-                         fn(∀ α, [α]%EL ++ E; envty, &uniq{α} ty1) → &uniq{α} ty2,
-                         envty)
+      (fn(∀ β, ∅; refmut β ty1,
+            fn(∀ α, ∅; envty, &uniq{α} ty1) → &uniq{α} ty2, envty)
        → refmut β ty2).
   Proof.
-    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>ref f env. simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT".
-    rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
-    iDestruct "HT" as "(Href & Hf & Henv)".
-    destruct ref as [[|lref|]|]; try done. iDestruct "Href" as "[Href Href†]".
+    intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#". iIntros (α ϝ ret arg).
+       inv_vec arg=>ref f env. simpl_subst.
+    iApply type_deref; [solve_typing..|]. iIntros (f'). simpl_subst.
+    iApply (type_newlft [ϝ]). iIntros (κ tid) "#LFT #HE Hna HL Hk (Hf & #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";
       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 (ν γ β ty') "(Hbor & #Hαβ & #Hinv & >Hν & Hγ)".
-    rewrite -(freeable_sz_split _ 1 1). iDestruct "Href†" as "[Href†1 Href†2]".
-    destruct (Qp_lower_bound qE 1) as (q & qE' & qν' & -> & EQ1).
-      rewrite [in (_).[ν]%I] EQ1.
-    iDestruct (@fractional_split with "HE") as "[[Hα HE] HE']".
-    iDestruct "Hν" as "[Hν Hν']".
-    remember (RecV "k" [] (ret [ LitV lref])%E)%V as k eqn:EQk.
-    iApply (wp_let' _ _ _ _ k). { subst. solve_to_val. } simpl_subst.
-    iApply (type_type ((α ⊓ ν) :: E)%EL []
-        [k ◁cont([], λ _:vec val 0, [ #lref ◁ own_ptr 2 (&uniq{α ⊓ ν}ty2)])]%CC
-        [ f ◁ box (fn(∀ α, [α]%EL ++ E; envty, &uniq{α}ty1) → &uniq{α}ty2);
-          #lref ◁ own_ptr 2 (&uniq{α ⊓ ν}ty1); env ◁ box envty ]%TC
-       with "[] LFT Hna [$HE Hα Hν] HL [Hk HE' Hν' Href↦2 Hγ Href†2]"); first last.
-    { rewrite 2!tctx_interp_cons tctx_interp_singleton !tctx_hasty_val.
-      iFrame. iApply tctx_hasty_val'. done. iFrame. iExists [_].
-      rewrite heap_mapsto_vec_singleton. by iFrame. }
-    { rewrite !cctx_interp_singleton /=. iIntros "HE". iIntros (args) "Hna HL HT".
-      inv_vec args. subst. simpl. wp_rec. iDestruct "HE" as "[Hαν HE]".
-      iDestruct (lft_tok_sep with "Hαν") as "[Hα Hν]".
-      iSpecialize ("Hk" with "[$HE' $Hα $HE]").
-      iApply ("Hk" $! [# #lref] with "Hna HL").
-      rewrite !tctx_interp_singleton !tctx_hasty_val' //.
-      iDestruct "HT" as "[Href Ḥref†2]".
-      rewrite /= -(freeable_sz_split _ 1 1). iFrame.
-      iDestruct "Href" as ([|[[|lv'|]|] [|]]) "[H↦ Href]"; auto.
-      iExists [ #lv'; #lrc].
-      rewrite (heap_mapsto_vec_cons _ _ _ [_]) !heap_mapsto_vec_singleton. iFrame.
-      iExists ν. rewrite /= EQ1. eauto 20 with iFrame. }
-    { rewrite /= -lft_tok_sep. iFrame. }
-    iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
-    iApply type_deref; [solve_typing..|]. iIntros (f'). simpl_subst.
-    iApply type_letalloc_1; [solve_typing..|]. iIntros (x). simpl_subst.
-    iApply (type_letcall); [simpl; solve_typing..|]. iIntros (r). simpl_subst.
-    iApply type_deref; [solve_typing..|]. iIntros (r'). simpl_subst.
-    iApply type_assign; [solve_typing..|].
+    wp_read. wp_let. wp_apply wp_new; first done.
+    iIntros (lx [|? []]) "(% & H† & Hlx)"; try (simpl in *; lia).
+    rewrite heap_mapsto_vec_singleton. wp_let. wp_write.
+    match goal with | |- context [(WP (_ [?k']) {{_, _}})%I] =>
+      assert (∃ k, to_val k' = Some k) as [k EQk] by (eexists; solve_to_val) end.
+    iApply (wp_let' _ _ _ _ k _ EQk). simpl_subst. iNext.
+    iDestruct (lctx_lft_incl_incl κ α with "HL HE") as "#Hκα"; [solve_typing..|].
+    iMod (bor_create _ κ (1).[ν] with "LFT [$Hν]") as "[Hb Hν]"; first done.
+    iAssert (κ ⊑ α ⊓ ν)%I with "[>Hb]" as "#Hκν".
+    { iApply (lft_incl_glb with "Hκα"). iApply (frac_bor_lft_incl with "LFT").
+      iApply (bor_fracture with "LFT [> -]"); first done. rewrite /= Qp_mult_1_r //. }
+    iApply (type_type ((κ ⊑ α ⊓ ν) :: (α ⊓ ν ⊑ α) :: _)%EL _
+        [k ◁cont(_, λ x:vec val 1, [ x!!!0 ◁ box (&uniq{α ⊓ ν}ty2)])]%CC
+        [ f' ◁ fn(∀ α, ∅; envty, &uniq{α}ty1) → &uniq{α}ty2;
+          #lx ◁ box (&uniq{α ⊓ ν}ty1); env ◁ box envty ]%TC
+       with "[] LFT [] Hna HL [-H† Hlx Henv Hbor]"); swap 1 2; swap 3 4.
+    { iSplitL; last iSplitL; [done|iApply lft_intersect_incl_l|iApply "HE"]. }
+    { iApply (type_call (α ⊓ ν)); solve_typing. }
+    { iFrame "∗#". iApply tctx_hasty_val'; first done. rewrite -freeable_sz_full.
+      iFrame. iExists [_]. rewrite heap_mapsto_vec_singleton. auto with iFrame. }
+    iIntros (? ->%elem_of_list_singleton arg) "Hna HL Hr". inv_vec arg=>r.
+    apply of_to_val in EQk. rewrite EQk. iApply wp_rec; try (done || apply _).
+    { repeat econstructor. by rewrite to_of_val. } simpl_subst.
+    rewrite /tctx_interp big_sepL_singleton (tctx_hasty_val _ r) ownptr_own.
+    iDestruct "Hr" as (lr vr) "(% & Hlr & Hvr & H†)". subst. inv_vec vr=>r'. iNext.
+    rewrite heap_mapsto_vec_singleton. wp_read. wp_let.
+    wp_apply (wp_delete _ _ _ [_] with "[Hlr H†]"). done.
+    { rewrite heap_mapsto_vec_singleton freeable_sz_full. iFrame. }
+    iIntros "_". wp_seq. wp_bind Endlft. iDestruct "HL" as "[Hκ HL]".
+    iDestruct "Hκ" as (κ') "(% & Hκ' & #Hκ'†)". simpl in *. subst κ.
+    iSpecialize ("Hκ'†" with "Hκ'").
+    iApply wp_mask_mono; last iApply (wp_step_fupd with "Hκ'†"); auto with ndisj.
+    wp_seq. iIntros "Hκ'† !>". iMod ("Hν" with "[Hκ'†]") as "Hν";
+      first by rewrite -lft_dead_or; auto. wp_seq. wp_write.
+    iApply (type_type _ _ _
+        [ f ◁ box (fn(∀ α, ∅; envty, &uniq{α}ty1) → &uniq{α}ty2);
+          #lref ◁ box (refmut α ty2) ]%TC
+       with "[] LFT HE Hna HL Hk"); first last.
+    { iFrame. rewrite big_sepL_singleton tctx_hasty_val. iExists _. iSplit. done.
+      iFrame. iExists [_;_]. rewrite heap_mapsto_vec_cons heap_mapsto_vec_singleton.
+      iFrame. destruct r' as [[]|]=>//=. auto 10 with iFrame. }
     iApply type_delete; [solve_typing..|].
     iApply type_jump; solve_typing.
   Qed.
diff --git a/theories/typing/lib/rwlock/rwlock.v b/theories/typing/lib/rwlock/rwlock.v
index 049e5ce511eff50838d2077d003f0fa272fb9c19..db17b3e5830be1ea7c2944aa7cf6055eda0c95d5 100644
--- a/theories/typing/lib/rwlock/rwlock.v
+++ b/theories/typing/lib/rwlock/rwlock.v
@@ -58,17 +58,18 @@ Section rwlock_inv.
     intros n ???. eapply rwlock_inv_type_ne, type_dist_dist2. done.
   Qed.
 
-  Lemma rwlock_inv_proper E L tid l γ α ty1 ty2 :
+  Lemma rwlock_inv_proper E L ty1 ty2 q :
     eqtype E L ty1 ty2 →
-    elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗
-    rwlock_inv tid l γ α ty1 -∗ rwlock_inv tid l γ α ty2.
+    llctx_interp L q -∗ ∀ tid l γ α, □ (elctx_interp E -∗
+    rwlock_inv tid l γ α ty1 -∗ rwlock_inv tid l γ α ty2).
   Proof.
     (* TODO : this proof is essentially [solve_proper], but within the logic.
               It would easily gain from some automation. *)
-    iIntros (Hty%eqtype_unfold) "#HE #HL H". unfold rwlock_inv.
-    iDestruct (Hty with "HE HL") as "(% & Hown & Hshr)".
+    rewrite eqtype_unfold. iIntros (Hty) "HL".
+    iDestruct (Hty with "HL") as "#Hty". iIntros "* !# #HE H".
+    iDestruct ("Hty" with "HE") as "(% & #Hown & #Hshr)".
     iAssert (□ (&{α} shift_loc l 1 ↦∗: ty_own ty1 tid -∗
-                &{α} shift_loc l 1 ↦∗: ty_own ty2 tid))%I with "[]" as "#Hb".
+                &{α} shift_loc l 1 ↦∗: ty_own ty2 tid))%I as "#Hb".
     { iIntros "!# H". iApply bor_iff; last done.
       iSplit; iIntros "!>!#H"; iDestruct "H" as (vl) "[Hf H]"; iExists vl;
       iFrame; by iApply "Hown". }
@@ -148,6 +149,9 @@ Section rwlock.
     iExists _, _. iFrame. iApply lft_incl_trans; auto.
   Qed.
 
+  Global Instance join_handle_wf ty `{!TyWf ty} : TyWf (rwlock ty) :=
+    { ty_lfts := ty.(ty_lfts); ty_wf_E := ty.(ty_wf_E) }.
+
   Global Instance rwlock_type_ne : TypeNonExpansive rwlock.
   Proof.
     constructor;
@@ -164,14 +168,17 @@ Section rwlock.
   Proof.
     (* TODO : this proof is essentially [solve_proper], but within the logic.
               It would easily gain from some automation. *)
-    iIntros (ty1 ty2 EQ) "#HE #HL". pose proof EQ as EQ'%eqtype_unfold.
-    iDestruct (EQ' with "HE HL") as "(% & #Hown & #Hshr)".
+    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; iIntros "!# * H"].
     - iPureIntro. simpl. congruence.
     - destruct vl as [|[[]|]]; try done. iDestruct "H" as "[$ H]". by iApply "Hown".
     - iDestruct "H" as (a γ) "[Ha H]". iExists a, γ. iFrame.
-      iApply shr_bor_iff; last done.
-      iSplit; iIntros "!>!# H"; by iApply rwlock_inv_proper.
+      iApply shr_bor_iff; last done. 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).
diff --git a/theories/typing/lib/rwlock/rwlock_code.v b/theories/typing/lib/rwlock/rwlock_code.v
index 393e673f57fd86331bb5502481e1bfbbb7d9b7a0..1298abb58690d0cee78da37db4dd90009127f242 100644
--- a/theories/typing/lib/rwlock/rwlock_code.v
+++ b/theories/typing/lib/rwlock/rwlock_code.v
@@ -19,13 +19,13 @@ Section rwlock_functions.
       "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).
+  Lemma rwlock_new_type ty `{!TyWf ty} :
+    typed_val (rwlock_new ty) (fn(∅; ty) → rwlock ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ret ϝ arg). inv_vec arg=>x. simpl_subst.
     iApply type_new; [solve_typing..|].
-    iIntros (r tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    iIntros (r tid) "#LFT HE Hna HL Hk HT". simpl_subst.
     rewrite (Nat2Z.id (S ty.(ty_size))) tctx_interp_cons
             tctx_interp_singleton !tctx_hasty_val.
     iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done.
@@ -38,7 +38,7 @@ Section rwlock_functions.
     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)]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        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.
@@ -55,13 +55,13 @@ Section rwlock_functions.
       "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).
+  Lemma rwlock_into_inner_type ty `{!TyWf ty} :
+    typed_val (rwlock_into_inner ty) (fn(∅; rwlock ty) → ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (_ ret ϝ arg). inv_vec arg=>x. simpl_subst.
     iApply type_new; [solve_typing..|].
-    iIntros (r tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    iIntros (r tid) "#LFT HE Hna HL Hk HT". simpl_subst.
     rewrite (Nat2Z.id (ty.(ty_size))) tctx_interp_cons
             tctx_interp_singleton !tctx_hasty_val.
     iDestruct "HT" as "[Hr Hx]". destruct x as [[|lx|]|]; try done.
@@ -74,7 +74,7 @@ Section rwlock_functions.
     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]%TC
-        with "[] LFT Hna HE HL Hk [-]"); last first.
+        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.
@@ -89,13 +89,13 @@ Section rwlock_functions.
       "x" <- "x'" +â‚— #1;;
       return: ["x"].
 
-  Lemma rwlock_get_mut_type ty :
-    typed_val rwlock_get_mut (fn(∀ α, [α]; &uniq{α} (rwlock ty)) → &uniq{α} ty).
+  Lemma rwlock_get_mut_type ty `{!TyWf ty} :
+    typed_val rwlock_get_mut (fn(∀ α, ∅; &uniq{α} (rwlock ty)) → &uniq{α} ty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
-    iIntros (tid qE) "#LFT Hna HE HL HC HT".
+    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⌝) ∗
@@ -111,7 +111,7 @@ Section rwlock_functions.
     iDestruct "Hx" as (vl) "[Hx↦ Hx]". rewrite uninit_own. wp_op.
     iApply (type_type _ _ _
             [ #lx ◁ box (uninit 1); #(shift_loc lx' 1) ◁ &uniq{α}ty]%TC
-            with "[] LFT Hna HE HL HC [-]"); last first.
+            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..|].
@@ -139,45 +139,46 @@ Section rwlock_functions.
     cont: "k" [] :=
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma rwlock_try_read_type ty :
+  Lemma rwlock_try_read_type ty `{!TyWf ty} :
     typed_val rwlock_try_read
-        (fn(∀ α, [α]; &shr{α}(rwlock ty)) → option (rwlockreadguard α ty)).
+        (fn(∀ α, ∅; &shr{α}(rwlock ty)) → option (rwlockreadguard α ty)).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      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))])%TC);
+    iApply (type_cont [] [ϝ ⊑ []]%LL (λ _, [x ◁ box (&shr{α} rwlock ty);
+                                            r ◁ box (option (rwlockreadguard α ty))])%TC);
       [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 ◁ _])%TC);
+    iApply (type_cont [] [ϝ ⊑ []]%LL (λ _, [x ◁ _; x' ◁ _; r ◁ _])%TC);
       [iIntros (loop)|iIntros "/= !#"; iIntros (loop arg); inv_vec arg];
       simpl_subst.
     { iApply type_jump; solve_typing. }
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT".
+    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αβ Hinv]".
-    rewrite {1}/elctx_interp big_sepL_singleton.
-    iMod (lft_incl_acc with "Hαβ HE") as (qβ) "[[Hβtok1 Hβtok2] Hclose]". done.
+    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 (shr_bor_acc_tok with "LFT Hinv Hβtok1") as "[INV Hclose']"; try done.
+    iMod (shr_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.
+    iMod ("Hclose''" with "[Hlx INV]") as "Hβtok1"; first by iExists _; iFrame.
     iModIntro. wp_let. wp_op=>Hm1; wp_if.
-    - iApply (type_type _ _ _
+    - iMod ("Hclose'" with "[$]") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
+      iApply (type_type _ _ _
               [ x ◁ box (&shr{α}(rwlock ty)); r ◁ box (uninit 2) ]%TC
-              with "[] LFT Hna [> Hclose Hβtok1 Hβtok2] HL Hk"); first last.
+              with "[] LFT HE Hna HL Hk"); first last.
       { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
-      { rewrite {1}/elctx_interp big_sepL_singleton /=. iApply "Hclose". by iFrame. }
       iApply (type_sum_unit (option $ rwlockreadguard α ty));
         [solve_typing..|]; first last.
       simpl. iApply type_jump; solve_typing.
     - wp_op. wp_bind (CAS _ _ _).
-      iMod (shr_bor_acc_tok with "LFT Hinv Hβtok1") as "[INV Hclose']"; try done.
+      iMod (shr_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"); first done. iNext. iIntros "Hlx".
@@ -215,26 +216,24 @@ Section rwlock_functions.
             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".
+        iMod ("Hclose''" with "[$INV]") as "Hβtok1".
         iApply (wp_if _ true). iIntros "!>!>".
-        iMod ("Hclose" with "[$Hβtok1 $Hβtok2]") as "HE".
+        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]%TC
-              with "[] LFT Hna [> HE] HL Hk"); first last.
+              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. }
-        { rewrite {1}/elctx_interp big_sepL_singleton //. }
         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.
+        iMod ("Hclose''" with "[Hlx Hownst Hst]") as "Hβtok1"; first by iExists _; iFrame.
         iModIntro. iApply (wp_if _ false). iNext.
-        iMod ("Hclose" with "[$Hβtok1 $Hβtok2]") as "HE".
-        iSpecialize ("Hk" with "[HE] []").
-        { by rewrite {1}/elctx_interp big_sepL_singleton. } solve_typing.
+        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.
@@ -255,52 +254,52 @@ Section rwlock_functions.
     cont: "k" ["r"] :=
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma rwlock_try_write_type ty :
+  Lemma rwlock_try_write_type ty `{!TyWf ty} :
     typed_val rwlock_try_write
-        (fn(∀ α, [α]; &shr{α}(rwlock ty)) → option (rwlockwriteguard α ty)).
+        (fn(∀ α, ∅; &shr{α}(rwlock ty)) → option (rwlockwriteguard α ty)).
   Proof.
     intros. 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 ◁ box (option (rwlockwriteguard α ty))])%TC);
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
+    iApply (type_cont [_] [ϝ ⊑ []]%LL (λ r, [x ◁ box (&shr{α} rwlock ty);
+                                        r!!!0 ◁ box (option (rwlockwriteguard α ty))])%TC);
       [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 qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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αβ Hinv]".
-    rewrite {1}/elctx_interp big_sepL_singleton.
-    iMod (lft_incl_acc with "Hαβ HE") as (qβ) "[Hβtok Hclose]". done.
+    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 (shr_bor_acc_tok with "LFT Hinv Hβtok") as "(INV & Hclose')"; try done.
+    iMod (shr_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"). done. by destruct c as [|[[]]|].
       iNext. iIntros "Hlx".
-      iMod ("Hclose'" with "[Hlx Hownst Hst]") as "Hβtok"; first by iExists _; iFrame.
+      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. iApply (wp_if _ false). iNext.
       iApply (type_type _ _ _
               [ x ◁ box (&shr{α}(rwlock ty)); r ◁ box (uninit 2) ]%TC
-              with "[] LFT Hna [> Hclose Hβtok] HL Hk"); first last.
+              with "[] LFT HE Hna HL Hk"); first last.
       { rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val. iFrame. }
-      { rewrite {1}/elctx_interp big_sepL_singleton /=. iApply "Hclose". by 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"). done. 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. iApply (wp_if _ true). iNext. iMod ("Hclose" with "Hβtok") as "HE".
+      iMod ("Hclose''" with "[Hlx Hownst]") as "Hβtok". { iExists _. iFrame. }
+      iModIntro. iApply (wp_if _ true). iNext. 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]%TC
-              with "[] LFT Hna [> HE] HL Hk"); first last.
+              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 "∗#". }
-      { rewrite {1}/elctx_interp big_sepL_singleton //. }
       iApply (type_sum_assign (option $ rwlockwriteguard α ty)); [solve_typing..|].
       simpl. iApply type_jump; solve_typing.
   Qed.
diff --git a/theories/typing/lib/rwlock/rwlockreadguard.v b/theories/typing/lib/rwlock/rwlockreadguard.v
index 6388002fa6c586dd4d1362ab65f707e6ac094855..3f44a88e8d599be379c2eae4b57633218ad58169 100644
--- a/theories/typing/lib/rwlock/rwlockreadguard.v
+++ b/theories/typing/lib/rwlock/rwlockreadguard.v
@@ -61,6 +61,9 @@ Section rwlockreadguard.
     iApply lft_incl_refl.
   Qed.
 
+  Global Instance rwlockreadguard_wf α ty `{!TyWf ty} : TyWf (rwlockreadguard α ty) :=
+    { ty_lfts := [α]; ty_wf_E := ty.(ty_wf_E) ++ ty.(ty_outlives_E) α }.
+
   Global Instance rwlockreadguard_type_contractive α : TypeContractive (rwlockreadguard α).
   Proof.
     constructor;
@@ -75,19 +78,21 @@ Section rwlockreadguard.
   Global Instance rwlockreadguard_mono E L :
     Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockreadguard.
   Proof.
-    iIntros (α1 α2 Hα ty1 ty2 Hty) "#HE #HL".
-    iDestruct (proj1 Hty with "HE HL") as "(%&#Ho&#Hs)".
-    iDestruct (Hα with "HE HL") as "Hα1α2".
-    iSplit; [|iSplit; iAlways].
+    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; [|iSplit; iAlways].
     - done.
     - iIntros (tid [|[[]|][]]) "H"; try done.
-      iDestruct "H" as (ν q γ β) "(#Hshr & #H⊑ & #Hinv & Htok & Hown)".
-      iExists ν, q, γ, β. iFrame "∗#". iSplit; last iSplit.
+      iDestruct "H" as (ν q' γ β) "(#Hshr & #H⊑ & #Hinv & Htok & Hown)".
+      iExists ν, q', γ, β. 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 (shr_bor_iff with "[] Hinv").
-        iIntros "!> !#"; iSplit; iIntros "H"; by iApply rwlock_inv_proper.
+        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.
diff --git a/theories/typing/lib/rwlock/rwlockreadguard_code.v b/theories/typing/lib/rwlock/rwlockreadguard_code.v
index 08917456fbbd5185116ea1eb49b95748b0bd1332..e53896fcc304360e0cc85467fabea7d22c4903b9 100644
--- a/theories/typing/lib/rwlock/rwlockreadguard_code.v
+++ b/theories/typing/lib/rwlock/rwlockreadguard_code.v
@@ -18,30 +18,28 @@ Section rwlockreadguard_functions.
       letalloc: "r" <- "r'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma rwlockreadguard_deref_type ty :
+  Lemma rwlockreadguard_deref_type ty `{!TyWf ty} :
     typed_val rwlockreadguard_deref
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL
-                              [# &shr{α}(rwlockreadguard β ty)] (&shr{α}ty))).
+      (fn (fun '(α, β) => FP_wf ∅ [# &shr{α}(rwlockreadguard β ty)] (&shr{α}ty))).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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]".
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "(Hα & Hβ & #Hαβ)".
+    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↦]") 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);
                               #(shift_loc l' 1) ◁ &shr{α}ty]%TC
-      with "[] LFT Hna [Hα Hβ Hαβ] HL Hk"); first last.
+      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. done.
-      iApply lft_incl_refl. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
+      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.
@@ -60,24 +58,22 @@ Section rwlockreadguard_functions.
         let: "r" := new [ #0] in return: ["r"]
       else "loop" [].
 
-  Lemma rwlockreadguard_drop_type ty :
-    typed_val rwlockreadguard_drop
-              (fn(∀ α, [α]; rwlockreadguard α ty) → unit).
+  Lemma rwlockreadguard_drop_type ty `{!TyWf ty} :
+    typed_val rwlockreadguard_drop (fn(∀ α, ∅; rwlockreadguard α ty) → unit).
   Proof.
     intros. 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' ◁ _])%TC);
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
+    iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
+    iApply (type_cont [] [ϝ ⊑ []]%LL (λ _, [x ◁ _; x' ◁ _])%TC);
       [iIntros (loop)|iIntros "/= !#"; iIntros (loop arg); inv_vec arg];
       simpl_subst.
     { iApply type_jump; solve_typing. }
-    iIntros (tid qE) "#LFT Hna Hα HL Hk HT".
-    rewrite {1}/elctx_interp big_sepL_singleton tctx_interp_cons
-            tctx_interp_singleton !tctx_hasty_val.
+    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 γ β) "(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 (shr_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|].
@@ -126,22 +122,22 @@ Section rwlockreadguard_functions.
       iApply (wp_step_fupd with "INV"). done. 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 "HE". iApply (wp_if _ true). iIntros "!>!>".
+      iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
+      iApply (wp_if _ true). iIntros "!>!>".
       iApply (type_type _ _ _ [ x ◁ box (uninit 1)]%TC
-              with "[] LFT Hna [HE] HL Hk"); first last.
+              with "[] LFT HE Hna HL Hk"); first last.
       { rewrite tctx_interp_singleton tctx_hasty_val //. }
-      { rewrite /elctx_interp big_sepL_singleton //. }
       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 "HE". iApply (wp_if _ false). iIntros "!> !>".
+      iMod ("Hcloseα" with "Hβ") as "Hα". iMod ("Hclose" with "Hα HL") as "HL".
+      iApply (wp_if _ false). iIntros "!> !>".
       iApply (type_type _ _ _ [ x ◁ box (uninit 1); #lx' ◁ rwlockreadguard α ty]%TC
-              with "[] LFT Hna [HE] HL Hk"); first last.
+              with "[] LFT HE Hna HL Hk"); first last.
       { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val
-                tctx_hasty_val' //. iFrame. iExists _, _, _, _. iFrame "∗#". }
-      { rewrite /elctx_interp big_sepL_singleton //. }
+                tctx_hasty_val' //=. 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
index ca13f922e59c1437c7ba83a0107e2f9aee1b82b1..7003099dff1d60b3a7d5dc2842d6706f5bb52a11 100644
--- a/theories/typing/lib/rwlock/rwlockwriteguard.v
+++ b/theories/typing/lib/rwlock/rwlockwriteguard.v
@@ -61,6 +61,9 @@ Section rwlockwriteguard.
       iApply ty_shr_mono; try done. iApply lft_intersect_mono. iApply lft_incl_refl. done.
   Qed.
 
+  Global Instance rwlockwriteguard_wf α ty `{!TyWf ty} : TyWf (rwlockwriteguard α ty) :=
+    { ty_lfts := [α]; ty_wf_E := ty.(ty_wf_E) ++ ty.(ty_outlives_E) α }.
+
   Global Instance rwlockwriteguard_type_contractive α : TypeContractive (rwlockwriteguard α).
   Proof.
     constructor;
@@ -73,11 +76,12 @@ Section rwlockwriteguard.
   Global Instance rwlockwriteguard_mono E L :
     Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) rwlockwriteguard.
   Proof.
-    iIntros (α1 α2 Hα ty1 ty2 Hty) "#HE #HL".
-    pose proof Hty as Hty'%eqtype_unfold.
-    iDestruct (Hty' with "HE HL") as "(%&#Ho&#Hs)".
-    iDestruct (Hα with "HE HL") as "Hα1α2".
-    iSplit; [|iSplit; iAlways].
+    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; [|iSplit; iAlways].
     - done.
     - iIntros (tid [|[[]|][]]) "H"; try done.
       iDestruct "H" as (γ β) "(Hb & #H⊑ & #Hinv & Hown)".
@@ -87,7 +91,7 @@ Section rwlockwriteguard.
         iExists vl; iFrame; by iApply "Ho".
       + by iApply lft_incl_trans.
       + iApply shr_bor_iff; try done.
-        iIntros "!>!#"; iSplit; iIntros "H"; by iApply rwlock_inv_proper.
+        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]". set_solver.
diff --git a/theories/typing/lib/rwlock/rwlockwriteguard_code.v b/theories/typing/lib/rwlock/rwlockwriteguard_code.v
index d46cf2071e319bcf75d814de68dd9b14e0d5d478..7ccd5321e11c58d0e4a57a66caae4b64a2c2e5d7 100644
--- a/theories/typing/lib/rwlock/rwlockwriteguard_code.v
+++ b/theories/typing/lib/rwlock/rwlockwriteguard_code.v
@@ -18,35 +18,37 @@ Section rwlockwriteguard_functions.
       letalloc: "r" <- "r'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma rwlockwriteguard_deref_type ty :
+  Lemma rwlockwriteguard_deref_type ty `{!TyWf ty} :
     typed_val rwlockwriteguard_deref
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL
-                              [# &shr{α}(rwlockwriteguard β ty)] (&shr{α}ty))).
+      (fn (fun '(α, β) => FP_wf ∅ [# &shr{α}(rwlockwriteguard β ty)] (&shr{α}ty))).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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]".
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "([Hα1 Hα2] & Hβ & #Hαβ)".
+    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β]");
          [done| |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α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);
                               #(shift_loc l' 1) ◁ &shr{α}ty]%TC
-            with "[] LFT Hna [Hα1 Hα2 Hβ Hαβ] HL Hk"); last first.
+            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.
-      iApply lft_incl_refl. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
+        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.
@@ -60,21 +62,19 @@ Section rwlockwriteguard_functions.
       letalloc: "r" <- "r'" in
       delete [ #1; "x"];; return: ["r"].
 
-  Lemma rwlockwriteguard_derefmut_type ty :
+  Lemma rwlockwriteguard_derefmut_type ty `{!TyWf ty} :
     typed_val rwlockwriteguard_derefmut
-      (fn (fun '(α, β) => FP [α; β; α ⊑ β]%EL
-                              [# &uniq{α}(rwlockwriteguard β ty)] (&uniq{α}ty))).
+      (fn (fun '(α, β) => FP_wf ∅ [# &uniq{α}(rwlockwriteguard β ty)] (&uniq{α}ty))).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros ([α β] ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros ([α β] ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x').
-    iIntros (tid qE) "#LFT Hna HE HL Hk HT". simpl_subst.
+    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']".
-    rewrite {1}/elctx_interp 2!big_sepL_cons big_sepL_singleton.
-    iDestruct "HE" as "(Hα & Hβ & #Hαβ)". destruct x' as [[|lx'|]|]; try done.
+    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_tok with "LFT H Hα") as "[>[] _]".
     rewrite heap_mapsto_vec_singleton.
@@ -87,13 +87,13 @@ Section rwlockwriteguard_functions.
     wp_bind (!(LitV lx'))%E. iApply (wp_step_fupd with "[Hb]");
       [done| |by iApply (bor_unnest with "LFT Hb")|]; first done.
     wp_read. iIntros "Hb !>". wp_op. wp_let.
-    iMod ("Hcloseα" with "[$H↦]") as "[_ Hα]".
+    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); #(shift_loc l 1) ◁ &uniq{α}ty]%TC
-            with "[] LFT Hna [Hα Hβ Hαβ] HL Hk"); last first.
+            with "[] LFT HE Hna HL Hk"); last first.
     { rewrite tctx_interp_cons tctx_interp_singleton tctx_hasty_val tctx_hasty_val' //.
-      iFrame. iApply (bor_shorten with "[] Hb").
-      by iApply lft_incl_glb; [iApply lft_incl_trans|iApply lft_incl_refl]. }
-    { rewrite /elctx_interp 2!big_sepL_cons big_sepL_singleton. by iFrame. }
+      iFrame. 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.
@@ -107,19 +107,20 @@ Section rwlockwriteguard_functions.
       delete [ #1; "x"];;
       let: "r" := new [ #0] in return: ["r"].
 
-  Lemma rwlockwriteguard_drop_type ty :
+  Lemma rwlockwriteguard_drop_type ty `{!TyWf ty} :
     typed_val rwlockwriteguard_drop
-              (fn(∀ α, [α]; rwlockwriteguard α ty) → unit).
+              (fn(∀ α, ∅; rwlockwriteguard α ty) → unit).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (α ret arg). inv_vec arg=>x. simpl_subst.
+      iIntros (α ϝ ret arg). inv_vec arg=>x. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (x'). simpl_subst.
-    iIntros (tid qE) "#LFT Hna Hα HL Hk HT".
-    rewrite {1}/elctx_interp big_sepL_singleton tctx_interp_cons
-            tctx_interp_singleton !tctx_hasty_val.
+    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 (γ β) "(Hx' & #Hαβ & #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 (shr_bor_acc_tok with "LFT Hinv Hβ") as "[INV Hcloseβ]"; [done..|].
@@ -132,11 +133,11 @@ Section rwlockwriteguard_functions.
       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_empty, _. }
-    iModIntro. wp_seq. iMod ("Hcloseα" with "Hβ") as "HE".
+    iModIntro. wp_seq. iMod ("Hcloseα" with "Hβ") as "Hα".
+    iMod ("Hclose" with "Hα HL") as "HL".
     iApply (type_type _ _ _ [ x ◁ box (uninit 1)]%TC
-            with "[] LFT Hna [HE] HL Hk"); first last.
+            with "[] LFT HE Hna HL Hk"); first last.
     { rewrite tctx_interp_singleton tctx_hasty_val //. }
-    { rewrite /elctx_interp big_sepL_singleton //. }
     iApply type_delete; [solve_typing..|].
     iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
     iApply type_jump; solve_typing.
diff --git a/theories/typing/lib/spawn.v b/theories/typing/lib/spawn.v
index f65bb4995635e9e14ba75df26ea1e12f65d2765e..d2a67fab880097b13687536aaad9ab54238592c6 100644
--- a/theories/typing/lib/spawn.v
+++ b/theories/typing/lib/spawn.v
@@ -26,6 +26,9 @@ Section join_handle.
   Next Obligation. iIntros "* _ _ _ $". auto. Qed.
   Next Obligation. iIntros (?) "**"; auto. Qed.
 
+  Global Instance join_handle_wf ty `{!TyWf ty} : TyWf (join_handle ty) :=
+    { ty_lfts := ty.(ty_lfts); ty_wf_E := ty.(ty_wf_E) }.
+
   Lemma join_handle_subtype ty1 ty2 :
     ▷ type_incl ty1 ty2 -∗ type_incl (join_handle ty1) (join_handle ty2).
   Proof.
@@ -39,8 +42,8 @@ Section join_handle.
   Global Instance join_handle_mono E L :
     Proper (subtype E L ==> subtype E L) join_handle.
   Proof.
-    iIntros (ty1 ty2 Hsub) "#? #?". iApply join_handle_subtype.
-    iApply Hsub; done.
+    iIntros (ty1 ty2 Hsub ?) "HL". iDestruct (Hsub with "HL") as "#Hsub".
+    iIntros "!# #HE". iApply join_handle_subtype. iApply "Hsub"; done.
   Qed.
   Global Instance join_handle_proper E L :
     Proper (eqtype E L ==> eqtype E L) join_handle.
@@ -67,16 +70,18 @@ Section spawn.
       letalloc: "r" <- "join" in
       delete [ #1; "f"];; return: ["r"].
 
-  Lemma spawn_type `(!Send envty, !Send retty) :
-    typed_val spawn (fn([]; fn([] ; envty) → retty, envty) → join_handle retty).
+  Lemma spawn_type envty retty `{!TyWf envty, !TyWf retty}
+        `(!Send envty, !Send retty) :
+    let E ϝ := envty.(ty_outlives_E) static ++ retty.(ty_outlives_E) static in
+    typed_val spawn (fn(E; fn(∅; envty) → retty, envty) → join_handle retty).
   Proof. (* FIXME: typed_instruction_ty is not used for printing. *)
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>f env. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>f env. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (f'). simpl_subst.
     iApply (type_let _ _ _ _ ([f' ◁ _; env ◁ _]%TC)
                      (λ j, [j ◁ join_handle retty]%TC)); try solve_typing; [|].
     { (* The core of the proof: showing that spawn is safe. *)
-      iIntros (tid qE) "#LFT $ $ $".
+      iIntros (tid) "#LFT #HE $ $".
       rewrite tctx_interp_cons tctx_interp_singleton.
       iIntros "[Hf' Henv]". iApply (spawn_spec _ (join_inv tid retty) with "[-]");
                               first solve_to_val; last first; last simpl_subst.
@@ -84,22 +89,22 @@ Section spawn.
         iIntros "?". iExists retty. iFrame. iApply type_incl_refl. }
       iIntros (c) "Hfin". iMod na_alloc as (tid') "Htl". wp_let. wp_let.
       (* FIXME this is horrible. *)
-      assert (Hcall := type_call' [] [] [] f' [env:expr] (λ _:(), FP [] [# envty] retty)).
+      refine (let Hcall := type_call' _ [] [] f' [] [env:expr]
+                (λ _:(), FP_wf ∅ [# envty] retty) in _).
       specialize (Hcall (rec: "_k" ["r"] := finish [ #c; "r"])%V ()).
       erewrite of_val_unlock in Hcall; last done.
-      iApply (Hcall $! _ 1%Qp with "LFT Htl [] [] [Hfin]").
-      - apply elctx_sat_nil.
-      - rewrite /elctx_interp big_sepL_nil. done.
+      iApply (Hcall with "LFT HE Htl [] [Hfin]").
+      - constructor.
+      - solve_typing.
       - rewrite /llctx_interp big_sepL_nil. done.
-      - rewrite /cctx_interp. iIntros "_ * Hin".
+      - rewrite /cctx_interp. iIntros "* Hin".
         iDestruct "Hin" as %Hin%elem_of_list_singleton. subst x.
         rewrite /cctx_elt_interp. iIntros "* ?? Hret". inv_vec args=>arg /=.
         wp_rec. iApply (finish_spec with "[$Hfin Hret]"); last auto.
         rewrite tctx_interp_singleton tctx_hasty_val. by iApply @send_change_tid.
       - rewrite tctx_interp_cons tctx_interp_singleton. iSplitL "Hf'".
-        + rewrite !tctx_hasty_val.
-          iApply @send_change_tid. done.
-        + rewrite !tctx_hasty_val. iApply @send_change_tid. done. }
+        + rewrite !tctx_hasty_val. by iApply @send_change_tid.
+        + rewrite !tctx_hasty_val. by iApply @send_change_tid. }
     iIntros (v). simpl_subst.
     iApply type_new; [solve_typing..|]. iIntros (r). simpl_subst.
     iApply type_assign; [solve_typing..|].
@@ -113,15 +118,15 @@ Section spawn.
       let: "r" := join ["c'"] in
       delete [ #1; "c"];; return: ["r"].
 
-  Lemma join_type retty :
-    typed_val join (fn([]; join_handle retty) → retty).
+  Lemma join_type retty `{!TyWf retty} :
+    typed_val join (fn(∅; join_handle retty) → retty).
   Proof.
     intros. iApply type_fn; [solve_typing..|]. iIntros "/= !#".
-      iIntros (_ ret arg). inv_vec arg=>c. simpl_subst.
+      iIntros (_ ϝ ret arg). inv_vec arg=>c. simpl_subst.
     iApply type_deref; [solve_typing..|]. iIntros (c'); simpl_subst.
     iApply (type_let _ _ _ _ ([c' ◁ _]%TC)
                              (λ r, [r ◁ box retty]%TC)); try solve_typing; [|].
-    { iIntros (tid qE) "#LFT $ $ $".
+    { iIntros (tid) "#LFT _ $ $".
       rewrite tctx_interp_singleton tctx_hasty_val. iIntros "Hc'".
       destruct c' as [[|c'|]|]; try done. iDestruct "Hc'" as (ty') "[#Hsub Hc']".
       iApply (join_spec with "Hc'"). iNext. iIntros "* Hret".
diff --git a/theories/typing/own.v b/theories/typing/own.v
index f5458177e1f6b450659092b14b6107f5fc2bbb0b..1fee0f1222317aff9c295c4ef89583859ca1de1e 100644
--- a/theories/typing/own.v
+++ b/theories/typing/own.v
@@ -63,8 +63,8 @@ Section own.
          end%I;
        ty_shr κ tid l :=
          (∃ l':loc, &frac{κ}(λ q', l ↦{q'} #l') ∗
-            □ (∀ F q, ⌜↑shrN ∪ lftE ⊆ F⌝ -∗ q.[κ] ={F,F∖↑shrN}▷=∗ ty.(ty_shr) κ tid l' ∗ q.[κ]))%I
-    |}.
+            □ (∀ F q, ⌜↑shrN ∪ lftE ⊆ F⌝ -∗ q.[κ] ={F,F∖↑shrN}▷=∗
+                            ty.(ty_shr) κ tid l' ∗ q.[κ]))%I |}.
   Next Obligation. by iIntros (q ty tid [|[[]|][]]) "H". Qed.
   Next Obligation.
     move=>n ty N κ l tid ?? /=. iIntros "#LFT Hshr Htok".
@@ -88,14 +88,17 @@ Section own.
     by iApply (ty.(ty_shr_mono) with "Hκ").
   Qed.
 
+  Global Instance own_ptr_wf n ty `{!TyWf ty} : TyWf (own_ptr n ty) :=
+    { ty_lfts := ty.(ty_lfts); ty_wf_E := ty.(ty_wf_E) }.
+
   Lemma own_type_incl n m ty1 ty2 :
     ▷ ⌜n = m⌝ -∗ ▷ type_incl ty1 ty2 -∗ type_incl (own_ptr n ty1) (own_ptr m ty2).
   Proof.
     iIntros "#Heq (#Hsz & #Ho & #Hs)". iSplit; first done. iSplit; iAlways.
     - iIntros (?[|[[| |]|][]]) "H"; try done. simpl.
-      iDestruct "H" as "[Hmt H†]". iDestruct "Hmt" as (vl') "[Hmt Hown]". iNext.
-      iDestruct ("Ho" with "Hown") as "Hown". iDestruct ("Hsz") as %<-.
-      iDestruct "Heq" as %->. iFrame. iExists _. iFrame.
+      iDestruct "H" as "[Hmt H†]". iNext. iDestruct ("Hsz") as %<-.
+      iDestruct "Heq" as %->. iFrame. iApply (heap_mapsto_pred_wand with "Hmt").
+      iApply "Ho".
     - iIntros (???) "H". iDestruct "H" as (l') "[Hfb #Hvs]".
       iExists l'. iFrame. iIntros "!#". iIntros (F' q) "% Htok".
       iMod ("Hvs" with "[%] Htok") as "Hvs'". done. iModIntro. iNext.
@@ -105,8 +108,10 @@ Section own.
   Global Instance own_mono E L n :
     Proper (subtype E L ==> subtype E L) (own_ptr n).
   Proof.
-    intros ty1 ty2 Hincl. iIntros.
-    iApply own_type_incl; first by auto. iApply Hincl; auto.
+    intros ty1 ty2 Hincl. iIntros (qL) "HL".
+    iDestruct (Hincl with "HL") as "#Hincl".
+    iClear "∗". iIntros "!# #HE".
+    iApply own_type_incl; first by auto. iApply "Hincl"; auto.
   Qed.
   Lemma own_mono' E L n1 n2 ty1 ty2 :
     n1 = n2 → subtype E L ty1 ty2 → subtype E L (own_ptr n1 ty1) (own_ptr n2 ty2).
@@ -157,7 +162,10 @@ Section box.
   Global Instance box_mono E L :
     Proper (subtype E L ==> subtype E L) box.
   Proof.
-    intros ty1 ty2 Hincl. iIntros. iApply box_type_incl. iApply Hincl; auto.
+    intros ty1 ty2 Hincl. iIntros (qL) "HL".
+    iDestruct (Hincl with "HL") as "#Hincl".
+    iClear "∗". iIntros "!# #HE".
+    iApply box_type_incl. iApply "Hincl"; auto.
   Qed.
   Lemma box_mono' E L ty1 ty2 :
     subtype E L ty1 ty2 → subtype E L (box ty1) (box ty2).
@@ -181,7 +189,8 @@ Section util.
 
   Lemma ownptr_own n ty tid v :
     (own_ptr n ty).(ty_own) tid [v] ⊣⊢
-       ∃ (l : loc) (vl : vec val ty.(ty_size)), ⌜v = #l⌝ ∗ ▷ l ↦∗ vl ∗ ▷ ty.(ty_own) tid vl ∗ ▷ freeable_sz n ty.(ty_size) l.
+       ∃ (l : loc) (vl : vec val ty.(ty_size)),
+         ⌜v = #l⌝ ∗ ▷ l ↦∗ vl ∗ ▷ ty.(ty_own) tid vl ∗ ▷ freeable_sz n ty.(ty_size) l.
   Proof.
     iSplit.
     - iIntros "Hown". destruct v as [[|l|]|]; try done.
@@ -213,7 +222,7 @@ Section typing.
   Lemma write_own {E L} ty ty' n :
     ty.(ty_size) = ty'.(ty_size) → typed_write E L (own_ptr n ty') ty (own_ptr n ty).
   Proof.
-    iIntros (Hsz) "!#". iIntros ([[]|] tid F qE qL ?) "_ $ $ Hown"; try done.
+    iIntros (Hsz) "!#". iIntros ([[]|] tid F qL ?) "_ _ $ Hown"; try done.
     rewrite /= Hsz. iDestruct "Hown" as "[H↦ $]". iDestruct "H↦" as (vl) "[>H↦ Hown]".
     iDestruct (ty_size_eq with "Hown") as "#>%". iExists _, _. iFrame "H↦". auto.
   Qed.
@@ -221,7 +230,7 @@ Section typing.
   Lemma read_own_copy E L ty n :
     Copy ty → typed_read E L (own_ptr n ty) ty (own_ptr n ty).
   Proof.
-    iIntros (Hsz) "!#". iIntros ([[]|] tid F qE qL ?) "_ $ $ $ Hown"; try done.
+    iIntros (Hsz) "!#". iIntros ([[]|] tid F qL ?) "_ _ $ $ Hown"; try done.
     iDestruct "Hown" as "[H↦ H†]". iDestruct "H↦" as (vl) "[>H↦ #Hown]".
     iExists l, _, _. iFrame "∗#". iSplitR; first done. iIntros "!> Hl !>".
     iExists _. auto.
@@ -230,7 +239,7 @@ Section typing.
   Lemma read_own_move E L ty n :
     typed_read E L (own_ptr n ty) ty (own_ptr n $ uninit ty.(ty_size)).
   Proof.
-    iAlways. iIntros ([[]|] tid F qE qL ?) "_ $ $ $ Hown"; try done.
+    iAlways. iIntros ([[]|] tid F qL ?) "_ _ $ $ Hown"; try done.
     iDestruct "Hown" as "[H↦ H†]". iDestruct "H↦" as (vl) "[>H↦ Hown]".
     iDestruct (ty_size_eq with "Hown") as "#>%".
     iExists l, vl, _. iFrame "∗#". iSplitR; first done. iIntros "!> Hl !> !>".
@@ -242,7 +251,7 @@ Section typing.
     let n' := Z.to_nat n in
     typed_instruction_ty E L [] (new [ #n ]%E) (own_ptr n' (uninit n')).
   Proof.
-    iIntros (? tid q) "#LFT $ $ $ _".
+    iIntros (? tid) "#LFT #HE $ $ _".
     iApply wp_new; first done. iModIntro.
     iIntros (l vl) "(% & H† & Hlft)". subst. rewrite tctx_interp_singleton tctx_hasty_val.
     iNext. rewrite freeable_sz_full Z2Nat.id //. iFrame.
@@ -274,7 +283,7 @@ Section typing.
     Z.of_nat (ty.(ty_size)) = n →
     typed_instruction E L [p ◁ own_ptr ty.(ty_size) ty] (delete [ #n; p])%E (λ _, []).
   Proof.
-    iIntros (<- tid eq) "#LFT $ $ $ Hp". rewrite tctx_interp_singleton.
+    iIntros (<- tid) "#LFT #HE $ $ Hp". rewrite tctx_interp_singleton.
     wp_bind p. iApply (wp_hasty with "Hp"). iIntros ([[]|]) "_ Hown"; try done.
     iDestruct "Hown" as "[H↦: >H†]". iDestruct "H↦:" as (vl) "[>H↦ Hown]".
     rewrite tctx_interp_nil. iDestruct (ty_size_eq with "Hown") as "#>EQ".
diff --git a/theories/typing/product.v b/theories/typing/product.v
index 5a27d299b491583561519c03a94f85a89aa73daa..a51abc5358006379c1dc9bda92e4c8c9d0e4edc9 100644
--- a/theories/typing/product.v
+++ b/theories/typing/product.v
@@ -81,9 +81,11 @@ Section product.
   Global Instance product2_mono E L:
     Proper (subtype E L ==> subtype E L ==> subtype E L) product2.
   Proof.
-    iIntros (ty11 ty12 H1 ty21 ty22 H2). iIntros.
-    iDestruct (H1 with "[] []") as "(% & #Ho1 & #Hs1)"; [done..|]. clear H1.
-    iDestruct (H2 with "[] []") as "(% & #Ho2 & #Hs2)"; [done..|]. clear H2.
+    iIntros (ty11 ty12 H1 ty21 ty22 H2). iIntros (qL) "HL".
+    iDestruct (H1 with "HL") as "#H1". iDestruct (H2 with "HL") as "#H2".
+    iClear "∗". iIntros "!# #HE".
+    iDestruct ("H1" with "HE") as "#(% & #Ho1 & #Hs1)". clear H1.
+    iDestruct ("H2" with "HE") as "#(% & #Ho2 & #Hs2)". clear H2.
     iSplit; first by (iPureIntro; simpl; f_equal). iSplit; iAlways.
     - iIntros (??) "H". iDestruct "H" as (vl1 vl2) "(% & Hown1 & Hown2)".
       iExists _, _. iSplit. done. iSplitL "Hown1".
@@ -146,6 +148,9 @@ Section product.
   Definition product := foldr product2 unit0.
   Definition unit := product [].
 
+  Global Instance product_wf tyl `{!TyWfLst tyl} : TyWf (product tyl) :=
+    { ty_lfts := tyl.(tyl_lfts); ty_wf_E := tyl.(tyl_wf_E) }.
+
   Global Instance product_type_ne n: Proper (Forall2 (type_dist2 n) ==> type_dist2 n) product.
   Proof. intros ??. induction 1=>//=. by f_equiv. Qed.
   Global Instance product_ne : NonExpansive product.
@@ -182,39 +187,39 @@ Section typing.
 
   Global Instance prod2_assoc E L : Assoc (eqtype E L) product2.
   Proof.
-    split; (iIntros; (iSplit; first by rewrite /= assoc); iSplit; iAlways;
-            last iIntros (?); iIntros (??) "H").
+    intros ???. apply eqtype_unfold. iIntros (?) "_ !# _".
+    iSplit; first by rewrite /= assoc. iSplit; iIntros "!# *"; iSplit; iIntros "H".
     - iDestruct "H" as (vl1 vl') "(% & Ho1 & H)".
       iDestruct "H" as (vl2 vl3) "(% & Ho2 & Ho3)". subst.
       iExists _, _. iSplit. by rewrite assoc. iFrame. iExists _, _. by iFrame.
-    - iDestruct "H" as "(Hs1 & Hs2 & Hs3)". rewrite shift_loc_assoc_nat.
-      by iFrame.
     - iDestruct "H" as (vl1 vl') "(% & H & Ho3)".
       iDestruct "H" as (vl2 vl3) "(% & Ho1 & Ho2)". subst.
       iExists _, _. iSplit. by rewrite -assoc. iFrame. iExists _, _. by iFrame.
+    - iDestruct "H" as "(Hs1 & Hs2 & Hs3)". rewrite shift_loc_assoc_nat.
+      by iFrame.
     - iDestruct "H" as "[[Hs1 Hs2] Hs3]". rewrite /= shift_loc_assoc_nat.
       by iFrame.
   Qed.
 
   Global Instance prod2_unit_leftid E L : LeftId (eqtype E L) unit product2.
   Proof.
-    intros ty. split; (iIntros; (iSplit; first done); iSplit; iAlways;
-                  last iIntros (?); iIntros (??) "H").
+    intros ty. apply eqtype_unfold. iIntros (?) "_ !# _". iSplit; first done.
+    iSplit; iIntros "!# *"; iSplit; iIntros "H".
     - iDestruct "H" as (? ?) "(% & % & ?)". by subst.
+    - iExists [], _. eauto.
     - iDestruct "H" as "(? & ?)". rewrite shift_loc_0.
       iApply ty_shr_mono; last done.
       iApply lft_incl_refl.
-    - iExists [], _. eauto.
     - simpl. rewrite shift_loc_0. by iFrame.
   Qed.
 
   Global Instance prod2_unit_rightid E L : RightId (eqtype E L) unit product2.
   Proof.
-    intros ty. split; (iIntros; (iSplit; first by rewrite /= -plus_n_O); iSplit;
-                  iAlways; last iIntros (?); iIntros (??) "H").
+    intros ty. apply eqtype_unfold. iIntros (?) "_ !# _".
+    iSplit; first by rewrite /= -plus_n_O. iSplit; iIntros "!# *"; iSplit; iIntros "H".
     - iDestruct "H" as (? ?) "(% & ? & %)". subst. by rewrite app_nil_r.
-    - iDestruct "H" as "(? & _)". done.
     - iExists _, []. rewrite app_nil_r. eauto.
+    - iDestruct "H" as "(? & _)". done.
     - simpl. iFrame.
   Qed.
 
diff --git a/theories/typing/product_split.v b/theories/typing/product_split.v
index 47a6b6fa778ba783b97b5dfb891c3b27941a344d..68d42dbad3d694132601a388e0ab41a0d1108fbc 100644
--- a/theories/typing/product_split.v
+++ b/theories/typing/product_split.v
@@ -37,16 +37,16 @@ Section product_split.
     (∀ tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝) →
     ∀ p, tctx_incl E L [p ◁ ptr $ product tyl] (hasty_ptr_offsets p ptr tyl 0).
   Proof.
-    iIntros (Hsplit Hloc p tid q1 q2) "#LFT HE HL H".
+    iIntros (Hsplit Hloc p tid q) "#LFT #HE HL H".
     iInduction tyl as [|ty tyl IH] "IH" forall (p).
-    { iFrame "HE HL". rewrite tctx_interp_nil. done. }
-    rewrite product_cons. iMod (Hsplit with "LFT HE HL H") as "(HE & HL & H)".
+    { rewrite tctx_interp_nil. auto. }
+    rewrite product_cons. iMod (Hsplit with "LFT HE HL H") as "(HL & H)".
     cbn -[tctx_elt_interp]. rewrite tctx_interp_cons tctx_interp_singleton tctx_interp_cons.
     iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]". iDestruct "Hp" as %Hp. 
     iDestruct (Hloc with "Hty") as %[l [=->]].
     iAssert (tctx_elt_interp tid (p +ₗ #0 ◁ ptr ty)) with "[Hty]" as "$".
     { iExists #l. iSplit; last done. simpl; by rewrite Hp shift_loc_0. }
-    iMod ("IH" with "HE HL [Htyl]") as "($ & $ & Htyl)".
+    iMod ("IH" with "HL [Htyl]") as "($ & Htyl)".
     { rewrite tctx_interp_singleton //. }
     iClear "IH". rewrite (hasty_ptr_offsets_offset l) // -plus_n_O //.
   Qed.
@@ -59,7 +59,7 @@ Section product_split.
     (∀ tid ty vl, (ptr ty).(ty_own) tid vl -∗ ⌜∃ l : loc, vl = [(#l) : val]⌝) →
     ∀ p, tctx_incl E L (hasty_ptr_offsets p ptr tyl 0) [p ◁ ptr $ product tyl].
   Proof.
-    iIntros (Hptr Htyl Hmerge Hloc p tid q1 q2) "#LFT HE HL H".
+    iIntros (Hptr Htyl Hmerge Hloc p tid q) "#LFT #HE HL H".
     iInduction tyl as [|ty tyl IH] "IH" forall (p Htyl); first done.
     rewrite product_cons. rewrite /= tctx_interp_singleton tctx_interp_cons.
     iDestruct "H" as "[Hty Htyl]". iDestruct "Hty" as (v) "[Hp Hty]".
@@ -68,16 +68,15 @@ Section product_split.
     { move:Hp. simpl. clear. destruct (eval_path p); last done.
       destruct v; try done. destruct l0; try done. rewrite shift_loc_0. done. }
     clear Hp. destruct tyl.
-    { iDestruct (elctx_interp_persist with "HE") as "#HE'".
-      iDestruct (llctx_interp_persist with "HL") as "#HL'". iFrame "HE HL". 
-      iClear "IH Htyl". iExists #l. rewrite product_nil. iSplitR; first done.
-      assert (eqtype E L (ptr ty) (ptr (product2 ty unit))) as [Hincl _].
+    { assert (eqtype E L (ptr ty) (ptr (product2 ty unit))) as [Hincl _].
       { rewrite right_id. done. }
-      iDestruct (Hincl with "HE' HL'") as "#(_ & #Heq & _)". by iApply "Heq". }
-    iMod ("IH" with "[] HE HL [Htyl]") as "(HE & HL & Htyl)"; first done.
+      iDestruct (Hincl with "HL HE") as "#(_ & #Heq & _)".
+      iFrame. iClear "IH Htyl". iExists #l. rewrite product_nil. iSplitR; first done.
+      by iApply "Heq". }
+    iMod ("IH" with "[] HL [Htyl]") as "(HL & Htyl)"; first done.
     { change (ty_size ty) with (0+ty_size ty)%nat at 1.
       rewrite plus_comm -hasty_ptr_offsets_offset //. }
-    iClear "IH". iMod (Hmerge with "LFT HE HL [Hty Htyl]") as "($ & $ & ?)";
+    iClear "IH". iMod (Hmerge with "LFT HE HL [Hty Htyl]") as "($ & ?)";
                    last by rewrite tctx_interp_singleton.
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton. iFrame.
     iExists #l. iSplit; done.
@@ -88,7 +87,7 @@ Section product_split.
     tctx_incl E L [p ◁ own_ptr n $ product2 ty1 ty2]
                   [p ◁ own_ptr n ty1; p +ₗ #ty1.(ty_size) ◁ own_ptr n ty2].
   Proof.
-    iIntros (tid q1 q2) "#LFT $ $ H".
+    iIntros (tid q) "#LFT _ $ H".
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as ([[]|]) "[#Hp H]"; try done.
     iDestruct "H" as "[H >H†]". iDestruct "H" as (vl) "[>H↦ H]".
@@ -106,7 +105,7 @@ Section product_split.
     tctx_incl E L [p ◁ own_ptr n ty1; p +ₗ #ty1.(ty_size) ◁ own_ptr n ty2]
                   [p ◁ own_ptr n $ product2 ty1 ty2].
   Proof.
-    iIntros (tid q1 q2) "#LFT $ $ H".
+    iIntros (tid q) "#LFT _ $ H".
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[]|]) "(Hp1 & H1)"; try done.
     iDestruct "H1" as "(H↦1 & H†1)".
@@ -143,7 +142,7 @@ Section product_split.
     tctx_incl E L [p ◁ &uniq{κ} product2 ty1 ty2]
                   [p ◁ &uniq{κ} ty1; p +ₗ #ty1.(ty_size) ◁ &uniq{κ} ty2].
   Proof.
-    iIntros (tid q1 q2) "#LFT $ $ H".
+    iIntros (tid q) "#LFT _ $ H".
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as ([[]|]) "[Hp H]"; try done. iDestruct "Hp" as %Hp.
     rewrite /= split_prod_mt. iMod (bor_sep with "LFT H") as "[H1 H2]".
@@ -155,7 +154,7 @@ Section product_split.
     tctx_incl E L [p ◁ &uniq{κ} ty1; p +ₗ #ty1.(ty_size) ◁ &uniq{κ} ty2]
                   [p ◁ &uniq{κ} product2 ty1 ty2].
   Proof.
-    iIntros (tid q1 q2) "#LFT $ $ H".
+    iIntros (tid q) "#LFT _ $ H".
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[]|]) "[Hp1 H1]"; try done.
     iDestruct "Hp1" as %Hp1. iDestruct "H2" as (v2) "(Hp2 & H2)". rewrite /= Hp1.
@@ -192,7 +191,7 @@ Section product_split.
     tctx_incl E L [p ◁ &shr{κ} product2 ty1 ty2]
                   [p ◁ &shr{κ} ty1; p +ₗ #ty1.(ty_size) ◁ &shr{κ} ty2].
   Proof.
-    iIntros (tid q1 q2) "#LFT $ $ H".
+    iIntros (tid q) "#LFT _ $ H".
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as ([[]|]) "[Hp H]"; try iDestruct "H" as "[]".
     iDestruct "H" as "[H1 H2]". iDestruct "Hp" as %Hp.
@@ -203,7 +202,7 @@ Section product_split.
     tctx_incl E L [p ◁ &shr{κ} ty1; p +ₗ #ty1.(ty_size) ◁ &shr{κ} ty2]
                   [p ◁ &shr{κ} product2 ty1 ty2].
   Proof.
-    iIntros (tid q1 q2) "#LFT $ $ H".
+    iIntros (tid q) "#LFT _ $ H".
     rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as "[H1 H2]". iDestruct "H1" as ([[]|]) "[Hp1 Hown1]"; try done.
     iDestruct "Hp1" as %Hp1. iDestruct "H2" as ([[]|]) "[Hp2 Hown2]"; try done.
diff --git a/theories/typing/programs.v b/theories/typing/programs.v
index 72f5e56bc221503c83d022b51f19bd462c943c46..7f545608a5bbf6e91d971c98789512bd1ad479a4 100644
--- a/theories/typing/programs.v
+++ b/theories/typing/programs.v
@@ -10,9 +10,8 @@ Section typing.
   (* This is an iProp because it is also used by the function type. *)
   Definition typed_body (E : elctx) (L : llctx) (C : cctx) (T : tctx)
                         (e : expr) : iProp Σ :=
-    (∀ tid qE, lft_ctx -∗ na_own tid ⊤ -∗
-               elctx_interp E qE -∗ llctx_interp L 1 -∗
-               (elctx_interp E qE -∗ cctx_interp tid C) -∗ tctx_interp tid T -∗
+    (∀ tid, lft_ctx -∗ elctx_interp E -∗ na_own tid ⊤ -∗ llctx_interp L 1 -∗
+               cctx_interp tid C -∗ tctx_interp tid T -∗
                WP e {{ _, cont_postcondition }})%I.
   Global Arguments typed_body _%EL _%LL _%CC _%TC _%E.
 
@@ -35,10 +34,10 @@ Section typing.
            (typed_body E L).
   Proof.
     intros C1 C2 HC T1 T2 HT e ? <-. iIntros "H".
-    iIntros (tid qE) "#LFT Htl HE HL HC HT".
-    iMod (HT with "LFT HE HL HT") as "(HE & HL & HT)".
-    iApply ("H" with "LFT Htl HE HL [HC] HT").
-    iIntros "HE". by iApply (HC with "LFT HC").
+    iIntros (tid) "#LFT #HE Htl HL HC HT".
+    iMod (HT with "LFT HE HL HT") as "(HL & HT)".
+    iApply ("H" with "LFT HE Htl HL [HC] HT").
+    by iApply (HC with "LFT HE HC").
   Qed.
 
   Global Instance typed_body_mono_flip E L:
@@ -49,19 +48,19 @@ Section typing.
   (** Instruction *)
   Definition typed_instruction (E : elctx) (L : llctx)
              (T1 : tctx) (e : expr) (T2 : val → tctx) : iProp Σ :=
-    (∀ tid qE, lft_ctx -∗ na_own tid ⊤ -∗
-              elctx_interp E qE -∗ llctx_interp L 1 -∗ tctx_interp tid T1 -∗
-              WP e {{ v, na_own tid ⊤ ∗ elctx_interp E qE ∗
+    (∀ tid, lft_ctx -∗ elctx_interp E -∗ na_own tid ⊤ -∗
+              llctx_interp L 1 -∗ tctx_interp tid T1 -∗
+              WP e {{ v, na_own tid ⊤ ∗
                          llctx_interp L 1 ∗ tctx_interp tid (T2 v) }})%I.
   Global Arguments typed_instruction _%EL _%LL _%TC _%E _%TC.
 
   (** Writing and Reading **)
   Definition typed_write (E : elctx) (L : llctx) (ty1 ty ty2 : type) : iProp Σ :=
-    (□ ∀ v tid F qE qL, ⌜lftE ∪ (↑lrustN) ⊆ F⌝ →
-      lft_ctx -∗ elctx_interp E qE -∗ llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗
+    (□ ∀ v tid F qL, ⌜lftE ∪ (↑lrustN) ⊆ F⌝ →
+      lft_ctx -∗ elctx_interp E -∗ llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗
         ∃ (l : loc) vl, ⌜length vl = ty.(ty_size) ∧ v = #l⌝ ∗ l ↦∗ vl ∗
           (▷ l ↦∗: ty.(ty_own) tid ={F}=∗
-            elctx_interp E qE ∗ llctx_interp L qL ∗ ty2.(ty_own) tid [v]))%I.
+            llctx_interp L qL ∗ ty2.(ty_own) tid [v]))%I.
   Global Arguments typed_write _%EL _%LL _%T _%T _%T.
 
   (* Technically speaking, we could remvoe the vl quantifiaction here and use
@@ -70,11 +69,11 @@ Section typing.
      that nobody could possibly have changed the vl (because only half the
      fraction was given). So we go with the definition that is easier to prove. *)
   Definition typed_read (E : elctx) (L : llctx) (ty1 ty ty2 : type) : iProp Σ :=
-    (□ ∀ v tid F qE qL, ⌜lftE ∪ (↑lrustN) ⊆ F⌝ →
-      lft_ctx -∗ na_own tid F -∗
-      elctx_interp E qE -∗ llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗
+    (□ ∀ v tid F qL, ⌜lftE ∪ (↑lrustN) ⊆ F⌝ →
+      lft_ctx -∗ elctx_interp E -∗ na_own tid F -∗
+      llctx_interp L qL -∗ ty1.(ty_own) tid [v] ={F}=∗
         ∃ (l : loc) vl q, ⌜v = #l⌝ ∗ l ↦∗{q} vl ∗ ▷ ty.(ty_own) tid vl ∗
-              (l ↦∗{q} vl ={F}=∗ na_own tid F ∗ elctx_interp E qE ∗
+              (l ↦∗{q} vl ={F}=∗ na_own tid F ∗
                               llctx_interp L qL ∗ ty2.(ty_own) tid [v]))%I.
   Global Arguments typed_read _%EL _%LL _%T _%T _%T.
 End typing.
@@ -100,13 +99,11 @@ Section typing_rules.
     typed_body ((κ1 ⊑ κ2) :: (κ2 ⊑ κ1) :: E) L C T e →
     typed_body E ((κ1 ⊑ [κ2]%list) :: L) C T e.
   Proof.
-    iIntros (He tid qE) "#LFT Htl HE HL HC HT".
+    iIntros (He tid) "#LFT #HE Htl HL HC HT".
     rewrite /llctx_interp big_sepL_cons. iDestruct "HL" as "[Hκ HL]".
     iMod (lctx_equalize_lft with "LFT Hκ") as "[Hκ1 Hκ2]".
-    iApply (He with "LFT Htl [Hκ1 Hκ2 HE] HL [HC] HT").
-    - rewrite /elctx_interp !big_sepL_cons. by iFrame.
-    - rewrite /elctx_interp !big_sepL_cons. iIntros "(_ & _ & HE)".
-      iApply "HC". done.
+    iApply (He with "LFT [Hκ1 Hκ2 HE] Htl HL HC HT").
+    rewrite /elctx_interp !big_sepL_cons. by iFrame.
   Qed.
 
   Lemma type_let' E L T1 T2 (T : tctx) C xb e e' :
@@ -115,11 +112,11 @@ Section typing_rules.
     (∀ v : val, typed_body E L C (T2 v ++ T) (subst' xb v e')) -∗
     typed_body E L C (T1 ++ T) (let: xb := e in e').
   Proof.
-    iIntros (Hc) "He He'". iIntros (tid qE) "#LFT Htl HE HL HC HT". rewrite tctx_interp_app.
-    iDestruct "HT" as "[HT1 HT]". wp_bind e. iApply (wp_wand with "[He HE HL HT1 Htl]").
-    { iApply ("He" with "LFT Htl HE HL HT1"). }
-    iIntros (v) "/= (Htl & HE & HL & HT2)". iApply wp_let; first wp_done.
-    iModIntro. iApply ("He'" with "LFT Htl HE HL HC [HT2 HT]").
+    iIntros (Hc) "He He'". iIntros (tid) "#LFT #HE Htl HL HC HT". rewrite tctx_interp_app.
+    iDestruct "HT" as "[HT1 HT]". wp_bind e. iApply (wp_wand with "[He HL HT1 Htl]").
+    { iApply ("He" with "LFT HE Htl HL HT1"). }
+    iIntros (v) "/= (Htl & HL & HT2)". iApply wp_let; first wp_done.
+    iModIntro. iApply ("He'" with "LFT HE Htl HL HC [HT2 HT]").
     rewrite tctx_interp_app. by iFrame.
   Qed.
 
@@ -152,10 +149,10 @@ Section typing_rules.
     (∀ κ, typed_body E ((κ ⊑ κs) :: L) C T e) -∗
     typed_body E L C T (Newlft ;; e).
   Proof.
-    iIntros (Hc) "He". iIntros (tid qE) "#LFT Htl HE HL HC HT".
+    iIntros (Hc) "He". iIntros (tid) "#LFT #HE Htl HL HC HT".
     iMod (lft_create with "LFT") as (Λ) "[Htk #Hinh]"; first done.
-    set (κ' := foldr lft_intersect static κs). wp_seq.
-    iApply ("He" $! (κ' ⊓ Λ) with "LFT Htl HE [HL Htk] HC HT").
+    set (κ' := lft_intersect_list κs). wp_seq.
+    iApply ("He" $! (κ' ⊓ Λ) with "LFT HE Htl [HL Htk] HC HT").
     rewrite /llctx_interp big_sepL_cons. iFrame "HL".
     iExists Λ. iSplit; first done. iFrame. done.
   Qed.
@@ -166,20 +163,20 @@ Section typing_rules.
     Closed [] e → UnblockTctx κ T1 T2 →
     typed_body E L C T2 e -∗ typed_body E ((κ ⊑ κs) :: L) C T1 (Endlft ;; e).
   Proof.
-    iIntros (Hc Hub) "He". iIntros (tid qE) "#LFT Htl HE".
+    iIntros (Hc Hub) "He". iIntros (tid) "#LFT #HE Htl".
     rewrite /llctx_interp big_sepL_cons. iIntros "[Hκ HL] HC HT".
     iDestruct "Hκ" as (Λ) "(% & Htok & #Hend)".
     iSpecialize ("Hend" with "Htok"). wp_bind Endlft.
     iApply (wp_mask_mono (↑lftN)); first done.
     iApply (wp_step_fupd with "Hend"). done. set_solver. wp_seq.
-    iIntros "#Hdead !>". wp_seq. iApply ("He" with "LFT Htl HE HL HC [> -]").
+    iIntros "#Hdead !>". wp_seq. iApply ("He" with "LFT HE Htl HL HC [> -]").
     iApply (Hub with "[] HT"). simpl in *. subst κ. rewrite -lft_dead_or. auto.
   Qed.
 
   Lemma type_path_instr {E L} p ty :
     typed_instruction_ty E L [p ◁ ty] p ty.
   Proof.
-    iIntros (??) "_ $$$ ?". rewrite tctx_interp_singleton.
+    iIntros (?) "_ _ $$ ?". rewrite tctx_interp_singleton.
     iApply (wp_hasty with "[-]"). done. iIntros (v) "_ Hv".
     rewrite tctx_interp_singleton. iExists v. iFrame. by rewrite eval_path_of_val.
   Qed.
@@ -195,7 +192,7 @@ Section typing_rules.
     typed_write E L ty1 ty ty1' →
     typed_instruction E L [p1 ◁ ty1; p2 ◁ ty] (p1 <- p2) (λ _, [p1 ◁ ty1']%TC).
   Proof.
-    iIntros (Hwrt tid qE) "#LFT $ HE HL".
+    iIntros (Hwrt tid) "#LFT #HE $ HL".
     rewrite tctx_interp_cons tctx_interp_singleton. iIntros "[Hp1 Hp2]".
     wp_bind p1. iApply (wp_hasty with "Hp1"). iIntros (v1) "% Hown1".
     wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2) "_ Hown2".
@@ -204,7 +201,7 @@ Section typing_rules.
     rewrite <-Hsz in *. destruct vl as [|v[|]]; try done.
     rewrite heap_mapsto_vec_singleton. iApply wp_fupd. wp_write.
     rewrite -heap_mapsto_vec_singleton.
-    iMod ("Hclose" with "[Hl Hown2]") as "($ & $ & Hown)".
+    iMod ("Hclose" with "[Hl Hown2]") as "($ & Hown)".
     { iExists _. iFrame. }
     rewrite tctx_interp_singleton tctx_hasty_val' //.
   Qed.
@@ -221,10 +218,10 @@ Section typing_rules.
     ty.(ty_size) = 1%nat → typed_read E L ty1 ty ty1' →
     typed_instruction E L [p ◁ ty1] (!p) (λ v, [p ◁ ty1'; v ◁ ty]%TC).
   Proof.
-    iIntros (Hsz Hread tid qE) "#LFT Htl HE HL Hp".
+    iIntros (Hsz Hread tid) "#LFT #HE Htl HL Hp".
     rewrite tctx_interp_singleton. wp_bind p. iApply (wp_hasty with "Hp").
     iIntros (v) "% Hown".
-    iMod (Hread with "[] LFT Htl HE HL Hown") as
+    iMod (Hread with "[] LFT HE Htl HL Hown") as
         (l vl q) "(% & Hl & Hown & Hclose)"; first done.
     subst v. iDestruct (ty_size_eq with "Hown") as "#>%". rewrite ->Hsz in *.
     destruct vl as [|v [|]]; try done.
@@ -243,29 +240,29 @@ Section typing_rules.
     typed_body E L C T (let: x := !p in e).
   Proof. iIntros. by iApply type_let; [apply type_deref_instr|solve_typing|]. Qed.
 
-  Lemma type_memcpy_iris E qE L qL tid ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 :
+  Lemma type_memcpy_iris E L qL tid ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 :
     Z.of_nat (ty.(ty_size)) = n →
     typed_write E L ty1 ty ty1' -∗ typed_read E L ty2 ty ty2' -∗
-    {{{ lft_ctx ∗ na_own tid ⊤ ∗ elctx_interp E qE ∗ llctx_interp L qL ∗
+    {{{ lft_ctx ∗ elctx_interp E ∗ na_own tid ⊤ ∗ llctx_interp L qL ∗
         tctx_elt_interp tid (p1 ◁ ty1) ∗ tctx_elt_interp tid (p2 ◁ ty2) }}}
       (p1 <-{n} !p2)
-    {{{ RET #(); na_own tid ⊤ ∗ elctx_interp E qE ∗ llctx_interp L qL ∗
+    {{{ RET #(); na_own tid ⊤ ∗ llctx_interp L qL ∗
                  tctx_elt_interp tid (p1 ◁ ty1') ∗ tctx_elt_interp tid (p2 ◁ ty2') }}}.
   Proof.
     iIntros (<-) "#Hwrt #Hread !#".
-    iIntros (Φ) "(#LFT & Htl & [HE1 HE2] & [HL1 HL2] & [Hp1 Hp2]) HΦ".
+    iIntros (Φ) "(#LFT & #HE & Htl & [HL1 HL2] & [Hp1 Hp2]) HΦ".
     wp_bind p1. iApply (wp_hasty with "Hp1"). iIntros (v1) "% Hown1".
     wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2) "% Hown2".
-    iMod ("Hwrt" with "[] LFT HE1 HL1 Hown1")
+    iMod ("Hwrt" with "[] LFT HE HL1 Hown1")
       as (l1 vl1) "([% %] & Hl1 & Hcl1)"; first done.
-    iMod ("Hread" with "[] LFT Htl HE2 HL2 Hown2")
+    iMod ("Hread" with "[] LFT HE Htl HL2 Hown2")
       as (l2 vl2 q2) "(% & Hl2 & Hown2 & Hcl2)"; first done.
     iDestruct (ty_size_eq with "Hown2") as "#>%". subst v1 v2. iApply wp_fupd.
     iApply (wp_memcpy with "[$Hl1 $Hl2]"); try congruence; [].
     iNext. iIntros "[Hl1 Hl2]". iApply ("HΦ" with "[> -]"). rewrite !tctx_hasty_val' //.
-    iMod ("Hcl1" with "[Hl1 Hown2]") as "($ & $ & $)".
+    iMod ("Hcl1" with "[Hl1 Hown2]") as "($ & $)".
     { iExists _. iFrame. }
-    iMod ("Hcl2" with "Hl2") as "($ & $ & $ & $)". done.
+    iMod ("Hcl2" with "Hl2") as "($ & $ & $)". done.
   Qed.
 
   Lemma type_memcpy_instr {E L} ty ty1 ty1' ty2 ty2' (n : Z) p1 p2 :
@@ -274,7 +271,7 @@ Section typing_rules.
     typed_instruction E L [p1 ◁ ty1; p2 ◁ ty2] (p1 <-{n} !p2)
                       (λ _, [p1 ◁ ty1'; p2 ◁ ty2']%TC).
   Proof.
-    iIntros (Hsz Hwrt Hread tid qE) "#LFT Htl HE HL HT".
+    iIntros (Hsz Hwrt Hread tid) "#LFT #HE Htl HL HT".
     iApply (type_memcpy_iris with "[] [] [$LFT $Htl $HE $HL HT]"); try done.
     { iApply Hwrt. }
     { iApply Hread. }
diff --git a/theories/typing/shr_bor.v b/theories/typing/shr_bor.v
index 2a1abdb433aafba85473faec86506b1732603b6c..a7cad9b8270abd495b23d11cae3809b9d2d945e6 100644
--- a/theories/typing/shr_bor.v
+++ b/theories/typing/shr_bor.v
@@ -16,13 +16,18 @@ Section shr_bor.
   Next Obligation. by iIntros (κ ty tid [|[[]|][]]) "H". Qed.
   Next Obligation. intros κ ty tid [|[[]|][]]; apply _. Qed.
 
+  Global Instance shr_bor_wf κ ty `{!TyWf ty} : TyWf (shr_bor κ ty) :=
+    { ty_lfts := [κ]; ty_wf_E := ty.(ty_wf_E) ++ ty.(ty_outlives_E) κ }.
+
   Global Instance shr_mono E L :
     Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) shr_bor.
   Proof.
     intros κ1 κ2 Hκ ty1 ty2 Hty. apply subtype_simple_type.
-    iIntros (? [|[[]|][]]) "#HE #HL H"; try done.
-    iDestruct (Hκ with "HE HL") as "#Hκ". iApply (ty2.(ty_shr_mono) with "Hκ").
-    iDestruct (Hty with "[] []") as "(_ & _ & #Hs1)"; [done..|clear Hty].
+    iIntros (?) "/= HL". iDestruct (Hκ with "HL") as "#Hincl".
+    iDestruct (Hty with "HL") as "#Hty". iIntros "!# #HE".
+    iIntros (? [|[[]|][]]) "H"; try done. iDestruct ("Hincl" with "HE") as "#Hκ".
+    iApply (ty2.(ty_shr_mono) with "Hκ").
+    iDestruct ("Hty" with "HE") as "(_ & _ & #Hs1)"; [done..|clear Hty].
     by iApply "Hs1".
   Qed.
   Global Instance shr_mono_flip E L :
@@ -64,11 +69,8 @@ Section typing.
     lctx_lft_incl E L κ' κ →
     tctx_incl E L [p ◁ &shr{κ}ty] [p ◁ &shr{κ'}ty; p ◁{κ} &shr{κ}ty].
   Proof.
-    iIntros (Hκκ' tid ??) "#LFT HE HL H".
-    iDestruct (elctx_interp_persist with "HE") as "#HE'".
-    iDestruct (llctx_interp_persist with "HL") as "#HL'". iFrame "HE HL".
-    iDestruct (Hκκ' with "HE' HL'") as "Hκκ'".
-    rewrite /tctx_interp big_sepL_singleton big_sepL_cons big_sepL_singleton.
+    iIntros (Hκκ' tid ?) "#LFT #HE HL H". iDestruct (Hκκ' with "HL HE") as "#Hκκ'".
+    iFrame. rewrite /tctx_interp big_sepL_singleton big_sepL_cons big_sepL_singleton.
     iDestruct "H" as ([[]|]) "[% #Hshr]"; try done. iModIntro. iSplit.
     - iExists _. iSplit. done. by iApply (ty_shr_mono with "Hκκ' Hshr").
     - iExists _. auto.
@@ -78,7 +80,7 @@ Section typing.
     Copy ty → lctx_lft_alive E L κ → typed_read E L (&shr{κ}ty) ty (&shr{κ}ty).
   Proof.
     iIntros (Hcopy Halive) "!#".
-    iIntros ([[]|] tid F qE qL ?) "#LFT Htl HE HL #Hshr"; try done.
+    iIntros ([[]|] tid F qL ?) "#LFT #HE Htl HL #Hshr"; try done.
     iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first set_solver.
     assert (↑shrN ⊆ (↑lrustN : coPset)) by solve_ndisj. (* set_solver needs some help. *)
     iMod (copy_shr_acc with "LFT Hshr Htl Hκ") as (q') "(Htl & H↦ & Hcl)".
diff --git a/theories/typing/soundness.v b/theories/typing/soundness.v
index eda8bb99162e6104b8c10f925b58bc00e81e45cc..42ba0dbbbb99af88b59b4ac1f5b0bcf5c72a705a 100644
--- a/theories/typing/soundness.v
+++ b/theories/typing/soundness.v
@@ -23,7 +23,7 @@ Proof. solve_inG. Qed.
 Section type_soundness.
   Definition exit_cont : val := λ: [<>], #().
 
-  Definition main_type `{typeG Σ} := (fn([]) → unit)%T.
+  Definition main_type `{typeG Σ} := (fn(∅) → unit)%T.
 
   Theorem type_soundness `{typePreG Σ} (main : val) σ t :
     (∀ `{typeG Σ}, typed_val main main_type) →
@@ -39,20 +39,23 @@ Section type_soundness.
     iMod lft_init as (?) "#LFT". done.
     iMod na_alloc as (tid) "Htl". set (Htype := TypeG _ _ _ _ _).
     wp_bind (of_val main). iApply (wp_wand with "[Htl]").
-    iApply (Hmain _ [] [] $! tid 1%Qp with "LFT Htl [] [] []").
+    iApply (Hmain _ [] [] $! tid with "LFT [] Htl [] []").
     { by rewrite /elctx_interp big_sepL_nil. }
     { by rewrite /llctx_interp big_sepL_nil. }
     { by rewrite tctx_interp_nil. }
-    clear Hrtc Hmain main. iIntros (main) "(Htl & HE & HL & Hmain)".
-    rewrite tctx_interp_singleton. iDestruct "Hmain" as (?) "[EQ Hmain]".
+    clear Hrtc Hmain main. iIntros (main) "(Htl & _ & Hmain & _)".
+    iDestruct "Hmain" as (?) "[EQ Hmain]".
     rewrite eval_path_of_val. iDestruct "EQ" as %[= <-].
     iDestruct "Hmain" as (f k x e ?) "(EQ & % & Hmain)". iDestruct "EQ" as %[= ->].
     destruct x; try done.
     iApply (wp_rec _ _ _ _ _ _ [exit_cont]%E); [done| |by simpl_subst|iNext].
     { repeat econstructor. apply to_of_val. }
-    iApply ("Hmain" $! () exit_cont [#] tid 1%Qp with "LFT Htl HE HL []");
+    iMod (lft_create with "LFT") as (ϝ) "Hϝ". done.
+    iApply ("Hmain" $! () ϝ exit_cont [#] tid with "LFT [] Htl [Hϝ] []");
       last by rewrite tctx_interp_nil.
-    iIntros "_". rewrite cctx_interp_singleton. simpl. iIntros (args) "_ _".
+    { by rewrite /elctx_interp /= big_sepL_nil. }
+    { rewrite /llctx_interp big_sepL_singleton. iExists ϝ. iFrame. by rewrite /= left_id. }
+    rewrite cctx_interp_singleton. simpl. iIntros (args) "_ _".
     inv_vec args. iIntros (x) "_ /=". by wp_lam.
   Qed.
 End type_soundness.
diff --git a/theories/typing/sum.v b/theories/typing/sum.v
index f2a92fd62d149d3179b2b8be38d5310e68006bae..450483311d2b6fea979b2e5da38beebea515a2a2 100644
--- a/theories/typing/sum.v
+++ b/theories/typing/sum.v
@@ -1,3 +1,4 @@
+From iris.base_logic Require Import big_op.
 From iris.proofmode Require Import tactics.
 From iris.algebra Require Import list.
 From iris.base_logic Require Import fractional.
@@ -19,6 +20,8 @@ Section sum.
   Qed.
   Next Obligation. iIntros (κ κ' tid l) "#Hord []". Qed.
 
+  Global Instance emp_wf : TyWf emp := { ty_lfts := []; ty_wf_E := [] }.
+
   Global Instance emp_empty : Empty type := emp.
 
   Global Instance emp_copy : Copy ∅.
@@ -91,6 +94,9 @@ Section sum.
     - iApply ((nth i tyl ∅).(ty_shr_mono) with "Hord"); done.
   Qed.
 
+  Global Instance sum_wf tyl `{!TyWfLst tyl} : TyWf (sum tyl) :=
+    { ty_lfts := tyl.(tyl_lfts); ty_wf_E := tyl.(tyl_wf_E) }.
+
   Global Instance sum_type_ne n : Proper (Forall2 (type_dist2 n) ==> type_dist2 n) sum.
   Proof.
     intros x y EQ.
@@ -130,19 +136,27 @@ Section sum.
   Global Instance sum_mono E L :
     Proper (Forall2 (subtype E L) ==> subtype E L) sum.
   Proof.
-    iIntros (tyl1 tyl2 Htyl) "#? %".
-    iAssert (⌜max_list_with ty_size tyl1 = max_list_with ty_size tyl2⌝%I) with "[#]" as %Hleq.
-    { iInduction Htyl as [|???? Hsub] "IH"; first done.
-      iDestruct (Hsub with "[] []") as "(% & _ & _)"; [done..|].
-      iDestruct "IH" as %IH. iPureIntro. simpl. inversion_clear IH. by f_equal. }
+    iIntros (tyl1 tyl2 Htyl qL) "HL".
+    iAssert (□ (lft_contexts.elctx_interp E -∗ ⌜max_list_with ty_size tyl1 = max_list_with ty_size tyl2⌝))%I with "[#]" as "#Hleq".
+    { iInduction Htyl as [|???? Hsub] "IH".
+      { iClear "∗". iIntros "!# _". done. }
+      iDestruct (Hsub with "HL") as "#Hsub". iDestruct ("IH" with "HL") as "{IH} #IH".
+      (* FIXME: WHy does the previous iDestruvt remove HL? *)
+      iAlways. iIntros "#HE". iDestruct ("Hsub" with "HE") as "(% & _ & _)".
+      iDestruct ("IH" with "HE") as %IH. iPureIntro. simpl. inversion_clear IH. by f_equal. }
+    iDestruct (subtype_Forall2_llctx with "HL") as "#Htyl"; first done.
+    iClear "HL". iIntros "!# #HE".
+    iDestruct ("Hleq" with "HE") as %Hleq. iSpecialize ("Htyl" with "HE"). iClear "Hleq HE".
     iAssert (∀ i, type_incl (nth i tyl1 ∅) (nth i tyl2 ∅))%I with "[#]" as "#Hty".
       { iIntros (i). edestruct (nth_lookup_or_length tyl1 i) as [Hl1|Hl]; last first.
         { rewrite nth_overflow // nth_overflow; first by iApply type_incl_refl.
           by erewrite <-Forall2_length. }
-        edestruct @Forall2_lookup_l as (ty2 & Hl2 & Hty2); [done..|].
-        rewrite (nth_lookup_Some tyl2 _ _ ty2) //.
-        by iApply (Hty2 with "[] []"). }
-    clear -Hleq. iSplit; last iSplit.
+        edestruct @Forall2_lookup_l as (ty2 & Hl2 & _); [done..|].
+        iDestruct (big_sepL_lookup with "Htyl") as "Hty".
+        { rewrite lookup_zip_with. erewrite Hl1. simpl.
+          rewrite Hl2 /=. done. }
+        rewrite (nth_lookup_Some tyl2 _ _ ty2) //. }
+    clear -Hleq. iClear "∗". iSplit; last iSplit.
     - simpl. by rewrite Hleq.
     - iAlways. iIntros (tid vl) "H". iDestruct "H" as (i vl' vl'') "(% & % & Hown)".
       iExists i, vl', vl''. iSplit; first done.
@@ -174,12 +188,10 @@ Section sum.
   Lemma emp_sum E L :
     eqtype E L emp (sum []).
   Proof.
-    split; (iIntros; iSplit; first done; iSplit; iAlways).
-    - iIntros (??) "[]".
-    - iIntros (κ tid l) "[]".
-    - iIntros (??) "H". iDestruct "H" as (i vl' vl'') "(% & % & Hown)".
-      by rewrite nth_empty.
-    - iIntros (???) "H". iDestruct "H" as (i) "(_ & Hshr)". by rewrite nth_empty.
+    apply eqtype_unfold. iIntros (?) "_ !# _".
+    iSplit; first done; iSplit; iAlways; iIntros; iSplit; try by iIntros "[]".
+    - iIntros "H". iDestruct "H" as (i vl' vl'') "(% & % & Hown)". by rewrite nth_empty.
+    - iIntros "H". iDestruct "H" as (i) "(_ & Hshr)". by rewrite nth_empty.
   Qed.
 
   Global Instance sum_copy tyl : LstCopy tyl → Copy (sum tyl).
diff --git a/theories/typing/type.v b/theories/typing/type.v
index bb4e356375148a25b07d814ee8b65be0f697067d..401392ae02e3b6a3ce1bd5938e165ecf789a0e56 100644
--- a/theories/typing/type.v
+++ b/theories/typing/type.v
@@ -53,6 +53,61 @@ Instance: Params (@ty_shr) 2.
 
 Arguments ty_own {_ _} _ _ !_ /.
 
+Class TyWf `{typeG Σ} (ty : type) := { ty_lfts : list lft; ty_wf_E : elctx }.
+Arguments ty_lfts {_ _} _ {_}.
+Arguments ty_wf_E {_ _} _ {_}.
+
+Definition ty_outlives_E `{typeG Σ} ty `{!TyWf ty} (κ : lft) : elctx :=
+  (λ α, κ ⊑ α)%EL <$> ty.(ty_lfts).
+
+Lemma ty_outlives_E_elctx_sat `{typeG Σ} E L ty `{!TyWf ty} α β :
+  ty.(ty_outlives_E) β ⊆+ E →
+  lctx_lft_incl E L α β →
+  elctx_sat E L (ty.(ty_outlives_E) α).
+Proof.
+  unfold ty_outlives_E. induction ty.(ty_lfts) as [|κ l IH]=>/= Hsub Hαβ.
+  - solve_typing.
+  - apply elctx_sat_lft_incl.
+    + etrans; first done. eapply lctx_lft_incl_external, elem_of_submseteq, Hsub.
+      set_solver.
+    + apply IH, Hαβ. etrans; last done. by apply submseteq_cons.
+Qed.
+
+Inductive TyWfLst `{typeG Σ} : list type → Type :=
+| tyl_wf_nil : TyWfLst []
+| tyl_wf_cons ty tyl `{!TyWf ty, !TyWfLst tyl} : TyWfLst (ty::tyl).
+Existing Class TyWfLst.
+Existing Instances tyl_wf_nil tyl_wf_cons.
+
+Fixpoint tyl_lfts `{typeG Σ} tyl {WF : TyWfLst tyl} : list lft :=
+  match WF with
+  | tyl_wf_nil => []
+  | tyl_wf_cons ty tyl _ _ => ty.(ty_lfts) ++ tyl.(tyl_lfts)
+  end.
+
+Fixpoint tyl_wf_E `{typeG Σ} tyl {WF : TyWfLst tyl} : elctx :=
+  match WF with
+  | tyl_wf_nil => []
+  | tyl_wf_cons ty tyl _ _ => ty.(ty_wf_E) ++ tyl.(tyl_wf_E)
+  end.
+
+Fixpoint tyl_outlives_E `{typeG Σ} tyl {WF : TyWfLst tyl} (κ : lft) : elctx :=
+  match WF with
+  | tyl_wf_nil => []
+  | tyl_wf_cons ty tyl _ _ => ty.(ty_outlives_E) κ ++ tyl.(tyl_outlives_E) κ
+  end.
+
+Lemma tyl_outlives_E_elctx_sat `{typeG Σ} E L tyl {WF : TyWfLst tyl} α β :
+  tyl.(tyl_outlives_E) β ⊆+ E →
+  lctx_lft_incl E L α β →
+  elctx_sat E L (tyl.(tyl_outlives_E) α).
+Proof.
+  induction WF as [|???? IH]=>/=.
+  - solve_typing.
+  - intros. apply elctx_sat_app, IH; [eapply ty_outlives_E_elctx_sat| |]=>//;
+      (etrans; [|done]); solve_typing.
+Qed.
+
 Record simple_type `{typeG Σ} :=
   { st_own : thread_id → list val → iProp Σ;
     st_size_eq tid vl : st_own tid vl -∗ ⌜length vl = 1%nat⌝;
@@ -194,8 +249,7 @@ Section ofe.
 
   Global Instance ty_of_st_ne : NonExpansive ty_of_st.
   Proof.
-    intros n ?? EQ. constructor. done. simpl.
-    - intros tid l. apply EQ.
+    intros n ?? EQ. constructor; try apply EQ. done.
     - simpl. intros; repeat f_equiv. apply EQ.
   Qed.
   Global Instance ty_of_st_proper : Proper ((≡) ==> (≡)) ty_of_st.
@@ -472,7 +526,7 @@ Instance: Params (@type_incl) 2.
 (* Typeclasses Opaque type_incl. *)
 
 Definition subtype `{typeG Σ} (E : elctx) (L : llctx) (ty1 ty2 : type) : Prop :=
-  elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗ type_incl ty1 ty2.
+  ∀ qL, llctx_interp L qL -∗ □ (elctx_interp E -∗ type_incl ty1 ty2).
 Instance: Params (@subtype) 4.
 
 (* TODO: The prelude should have a symmetric closure. *)
@@ -505,12 +559,30 @@ Section subtyping.
   Qed.
 
   Lemma subtype_refl E L ty : subtype E L ty ty.
-  Proof. iIntros. iApply type_incl_refl. Qed.
+  Proof. iIntros (?) "_ !# _". iApply type_incl_refl. Qed.
   Global Instance subtype_preorder E L : PreOrder (subtype E L).
   Proof.
     split; first by intros ?; apply subtype_refl.
-    intros ty1 ty2 ty3 H12 H23. iIntros.
-    iApply type_incl_trans. by iApply H12. by iApply H23.
+    intros ty1 ty2 ty3 H12 H23. iIntros (?) "HL".
+    iDestruct (H12 with "HL") as "#H12".
+    iDestruct (H23 with "HL") as "#H23".
+    iClear "∗". iIntros "!# #HE".
+    iApply (type_incl_trans with "[#]"). by iApply "H12". by iApply "H23".
+  Qed.
+
+  Lemma subtype_Forall2_llctx E L tys1 tys2 qL :
+    Forall2 (subtype E L) tys1 tys2 →
+    llctx_interp L qL -∗ □ (elctx_interp E -∗
+           [∗ list] tys ∈ (zip tys1 tys2), type_incl (tys.1) (tys.2)).
+  Proof.
+    iIntros (Htys) "HL".
+    iAssert ([∗ list] tys ∈ zip tys1 tys2,
+             □ (elctx_interp _ -∗ type_incl (tys.1) (tys.2)))%I as "#Htys".
+    { iApply big_sepL_forall. iIntros (k [ty1 ty2] Hlookup).
+      move:Htys => /Forall2_Forall /Forall_forall=>Htys.
+      iApply (Htys (ty1, ty2)); first by exact: elem_of_list_lookup_2. done. }
+    iClear "∗". iIntros "!# #HE". iApply (big_sepL_impl with "[$Htys]").
+    iIntros "!# * % #Hincl". by iApply "Hincl".
   Qed.
 
   Lemma equiv_subtype E L ty1 ty2 : ty1 ≡ ty2 → subtype E L ty1 ty2.
@@ -518,21 +590,27 @@ Section subtyping.
 
   Lemma eqtype_unfold E L ty1 ty2 :
     eqtype E L ty1 ty2 ↔
-    (elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗
+    (∀ qL, llctx_interp L qL -∗ □ (elctx_interp E -∗
       (⌜ty1.(ty_size) = ty2.(ty_size)⌝ ∗
       (□ ∀ tid vl, ty1.(ty_own) tid vl ↔ ty2.(ty_own) tid vl) ∗
-      (□ ∀ κ tid l, ty1.(ty_shr) κ tid l ↔ ty2.(ty_shr) κ tid l))%I).
+      (□ ∀ κ tid l, ty1.(ty_shr) κ tid l ↔ ty2.(ty_shr) κ tid l)))%I).
   Proof.
     split.
-    - iIntros ([EQ1 EQ2]) "#HE #HL".
-      iDestruct (EQ1 with "HE HL") as "[#Hsz [#H1own #H1shr]]".
-      iDestruct (EQ2 with "HE HL") as "[_ [#H2own #H2shr]]".
+    - iIntros ([EQ1 EQ2] qL) "HL".
+      iDestruct (EQ1 with "HL") as "#EQ1".
+      iDestruct (EQ2 with "HL") as "#EQ2".
+      iClear "∗". iIntros "!# #HE".
+      iDestruct ("EQ1" with "HE") as "[#Hsz [#H1own #H1shr]]".
+      iDestruct ("EQ2" with "HE") as "[_ [#H2own #H2shr]]".
       iSplit; last iSplit.
       + done.
       + by iIntros "!#*"; iSplit; iIntros "H"; [iApply "H1own"|iApply "H2own"].
       + by iIntros "!#*"; iSplit; iIntros "H"; [iApply "H1shr"|iApply "H2shr"].
-    - intros EQ. split; iIntros "#HE #HL"; (iSplit; last iSplit);
-      iDestruct (EQ with "HE HL") as "[% [#Hown #Hshr]]".
+    - intros EQ. split; (iIntros (qL) "HL";
+      iDestruct (EQ with "HL") as "#EQ";
+      iClear "∗"; iIntros "!# #HE";
+      iDestruct ("EQ" with "HE") as "[% [#Hown #Hshr]]";
+      (iSplit; last iSplit)).
       + done.
       + iIntros "!#* H". by iApply "Hown".
       + iIntros "!#* H". by iApply "Hshr".
@@ -560,31 +638,31 @@ Section subtyping.
   Qed.
 
   Lemma subtype_simple_type E L (st1 st2 : simple_type) :
-    (∀ tid vl, elctx_interp_0 E -∗ ⌜llctx_interp_0 L⌝ -∗
-                 st1.(st_own) tid vl -∗ st2.(st_own) tid vl) →
+    (∀ qL, llctx_interp L qL -∗ □ (elctx_interp E -∗
+       ∀ tid vl, st1.(st_own) tid vl -∗ st2.(st_own) tid vl)) →
     subtype E L st1 st2.
   Proof.
-    intros Hst. iIntros. iSplit; first done. iSplit; iAlways.
-    - iIntros. iApply Hst; done.
+    intros Hst. iIntros (qL) "HL". iDestruct (Hst with "HL") as "#Hst".
+    iClear "∗". iIntros "!# #HE". iSplit; first done. iSplit; iAlways.
+    - iIntros. iApply "Hst"; done.
     - iIntros (???). iDestruct 1 as (vl) "[Hf Hown]". iExists vl. iFrame "Hf".
-      by iApply Hst.
+      by iApply "Hst".
   Qed.
 
   Lemma subtype_weaken E1 E2 L1 L2 ty1 ty2 :
     E1 ⊆+ E2 → L1 ⊆+ L2 →
     subtype E1 L1 ty1 ty2 → subtype E2 L2 ty1 ty2.
   Proof.
-    iIntros (HE12 ? Hsub) "HE HL".
-    iApply (Hsub with "[HE] [HL]").
-    - rewrite /elctx_interp_0. by iApply big_sepL_submseteq.
-    - iDestruct "HL" as %HL. iPureIntro. intros ??. apply HL.
-      eauto using elem_of_submseteq.
+    iIntros (HE12 ? Hsub qL) "HL". iDestruct (Hsub with "[HL]") as "#Hsub".
+    { rewrite /llctx_interp. by iApply big_sepL_submseteq. }
+    iClear "∗". iIntros "!# #HE". iApply "Hsub".
+    rewrite /elctx_interp. by iApply big_sepL_submseteq.
   Qed.
 End subtyping.
 
 Section type_util.
   Context `{typeG Σ}.
-  
+
   Lemma heap_mapsto_ty_own l ty tid :
     l ↦∗: ty_own ty tid ⊣⊢ ∃ (vl : vec val ty.(ty_size)), l ↦∗ vl ∗ ty_own ty tid vl.
   Proof.
@@ -594,7 +672,9 @@ Section type_util.
       iExists (list_to_vec vl). rewrite vec_to_list_of_list. iFrame.
     - iIntros "H". iDestruct "H" as (vl) "[Hl Hown]". eauto with iFrame.
   Qed.
+
 End type_util.
 
+Hint Resolve ty_outlives_E_elctx_sat tyl_outlives_E_elctx_sat : lrust_typing.
 Hint Resolve subtype_refl eqtype_refl : lrust_typing.
 Hint Opaque subtype eqtype : lrust_typing.
diff --git a/theories/typing/type_context.v b/theories/typing/type_context.v
index 92ef13e3e0da06335ed04899ac9f69f6eb3824fd..077f13ce1dd0947719f7be9d0f5bf73841f3059f 100644
--- a/theories/typing/type_context.v
+++ b/theories/typing/type_context.v
@@ -176,32 +176,31 @@ Section type_context.
 
   (** Type context inclusion *)
   Definition tctx_incl (E : elctx) (L : llctx) (T1 T2 : tctx): Prop :=
-    ∀ tid q1 q2, lft_ctx -∗ elctx_interp E q1 -∗ llctx_interp L q2 -∗
-              tctx_interp tid T1 ={⊤}=∗ elctx_interp E q1 ∗ llctx_interp L q2 ∗
-                                     tctx_interp tid T2.
+    ∀ tid q2, lft_ctx -∗ elctx_interp E -∗ llctx_interp L q2 -∗
+              tctx_interp tid T1 ={⊤}=∗ llctx_interp L q2 ∗ tctx_interp tid T2.
   Global Arguments tctx_incl _%EL _%LL _%TC _%TC.
   Global Instance : ∀ E L, RewriteRelation (tctx_incl E L).
 
   Global Instance tctx_incl_preorder E L : PreOrder (tctx_incl E L).
   Proof.
     split.
-    - by iIntros (????) "_ $ $ $".
-    - iIntros (??? H1 H2 ???) "#LFT HE HL H".
-      iMod (H1 with "LFT HE HL H") as "(HE & HL & H)".
-      by iMod (H2 with "LFT HE HL H") as "($ & $ & $)".
+    - by iIntros (???) "_ _ $ $".
+    - iIntros (??? H1 H2 ??) "#LFT #HE HL H".
+      iMod (H1 with "LFT HE HL H") as "(HL & H)".
+      by iMod (H2 with "LFT HE HL H") as "($ & $)".
   Qed.
 
   Lemma contains_tctx_incl E L T1 T2 : T1 ⊆+ T2 → tctx_incl E L T2 T1.
   Proof.
-    rewrite /tctx_incl. iIntros (Hc ???) "_ $ $ H". by iApply big_sepL_submseteq.
+    rewrite /tctx_incl. iIntros (Hc ??) "_ _ $ H". by iApply big_sepL_submseteq.
   Qed.
 
   Global Instance tctx_incl_app E L :
     Proper (tctx_incl E L ==> tctx_incl E L ==> tctx_incl E L) app.
   Proof.
-    intros ?? Hincl1 ?? Hincl2 ???. rewrite /tctx_interp !big_sepL_app.
-    iIntros "#LFT HE HL [H1 H2]".
-    iMod (Hincl1 with "LFT HE HL H1") as "(HE & HL & $)".
+    intros ?? Hincl1 ?? Hincl2 ??. rewrite /tctx_interp !big_sepL_app.
+    iIntros "#LFT #HE HL [H1 H2]".
+    iMod (Hincl1 with "LFT HE HL H1") as "(HL & $)".
     iApply (Hincl2 with "LFT HE HL H2").
   Qed.
   Global Instance tctx_incl_cons E L x :
@@ -218,7 +217,7 @@ Section type_context.
   Lemma copy_tctx_incl E L p `{!Copy ty} :
     tctx_incl E L [p ◁ ty] [p ◁ ty; p ◁ ty].
   Proof.
-    iIntros (???) "_ $ $ *". rewrite /tctx_interp !big_sepL_cons big_sepL_nil.
+    iIntros (??) "_ _ $ *". rewrite /tctx_interp !big_sepL_cons big_sepL_nil.
     by iIntros "[#$ $]".
   Qed.
 
@@ -234,11 +233,9 @@ Section type_context.
   Lemma subtype_tctx_incl E L p ty1 ty2 :
     subtype E L ty1 ty2 → tctx_incl E L [p ◁ ty1] [p ◁ ty2].
   Proof.
-    iIntros (Hst ???) "#LFT HE HL H". rewrite /tctx_interp !big_sepL_singleton /=.
-    iDestruct (elctx_interp_persist with "HE") as "#HE'".
-    iDestruct (llctx_interp_persist with "HL") as "#HL'".
-    iFrame "HE HL". iDestruct "H" as (v) "[% H]". iExists _. iFrame "%".
-    iDestruct (Hst with "[] []") as "(_ & #Ho & _)"; [done..|by iApply "Ho"].
+    iIntros (Hst ??) "#LFT #HE HL H". rewrite /tctx_interp !big_sepL_singleton /=.
+    iDestruct "H" as (v) "[% H]". iDestruct (Hst with "HL HE") as "#(_ & Ho & _)". 
+    iFrame "HL". iExists _. iFrame "%". by iApply "Ho".
   Qed.
 
   (* Extracting from a type context. *)
diff --git a/theories/typing/type_sum.v b/theories/typing/type_sum.v
index 14ea186cc53c61a16041277cc7bccc2b9efdb93d..93dbce64e029532fc88a78e3b2863db952aeedf1 100644
--- a/theories/typing/type_sum.v
+++ b/theories/typing/type_sum.v
@@ -16,7 +16,7 @@ Section case.
       typed_body E L C ((p ◁ own_ptr n (sum tyl)) :: T) e) tyl el →
     typed_body E L C ((p ◁ own_ptr n (sum tyl)) :: T) (case: !p of el).
   Proof.
-    iIntros (Hel tid qE) "#LFT Hna HE HL HC HT". wp_bind p.
+    iIntros (Hel tid) "#LFT #HE Hna HL HC HT". wp_bind p.
     rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]".
     iApply (wp_hasty with "Hp"). iIntros ([[]|] Hv) "Hp"; try done.
     iDestruct "Hp" as "[H↦ Hf]". iDestruct "H↦" as (vl) "[H↦ Hown]".
@@ -30,7 +30,7 @@ Section case.
     destruct (tyl !! i) as [ty|] eqn:EQty; last iDestruct "Hown" as ">[]".
     edestruct @Forall2_lookup_l as (e & He & Hety); eauto.
     wp_read. iApply wp_case; [lia|by rewrite Nat2Z.id|]. iNext.
-    destruct Hety as [Hety|Hety]; iApply (Hety with "LFT Hna HE HL HC");
+    destruct Hety as [Hety|Hety]; iApply (Hety with "LFT HE Hna HL HC");
       rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //; iFrame "HT".
     - iDestruct (ty.(ty_size_eq) with "Hown") as %<-.
       iSplitL "H↦i Hfi"; last iSplitR "H↦vl'' Hfvl''".
@@ -61,7 +61,7 @@ Section case.
       typed_body E L C ((p ◁ &uniq{κ}sum tyl) :: T) e) tyl el →
     typed_body E L C ((p ◁ &uniq{κ}sum tyl) :: T) (case: !p of el).
   Proof.
-    iIntros (Halive Hel tid qE) "#LFT Hna HE HL HC HT". wp_bind p.
+    iIntros (Halive Hel tid) "#LFT #HE Hna HL HC HT". wp_bind p.
     rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]".
     iApply (wp_hasty with "Hp"). iIntros ([[]|] Hv) "Hp"; try iDestruct "Hp" as "[]".
     iMod (Halive with "HE HL") as (q) "[Htok Hclose]". done.
@@ -85,16 +85,16 @@ Section case.
         iFrame. iExists _, _, _. iSplit. by auto.
         rewrite /= -EQlen !app_length EQlenvl' EQlenvl'2 nth_lookup EQty /=. auto. }
       { iExists vl'. iFrame. }
-      iMod ("Hclose" with "Htok") as "[HE HL]".
-      iApply (Hety with "LFT Hna HE HL HC").
+      iMod ("Hclose" with "Htok") as "HL".
+      iApply (Hety with "LFT HE Hna HL HC").
       rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //. iFrame.
     - iMod ("Hclose'" with "[] [H↦i H↦vl' H↦vl'' Hown]") as "[Hb Htok]";
         [by iIntros "!>$"| |].
       { iExists (#i::vl'++vl'').
         rewrite heap_mapsto_vec_cons heap_mapsto_vec_app /= -EQlen. iFrame. iNext.
         iExists i, vl', vl''. rewrite nth_lookup EQty. auto. }
-      iMod ("Hclose" with "Htok") as "[HE HL]".
-      iApply (Hety with "LFT Hna HE HL HC").
+      iMod ("Hclose" with "Htok") as "HL".
+      iApply (Hety with "LFT HE Hna HL HC").
       rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //. iFrame.
   Qed.
 
@@ -114,7 +114,7 @@ Section case.
       typed_body E L C ((p ◁ &shr{κ}sum tyl) :: T) e) tyl el →
     typed_body E L C ((p ◁ &shr{κ}sum tyl) :: T) (case: !p of el).
   Proof.
-    iIntros (Halive Hel tid qE) "#LFT Hna HE HL HC HT". wp_bind p.
+    iIntros (Halive Hel tid) "#LFT #HE Hna HL HC HT". wp_bind p.
     rewrite tctx_interp_cons. iDestruct "HT" as "[Hp HT]".
     iApply (wp_hasty with "Hp"). iIntros ([[]|] Hv) "Hp"; try done.
     iDestruct "Hp" as (i) "[#Hb Hshr]".
@@ -125,8 +125,8 @@ Section case.
     edestruct @Forall2_lookup_l as (e & He & Hety); eauto.
     wp_read. iApply wp_case; [lia|by rewrite Nat2Z.id|]. iNext.
     iMod ("Hclose'" with "[$H↦i $H↦vl'']") as "Htok".
-    iMod ("Hclose" with "Htok") as "[HE HL]".
-    destruct Hety as [Hety|Hety]; iApply (Hety with "LFT Hna HE HL HC");
+    iMod ("Hclose" with "Htok") as "HL".
+    destruct Hety as [Hety|Hety]; iApply (Hety with "LFT HE Hna HL HC");
       rewrite !tctx_interp_cons !tctx_hasty_val' /= ?Hv //; iFrame.
     iExists _. rewrite ->nth_lookup, EQty. auto.
   Qed.
@@ -146,7 +146,7 @@ Section case.
     typed_write E L ty1 (sum tyl) ty2 →
     typed_instruction E L [p1 ◁ ty1; p2 ◁ ty] (p1 <-{Σ i} p2) (λ _, [p1 ◁ ty2]%TC).
   Proof.
-    iIntros (Hty Hw tid qE) "#LFT $ HE HL Hp".
+    iIntros (Hty Hw tid) "#LFT #HE $ HL Hp".
     rewrite tctx_interp_cons tctx_interp_singleton.
     iDestruct "Hp" as "[Hp1 Hp2]". iDestruct (closed_hasty with "Hp1") as "%".
     iDestruct (closed_hasty with "Hp2") as "%". wp_bind p1.
@@ -186,7 +186,7 @@ Section case.
     typed_write E L ty1 (sum tyl) ty2 →
     typed_instruction E L [p ◁ ty1] (p <-{Σ i} ()) (λ _, [p ◁ ty2]%TC).
   Proof.
-    iIntros (Hty Hw tid qE) "#LFT $ HE HL Hp". rewrite tctx_interp_singleton.
+    iIntros (Hty Hw tid) "#LFT #HE $ HL Hp". rewrite tctx_interp_singleton.
     wp_bind p. iApply (wp_hasty with "Hp"). iIntros (v Hv) "Hty".
     iMod (Hw with "[] LFT HE HL Hty") as (l vl) "(H & H↦ & Hw)". done.
     simpl. destruct vl as [|? vl]; iDestruct "H" as %[[= Hlen] ->].
@@ -217,17 +217,17 @@ Section case.
     typed_instruction E L [p1 ◁ ty1; p2 ◁ ty2]
                (p1 <-{ty.(ty_size),Σ i} !p2) (λ _, [p1 ◁ ty1'; p2 ◁ ty2']%TC).
   Proof.
-    iIntros (Hty Hw Hr tid qE) "#LFT Htl [HE1 HE2] [HL1 HL2] Hp".
+    iIntros (Hty Hw Hr tid) "#LFT #HE Htl [HL1 HL2] Hp".
     rewrite tctx_interp_cons tctx_interp_singleton.
     iDestruct "Hp" as "[Hp1 Hp2]". iDestruct (closed_hasty with "Hp1") as "%".
     iDestruct (closed_hasty with "Hp2") as "%". wp_bind p1.
     iApply (wp_hasty with "Hp1"). iIntros (v1 Hv1) "Hty1".
-    iMod (Hw with "[] LFT HE1 HL1 Hty1") as (l1 vl1) "(H & H↦ & Hw)"=>//=.
+    iMod (Hw with "[] LFT HE HL1 Hty1") as (l1 vl1) "(H & H↦ & Hw)"=>//=.
     destruct vl1 as [|? vl1]; iDestruct "H" as %[[= Hlen] ->].
     rewrite heap_mapsto_vec_cons -wp_fupd. iDestruct "H↦" as "[H↦0 H↦vl1]". wp_write.
     wp_bind p1. iApply (wp_wand with "[]"); first by iApply (wp_eval_path). iIntros (? ->).
     wp_op. wp_bind p2. iApply (wp_hasty with "Hp2"). iIntros (v2 Hv2) "Hty2".
-    iMod (Hr with "[] LFT Htl HE2 HL2 Hty2") as (l2 vl2 q) "(% & H↦2 & Hty & Hr)"=>//=.
+    iMod (Hr with "[] LFT HE Htl HL2 Hty2") as (l2 vl2 q) "(% & H↦2 & Hty & Hr)"=>//=.
     subst. assert (ty.(ty_size) ≤ length vl1).
     { revert i Hty. rewrite Hlen. clear. induction tyl=>//= -[|i] /=.
       - intros [= ->]. lia.
@@ -239,7 +239,7 @@ Section case.
     { rewrite take_length. lia. }
     iNext; iIntros "[H↦vl1 H↦2]".
     rewrite tctx_interp_cons tctx_interp_singleton !tctx_hasty_val' //.
-    iMod ("Hr" with "H↦2") as "($ & $ & $ & $)". iApply "Hw". iNext.
+    iMod ("Hr" with "H↦2") as "($ & $ & $)". iApply "Hw". iNext.
     rewrite split_sum_mt /is_pad. iExists i. rewrite nth_lookup Hty. iFrame.
     iSplitL "H↦pad".
     - rewrite (shift_loc_assoc_nat _ 1) take_length Nat.min_l; last lia.
diff --git a/theories/typing/uninit.v b/theories/typing/uninit.v
index 294fb3721164eff16ac29a7041e4ed8fb3fa584d..b8852230083a62bfcaeac3e7c47d88c62ecadd6f 100644
--- a/theories/typing/uninit.v
+++ b/theories/typing/uninit.v
@@ -17,7 +17,7 @@ Section uninit.
     Π (replicate n uninit_1).
 
   Lemma uninit0_sz n : ty_size (uninit0 n) = n.
-  Proof. induction n. done. simpl. by f_equal. Qed.
+  Proof. induction n=>//=. by f_equal. Qed.
 
   (* We redefine uninit as an alias of uninit0, so that the size
      computes directly to [n] *)
@@ -28,6 +28,9 @@ Section uninit.
   Next Obligation. intros. by apply ty_share. Qed.
   Next Obligation. intros. by apply ty_shr_mono. Qed.
 
+  Global Instance uninit_wf n : TyWf (uninit n) :=
+    { ty_lfts := []; ty_wf_E := [] }.
+
   Global Instance uninit_copy n : Copy (uninit n).
   Proof.
     destruct (product_copy (replicate n uninit_1)) as [A B].
diff --git a/theories/typing/uniq_bor.v b/theories/typing/uniq_bor.v
index d39e3a47fcfa487d0cb4c902618a0d43bc44c357..e8f07a2d160210416c3cef2164ea2b6f826dde12 100644
--- a/theories/typing/uniq_bor.v
+++ b/theories/typing/uniq_bor.v
@@ -43,12 +43,17 @@ Section uniq_bor.
     by iApply (ty_shr_mono with "Hκ0").
   Qed.
 
+  Global Instance uniq_bor_wf κ ty `{!TyWf ty} : TyWf (uniq_bor κ ty) :=
+    { ty_lfts := [κ]; ty_wf_E := ty.(ty_wf_E) ++ ty.(ty_outlives_E) κ }.
+
   Global Instance uniq_mono E L :
     Proper (flip (lctx_lft_incl E L) ==> eqtype E L ==> subtype E L) uniq_bor.
   Proof.
-    intros κ1 κ2 Hκ ty1 ty2 Hty%eqtype_unfold. iIntros. iSplit; first done.
-    iDestruct (Hty with "[] []") as "(_ & #Ho & #Hs)"; [done..|clear Hty].
-    iDestruct (Hκ with "[] []") as "#Hκ"; [done..|]. iSplit; iAlways.
+    intros κ1 κ2 Hκ ty1 ty2. rewrite eqtype_unfold=>Hty. iIntros (?) "HL".
+    iDestruct (Hty with "HL") as "#Hty". iDestruct (Hκ with "HL") as "#Hκ".
+    iIntros "!# #HE". iSplit; first done.
+    iDestruct ("Hty" with "HE") as "(_ & #Ho & #Hs)"; [done..|clear Hty].
+    iSpecialize ("Hκ" with "HE"). iSplit; iAlways.
     - iIntros (? [|[[]|][]]) "H"; try done.
       iApply (bor_shorten with "Hκ"). iApply bor_iff; last done.
       iSplit; iIntros "!>!# H"; iDestruct "H" as (vl) "[??]";
@@ -109,7 +114,7 @@ Section typing.
   Lemma tctx_share E L p κ ty :
     lctx_lft_alive E L κ → tctx_incl E L [p ◁ &uniq{κ}ty] [p ◁ &shr{κ}ty].
   Proof.
-    iIntros (Hκ ???) "#LFT HE HL Huniq".
+    iIntros (Hκ ??) "#LFT #HE HL Huniq".
     iMod (Hκ with "HE HL") as (q) "[Htok Hclose]"; [try done..|].
     rewrite !tctx_interp_singleton /=.
     iDestruct "Huniq" as ([[]|]) "[% Huniq]"; try done.
@@ -121,11 +126,8 @@ Section typing.
     lctx_lft_incl E L κ' κ →
     tctx_incl E L [p ◁ &uniq{κ}ty] [p ◁ &uniq{κ'}ty; p ◁{κ'} &uniq{κ}ty].
   Proof.
-    iIntros (Hκκ' tid ??) "#LFT HE HL H".
-    iDestruct (elctx_interp_persist with "HE") as "#HE'".
-    iDestruct (llctx_interp_persist with "HL") as "#HL'". iFrame "HE HL".
-    iDestruct (Hκκ' with "HE' HL'") as "Hκκ'".
-    rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
+    iIntros (Hκκ' tid ?) "#LFT HE HL H". iDestruct (Hκκ' with "HL HE") as "#Hκκ'".
+    iFrame. rewrite tctx_interp_singleton tctx_interp_cons tctx_interp_singleton.
     iDestruct "H" as ([[]|]) "[% Hb]"; try done.
     iMod (rebor with "LFT Hκκ' Hb") as "[Hb Hext]". done. iModIntro.
     iSplitL "Hb"; iExists _; auto.
@@ -174,7 +176,7 @@ Section typing.
     Copy ty → lctx_lft_alive E L κ → typed_read E L (&uniq{κ}ty) ty (&uniq{κ}ty).
   Proof.
     iIntros (Hcopy Halive) "!#".
-    iIntros ([[]|] tid F qE qL ?) "#LFT Htl HE HL Hown"; try done.
+    iIntros ([[]|] tid F qL ?) "#LFT #HE Htl HL Hown"; try done.
     iMod (Halive with "HE HL") as (q) "[Hκ Hclose]"; first set_solver.
     iMod (bor_acc with "LFT Hown Hκ") as "[H↦ Hclose']"; first set_solver.
     iDestruct "H↦" as (vl) "[>H↦ #Hown]".
@@ -188,7 +190,7 @@ Section typing.
     lctx_lft_alive E L κ → typed_write E L (&uniq{κ}ty) ty (&uniq{κ}ty).
   Proof.
     iIntros (Halive) "!#".
-    iIntros ([[]|] tid F qE qL ?) "#LFT HE HL Hown"; try done.
+    iIntros ([[]|] tid F qL ?) "#LFT HE HL Hown"; try done.
     iMod (Halive with "HE HL") as (q) "[Htok Hclose]"; first set_solver.
     iMod (bor_acc with "LFT Hown Htok") as "[H↦ Hclose']". set_solver.
     iDestruct "H↦" as (vl) "[>H↦ Hown]". rewrite ty.(ty_size_eq).