From 6fc57892f8ef6821d3d2d4d14604456bfd2a5082 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Thu, 28 May 2020 17:31:51 +0200 Subject: [PATCH] Bump Iris (Z_scope). --- opam | 2 +- theories/examples/cell.v | 6 +++--- theories/examples/namegen.v | 10 +++++----- theories/examples/or.v | 4 ++-- theories/examples/symbol.v | 10 +++++----- theories/examples/ticket_lock.v | 6 +++--- theories/examples/various.v | 12 ++++++------ theories/experimental/helping/offers.v | 2 +- theories/experimental/hocap/counter.v | 6 +++--- theories/experimental/hocap/ticket_lock.v | 2 +- theories/logic/adequacy.v | 2 +- theories/logic/proofmode/tactics.v | 2 +- theories/logic/spec_rules.v | 2 +- theories/typing/contextual_refinement.v | 6 +++--- theories/typing/fundamental.v | 4 ++-- theories/typing/interp.v | 6 +++--- theories/typing/types.v | 4 ++-- 17 files changed, 43 insertions(+), 43 deletions(-) diff --git a/opam b/opam index 4ca4f67..5e63a3d 100644 --- a/opam +++ b/opam @@ -9,6 +9,6 @@ build: [make "-j%{jobs}%"] install: [make "install"] remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/reloc"] depends: [ - "coq-iris" { (= "dev.2020-05-26.1.d80e7abf") | (= "dev") } + "coq-iris" { (= "dev.2020-05-28.3.38f177d3") | (= "dev") } "coq-autosubst" { = "dev.coq86" } ] diff --git a/theories/examples/cell.v b/theories/examples/cell.v index 38f290f..9936ba5 100644 --- a/theories/examples/cell.v +++ b/theories/examples/cell.v @@ -7,9 +7,9 @@ From reloc.lib Require Export lock. (** A type of cells -- basically an abstract type of references. *) (* ∀ α, ∃ β, (α → β) × (β → α) × (β → α → ()) *) Definition cellÏ„ : type := - ∀: ∃: (TVar 1%nat → TVar 0%nat) - * (TVar 0%nat → TVar 1%nat) - * (TVar 0%nat → TVar 1%nat → TUnit). + ∀: ∃: (TVar 1 → TVar 0) + * (TVar 0 → TVar 1) + * (TVar 0 → TVar 1 → TUnit). (** We show that the canonical implementation `cell1` is equivalent to an implementation using two alternating slots *) diff --git a/theories/examples/namegen.v b/theories/examples/namegen.v index 0694e94..aa759d1 100644 --- a/theories/examples/namegen.v +++ b/theories/examples/namegen.v @@ -11,8 +11,8 @@ version uses integers. *) (* ^ new name ^ *) (* | compare names for equality *) Definition nameGenTy : type := - ∃: (TUnit → TVar 0%nat) - * (TVar 0%nat → TVar 0%nat → TBool). + ∃: (TUnit → TVar 0) + * (TVar 0 → TVar 0 → TBool). (* TODO: cannot be a value *) Definition nameGen1 : expr := @@ -35,7 +35,7 @@ Section namegen_refinement. (∃ (n : nat) (L : gset (loc * nat)), BIJ γ L ∗ c ↦ₛ #n ∗ [∗ set] lk ∈ L, match lk with - | (l, k) => l ↦ #() ∗ ⌜k ≤ nâŒ%nat + | (l, k) => l ↦ #() ∗ ⌜k ≤ n⌠end)%I. Lemma nameGen_ref1 : @@ -46,7 +46,7 @@ Section namegen_refinement. iMod alloc_empty_bij as (γ) "HB". pose (N:=relocN.@"ng"). iMod (inv_alloc N _ (ng_Inv γ c) with "[-]") as "#Hinv". - { iNext. iExists 0%nat, ∅. iFrame. + { iNext. iExists 0, ∅. iFrame. by rewrite big_sepS_empty. } iApply (refines_exists (ngR γ)). do 2 rel_pure_r. @@ -82,7 +82,7 @@ Section namegen_refinement. intros [l x] Hlx. apply bi.sep_mono_r, bi.pure_mono. lia. } rel_values. iModIntro. replace (Z.of_nat n + 1)%Z with (Z.of_nat (S n)); last lia. - iExists l', (S n)%nat; eauto. + iExists l', (S n); eauto. - (* Name comparison *) rel_pure_l. rel_pure_r. iApply refines_arrow_val. diff --git a/theories/examples/or.v b/theories/examples/or.v index 5ce043a..2b70b28 100644 --- a/theories/examples/or.v +++ b/theories/examples/or.v @@ -325,12 +325,12 @@ Proof. + econstructor; cbn; eauto. econstructor; cbn; eauto. econstructor; cbn; eauto. - change Z0 with (Z.of_nat 0%nat). + change Z0 with (Z.of_nat 0). econstructor; cbn; eauto. econstructor; cbn; eauto. * econstructor; cbn; eauto. eapply Seq_typed. - ** change 1%Z with (Z.of_nat 1%nat). + ** change 1%Z with (Z.of_nat 1). repeat (econstructor; cbn; eauto). ** repeat (econstructor; cbn; eauto). * repeat (econstructor; cbn; eauto). diff --git a/theories/examples/symbol.v b/theories/examples/symbol.v index d136a7e..5d173f2 100644 --- a/theories/examples/symbol.v +++ b/theories/examples/symbol.v @@ -60,7 +60,7 @@ Section rules. Lemma size_le (n m : nat) : own γ (â— (n : mnat)) -∗ own γ (â—¯ (m : mnat)) -∗ - ⌜m ≤ nâŒ%nat. + ⌜m ≤ nâŒ. Proof. iIntros "Hn Hm". by iDestruct (own_valid_2 with "Hn Hm") @@ -225,7 +225,7 @@ Section proof. rel_alloc_r size2 as "[Hs2 Hs2']"; repeat rel_pure_r. rel_alloc_l tbl1 as "[Htbl1 Htbl1']"; repeat rel_pure_l. rel_alloc_r tbl2 as "[Htbl2 Htbl2']"; repeat rel_pure_r. - iMod (own_alloc (â— (0%nat : mnat))) as (γ) "Ha". + iMod (own_alloc (â— (0 : mnat))) as (γ) "Ha". { by apply auth_auth_valid. } iMod (inv_alloc sizeN _ (table_inv γ size1 size2 tbl1 tbl2) with "[Hs1 Hs2 Htbl1 Htbl2 Ha]") as "#Hinv". { iNext. iExists _,_. iFrame. iApply lrel_list_nil. } @@ -233,7 +233,7 @@ Section proof. iIntros (l2) "Hl2". rel_pure_r. rel_pure_r. pose (N:=(relocN.@"lock1")). rel_apply_l (refines_newlock_l N (lok_inv size1 size2 tbl1 tbl2 l2)%I with "[Hs1' Hs2' Htbl1' Htbl2' Hl2]"). - { iExists 0%nat,_. by iFrame. } + { iExists 0,_. by iFrame. } iNext. iIntros (l1 γl) "#Hl1". rel_pure_l. rel_pure_l. iApply (refines_exists (tableR γ)). repeat iApply refines_pair. @@ -304,7 +304,7 @@ Section proof. rel_alloc_r size2 as "[Hs2 Hs2']"; repeat rel_pure_r. rel_alloc_l tbl1 as "[Htbl1 Htbl1']"; repeat rel_pure_l. rel_alloc_r tbl2 as "[Htbl2 Htbl2']"; repeat rel_pure_r. - iMod (own_alloc (â— (0%nat : mnat))) as (γ) "Ha". + iMod (own_alloc (â— (0 : mnat))) as (γ) "Ha". { by apply auth_auth_valid. } iMod (inv_alloc sizeN _ (table_inv γ size1 size2 tbl1 tbl2) with "[Hs1 Hs2 Htbl1 Htbl2 Ha]") as "#Hinv". { iNext. iExists _,_. iFrame. iApply lrel_list_nil. } @@ -312,7 +312,7 @@ Section proof. iIntros (l2) "Hl2". rel_pure_r. rel_pure_r. pose (N:=(relocN.@"lock1")). rel_apply_l (refines_newlock_l N (lok_inv size1 size2 tbl1 tbl2 l2)%I with "[Hs1' Hs2' Htbl1' Htbl2' Hl2]"). - { iExists 0%nat,_. by iFrame. } + { iExists 0,_. by iFrame. } iNext. iIntros (l1 γl) "#Hl1". rel_pure_l. rel_pure_l. iApply (refines_exists (tableR γ)). repeat iApply refines_pair. diff --git a/theories/examples/ticket_lock.v b/theories/examples/ticket_lock.v index edecad1..3230b98 100644 --- a/theories/examples/ticket_lock.v +++ b/theories/examples/ticket_lock.v @@ -118,7 +118,7 @@ Section refinement. (* Establishing the invariant *) iMod newIssuedTickets as (γ) "Hγ". iMod (inv_alloc N _ (lockInv lo ln γ l') with "[-]") as "#Hinv". - { iNext. iExists 0%nat,0%nat,_. by iFrame. } + { iNext. iExists 0,0,_. by iFrame. } rel_pure_l. rel_values. iExists _,_,_. iFrame "Hinv". eauto. @@ -335,8 +335,8 @@ Section refinement. iIntros "Hl'". iMod ("Hcl" with "[-]") as "_". { iNext. - replace (o' + 1)%Z with (Z.of_nat (o' + 1))%nat by lia. - iExists (o' + 1)%nat,_,_. by iFrame. } + replace (o' + 1)%Z with (Z.of_nat (o' + 1)) by lia. + iExists (o' + 1),_,_. by iFrame. } rel_values. Qed. diff --git a/theories/examples/various.v b/theories/examples/various.v index d986b63..146fdee 100644 --- a/theories/examples/various.v +++ b/theories/examples/various.v @@ -190,7 +190,7 @@ Section proofs. rel_alloc_r x' as "Hx'". rel_pure_r. rel_pure_r. iMod (inv_alloc i3n _ (i3 x x' b b') with "[-]") as "#Hinv". - { iNext. unfold i3. iExists 0%nat. + { iNext. unfold i3. iExists 0. iDestruct "Hx" as "[$ Hx]". iDestruct "Hx'" as "[$ Hx']". iLeft. iFrame. } @@ -388,7 +388,7 @@ Section proofs. { iIntros "[Hc1 [(Hc2 & Ht) | (Hc2 & Ht & Ht' & %)]]"; iApply ("Hcl" with "[-]"); iNext. + iExists n. iLeft. iFrame. - + iExists (n-1)%nat. iRight. + + iExists (n-1). iRight. replace (Z.of_nat (n - 1)) with (Z.of_nat n - 1)%Z by lia. replace (n - 1 + 1)%Z with (Z.of_nat n) by lia. iFrame. } @@ -396,7 +396,7 @@ Section proofs. iDestruct "Hc" as "[[Hc2 Ht] | [Hc2 [Ht [Ht' %]]]]". - rel_apply_r (FG_increment_r with "Hc2"). iIntros "Hc2". iMod ("Hcl" with "[-]") as "_". - { iNext. iExists (n + 1)%nat. + { iNext. iExists (n + 1). replace (Z.of_nat (n + 1)) with (Z.of_nat n + 1)%Z by lia. iLeft; iFrame. } rel_values. @@ -406,7 +406,7 @@ Section proofs. { iNext. iExists n. iRight; iFrame. by replace ((n - 1)%nat + 1)%Z with (Z.of_nat n) by lia. } rel_values. } - - iModIntro. iExists (n + 1)%nat. + - iModIntro. iExists (n + 1). replace (Z.of_nat n + 1)%Z with (Z.of_nat (n + 1)) by lia. iFrame. iSplitL "Hc2 Ht". { iRight. replace ((n + 1)%nat - 1)%Z with (Z.of_nat n) by lia. @@ -463,7 +463,7 @@ Section proofs. Proof. iDestruct 1 as (m) "[Hc1 Hc2]". iDestruct "Hc2" as "[[Hc2 Hp] | (Hc2 & Hs & Ht & %)]"; - [iExists m; iLeft | iExists (m - 1)%nat; iRight]; iFrame. + [iExists m; iLeft | iExists (m - 1); iRight]; iFrame. replace ((m - 1)%nat + 1)%Z with (Z.of_nat m) by lia. replace (Z.of_nat (m - 1)) with (Z.of_nat m - 1)%Z by lia. iFrame. @@ -540,7 +540,7 @@ Section proofs. iMod new_pending as (γ) "Ht". iMod (own_alloc (Excl ())) as (γ') "Ht'"; first done. iMod (inv_alloc shootN _ (i6 c1 c2 γ γ') with "[Hc1 Hc2 Ht]") as "#Hinv". - { iNext. iExists 0%nat. iLeft. iFrame. } + { iNext. iExists 0. iLeft. iFrame. } repeat rel_pure_l. repeat rel_pure_r. iApply (refines_seq lrel_unit). { iApply (refines_app with "Hf"). diff --git a/theories/experimental/helping/offers.v b/theories/experimental/helping/offers.v index 7403f5f..2bb6567 100644 --- a/theories/experimental/helping/offers.v +++ b/theories/experimental/helping/offers.v @@ -70,7 +70,7 @@ Section rules. rel_pures_l. iApply ("Hcont" with "[-]"). iExists _; iFrame. - iIntros (?). simplify_eq/=. - assert (c = 0%nat) as -> by lia. (* :) *) + assert (c = 0) as -> by lia. (* :) *) iNext. iIntros "Hl". iDestruct "Hoff" as "[[% HP] | [[% ?] | [% ?]]]"; [| congruence..]. rel_pures_l. diff --git a/theories/experimental/hocap/counter.v b/theories/experimental/hocap/counter.v index 37a1839..ba83aad 100644 --- a/theories/experimental/hocap/counter.v +++ b/theories/experimental/hocap/counter.v @@ -280,7 +280,7 @@ Section refinement. iIntros "Hl Hlk". iMod ("Hcl" with "[-]") as "_". { iNext. iExists (n+1); try iFrame. - assert (Z.of_nat (n + 1)%nat = Z.of_nat n + 1)%Z as -> by lia. + assert (Z.of_nat (n + 1) = Z.of_nat n + 1)%Z as -> by lia. done. } rel_values. Qed. @@ -317,12 +317,12 @@ Section refinement. rel_alloc_r c' as "Hcnt'". rel_alloc_l c as "Hcnt". simpl. - iMod (Cnt_alloc N _ 0%nat with "Hcnt") as (γ) "[#HCnt Hc]". + iMod (Cnt_alloc N _ 0 with "Hcnt") as (γ) "[#HCnt Hc]". (* establishing the invariant *) iMod (inv_alloc N2 _ (∃ m, is_locked_r lk false ∗ cnt γ 1 m ∗ c' ↦ₛ #m)%I with "[-]") as "#Hinv". - { iNext. iExists 0%nat. by iFrame. } + { iNext. iExists 0. by iFrame. } (* TODO: here we have to do /exactly/ 4 steps. The next step will reduce `(Val v1, Val v2)` to `Val (v1, v2)`, and the compatibility rule wouldn't be applicable *) diff --git a/theories/experimental/hocap/ticket_lock.v b/theories/experimental/hocap/ticket_lock.v index d0f2058..2aa85b9 100644 --- a/theories/experimental/hocap/ticket_lock.v +++ b/theories/experimental/hocap/ticket_lock.v @@ -117,7 +117,7 @@ Section refinement. iMod (Cnt_alloc (N.@"cnt1") _ 0 with "Hlo") as (γlo) "[#Hc2 Hlo]". iMod newIssuedTickets as (γ) "Hγ". iMod (inv_alloc (N.@"lock") _ (lockInv γln γlo γ l') with "[-]") as "#Hinv". - { iNext. iExists 0%nat,0%nat,_. iFrame. } + { iNext. iExists 0,0,_. iFrame. } rel_values. iModIntro. iExists _,_,_,_,_. iSplit; eauto with iFrame. Qed. diff --git a/theories/logic/adequacy.v b/theories/logic/adequacy.v index bb0ac2a..e82a01f 100644 --- a/theories/logic/adequacy.v +++ b/theories/logic/adequacy.v @@ -44,7 +44,7 @@ Proof. - iPoseProof (Hlog _) as "Hrel". rewrite refines_eq /refines_def /spec_ctx. iApply fupd_wp. - iSpecialize ("Hrel" $! 0%nat [] with "[Hcfg]"); first by eauto. + iSpecialize ("Hrel" $! 0 [] with "[Hcfg]"); first by eauto. iApply "Hrel". rewrite tpool_mapsto_eq /tpool_mapsto_def. iFrame. by rewrite /to_tpool /= insert_empty map_fmap_singleton //. diff --git a/theories/logic/proofmode/tactics.v b/theories/logic/proofmode/tactics.v index 3f2efed..54ea266 100644 --- a/theories/logic/proofmode/tactics.v +++ b/theories/logic/proofmode/tactics.v @@ -128,7 +128,7 @@ Lemma tac_rel_pure_l `{!relocG Σ} K e1 ℶ ℶ' E e e2 eres Ï• n m t A : e = fill K e1 → PureExec Ï• n e1 e2 → Ï• → - ((m = n ∧ E = ⊤) ∨ m = 0%nat) → + ((m = n ∧ E = ⊤) ∨ m = 0) → MaybeIntoLaterNEnvs m ℶ ℶ' → eres = fill K e2 → envs_entails ℶ' (REL eres << t @ E : A) → diff --git a/theories/logic/spec_rules.v b/theories/logic/spec_rules.v index f22546c..be1abfc 100644 --- a/theories/logic/spec_rules.v +++ b/theories/logic/spec_rules.v @@ -327,7 +327,7 @@ Section rules. iInv specN as (tp σ) ">[Hown %]" "Hclose". iDestruct (own_valid_2 with "Hown Hj") as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_both_valid. - assert (j < length tp)%nat by eauto using lookup_lt_Some. + assert (j < length tp) by eauto using lookup_lt_Some. iMod (own_update_2 with "Hown Hj") as "[Hown Hj]". { by eapply auth_update, prod_local_update_1, singleton_local_update, (exclusive_local_update _ (Excl (fill K #()))). } diff --git a/theories/typing/contextual_refinement.v b/theories/typing/contextual_refinement.v index 74c7248..ea12f0c 100644 --- a/theories/typing/contextual_refinement.v +++ b/theories/typing/contextual_refinement.v @@ -203,18 +203,18 @@ Inductive typed_ctx_item : | TP_CTX_Unfold Γ Ï„ : typed_ctx_item CTX_Unfold Γ (TRec Ï„) Γ Ï„.[(TRec Ï„)/] | TP_CTX_TLam Γ Ï„ : - typed_ctx_item CTX_TLam (Autosubst_Classes.subst (ren (+1%nat)) <$> Γ) Ï„ Γ (TForall Ï„) + typed_ctx_item CTX_TLam (Autosubst_Classes.subst (ren (+1)) <$> Γ) Ï„ Γ (TForall Ï„) | TP_CTX_TApp Γ Ï„ Ï„' : typed_ctx_item CTX_TApp Γ (TForall Ï„) Γ Ï„.[Ï„'/] (* | TP_CTX_Pack Γ Ï„ Ï„' : *) (* typed_ctx_item CTX_Pack Γ Ï„.[Ï„'/] Γ (TExists Ï„) *) | TP_CTX_UnpackL x e2 Γ Ï„ Ï„2 : - <[x:=Ï„]>(⤉ Γ) ⊢ₜ e2 : (Autosubst_Classes.subst (ren (+1%nat)) Ï„2) → + <[x:=Ï„]>(⤉ Γ) ⊢ₜ e2 : (Autosubst_Classes.subst (ren (+1)) Ï„2) → typed_ctx_item (CTX_UnpackL x e2) Γ (TExists Ï„) Γ Ï„2 | TP_CTX_UnpackR x e1 Γ Ï„ Ï„2 : Γ ⊢ₜ e1 : TExists Ï„ → typed_ctx_item (CTX_UnpackR x e1) - (<[x:=Ï„]>(⤉ Γ)) (Autosubst_Classes.subst (ren (+1%nat)) Ï„2) + (<[x:=Ï„]>(⤉ Γ)) (Autosubst_Classes.subst (ren (+1)) Ï„2) Γ Ï„2 . diff --git a/theories/typing/fundamental.v b/theories/typing/fundamental.v index 74d4d01..2eec8e2 100644 --- a/theories/typing/fundamental.v +++ b/theories/typing/fundamental.v @@ -194,7 +194,7 @@ Section fundamental. {Δ;Γ} ⊨ (e1;; e2) ≤log≤ (e1';; e2') : Ï„2. Proof. iIntros "He1 He2". - iApply (bin_log_related_seq lrel_true _ _ _ _ _ _ Ï„1.[ren (+1)%nat] with "[He1] He2"). + iApply (bin_log_related_seq lrel_true _ _ _ _ _ _ Ï„1.[ren (+1)] with "[He1] He2"). intro_clause. rewrite interp_ren -(interp_ren_up [] Δ Ï„1). by iApply "He1". @@ -472,7 +472,7 @@ Section fundamental. ({Δ;Γ} ⊨ e1 ≤log≤ e1' : ∃: Ï„) -∗ (∀ Ï„i : lrel Σ, {Ï„i::Δ;<[x:=Ï„]>(⤉Γ)} ⊨ - e2 ≤log≤ e2' : (Autosubst_Classes.subst (ren (+1)%nat) Ï„2)) -∗ + e2 ≤log≤ e2' : (Autosubst_Classes.subst (ren (+1)) Ï„2)) -∗ {Δ;Γ} ⊨ (unpack: x := e1 in e2) ≤log≤ (unpack: x := e1' in e2') : Ï„2. Proof. iIntros "IH1 IH2". diff --git a/theories/typing/interp.v b/theories/typing/interp.v index 2845ffc..7f52e37 100644 --- a/theories/typing/interp.v +++ b/theories/typing/interp.v @@ -83,7 +83,7 @@ Section interp_ren. (* TODO: why do I need to unfold lrel_car here? *) Lemma interp_ren_up (Δ1 Δ2 : list (lrel Σ)) Ï„ Ï„i : - interp Ï„ (Δ1 ++ Δ2) ≡ interp (Ï„.[upn (length Δ1) (ren (+1)%nat)]) (Δ1 ++ Ï„i :: Δ2). + interp Ï„ (Δ1 ++ Δ2) ≡ interp (Ï„.[upn (length Δ1) (ren (+1))]) (Δ1 ++ Ï„i :: Δ2). Proof. revert Δ1 Δ2. induction Ï„ => Δ1 Δ2; simpl; eauto; try by @@ -96,7 +96,7 @@ Section interp_ren. { by rewrite !lookup_app_l. } change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia..]. - assert ((length Δ1 + S (x - length Δ1) - length Δ1) = S (x - length Δ1))%nat as Hwat. + assert ((length Δ1 + S (x - length Δ1) - length Δ1) = S (x - length Δ1)) as Hwat. { lia. } rewrite Hwat. simpl. done. - intros v1 v2; unfold lrel_car; simpl; @@ -151,7 +151,7 @@ Section interp_ren. rewrite iter_up; case_decide; simpl; properness. { by rewrite !lookup_app_l. } rewrite !lookup_app_r; [|lia..]. - case EQ: (x - length Δ1)%nat => [|n]; simpl. + case EQ: (x - length Δ1)=> [|n]; simpl. { symmetry. pose (HW := interp_weaken [] Δ1 Δ2 Ï„' w1 w2). etrans; last by apply HW. diff --git a/theories/typing/types.v b/theories/typing/types.v index c94c6d6..a3c6ded 100644 --- a/theories/typing/types.v +++ b/theories/typing/types.v @@ -76,7 +76,7 @@ Reserved Notation "⊢ᵥ v : Ï„" (at level 20, v, Ï„ at next level). (* Shift all the indices in the context by one, used when inserting a new type interpretation in Δ. *) -Notation "⤉ Γ" := (Autosubst_Classes.subst (ren (+1)%nat) <$> Γ) (at level 10, format "⤉ Γ"). +Notation "⤉ Γ" := (Autosubst_Classes.subst (ren (+1)) <$> Γ) (at level 10, format "⤉ Γ"). (** We model type-level lambdas and applications as thunks *) Notation "Λ: e" := (λ: <>, e)%E (at level 200, only parsing). @@ -172,7 +172,7 @@ Inductive typed : stringmap type → expr → type → Prop := Γ ⊢ₜ e : (∃: Ï„) | TUnpack Γ e1 x e2 Ï„ Ï„2 : Γ ⊢ₜ e1 : (∃: Ï„) → - <[x:=Ï„]>(⤉ Γ) ⊢ₜ e2 : (Autosubst_Classes.subst (ren (+1%nat)) Ï„2) → + <[x:=Ï„]>(⤉ Γ) ⊢ₜ e2 : (Autosubst_Classes.subst (ren (+1)) Ï„2) → Γ ⊢ₜ (unpack: x := e1 in e2) : Ï„2 | TFork Γ e : Γ ⊢ₜ e : TUnit → Γ ⊢ₜ Fork e : TUnit | TAlloc Γ e Ï„ : Γ ⊢ₜ e : Ï„ → Γ ⊢ₜ Alloc e : Tref Ï„ -- GitLab