(* EXAMPLES *****************************************************************)
-(* Extended validy (basic?_2) vs. restricted validity (basic_1) *************)
+(* Extended validy (basic_2B) vs. restricted validity (basic_1A) ************)
(* Note: extended validity of a closure, height of cnv_appl > 1 *)
lemma cnv_extended (h) (p): ∀G,L,s. ⦃G, L.ⓛ⋆s.ⓛⓛ{p}⋆s.⋆s.ⓛ#0⦄ ⊢ ⓐ#2.#0 ![Ⓕ,h].
</body>
<table name="apps_2_sum"/>
<news class="alpha" date="2018 April 30.">
- The Models component is started for λδ version 2.
+ The Models component is started for λδ-2B.
</news>
<news class="alpha" date="2017 March 6.">
The Examples component is moved from the Core directory.
</news>
<news class="alpha" date="2011 December 20.">
The Functional component is started
- inside the specification of λδ version 2.
+ inside the specification of λδ-2A.
</news>
<news class="alpha" date="2011 December 12.">
The MLTT1 component is started.
lapply (cpms_inv_sort1 … H) -H #H destruct
/2 width=1 by cpcs_cprs_sn/
qed-.
-
/4 width=2 by cnv_cpms_nta, cpm_cpms, ex_intro/
qed-.
+(* Basic_1: was: ty3_typecheck *)
+lemma nta_typecheck (a) (h) (G) (L):
+ ∀T,U. ⦃G,L⦄ ⊢ T :[a,h] U → ∃T0. ⦃G,L⦄ ⊢ ⓝU.T :[a,h] T0.
+/3 width=1 by cnv_cast, cnv_nta_sn/ qed-.
+
+(* Basic_1: was: ty3_correct *)
+(* Basic_2A1: was: ntaa_fwd_correct *)
+lemma nta_fwd_correct (a) (h) (G) (L):
+ ∀T,U. ⦃G,L⦄ ⊢ T :[a,h] U → ∃T0. ⦃G,L⦄ ⊢ U :[a,h] T0.
+/3 width=2 by nta_fwd_cnv_dx, cnv_nta_sn/ qed-.
+
lemma nta_pure_cnv (h) (G) (L):
∀T,U. ⦃G,L⦄ ⊢ T :*[h] U →
∀V. ⦃G,L⦄ ⊢ ⓐV.U !*[h] → ⦃G,L⦄ ⊢ ⓐV.T :*[h] ⓐV.U.
]
qed-.
+(* Basic_1: uses: ty3_gen_cast *)
+lemma nta_inv_cast_sn_old (a) (h) (G) (L) (X2):
+ ∀T0,T1. ⦃G,L⦄ ⊢ ⓝT1.T0 :[a,h] X2 →
+ ∃∃T2. ⦃G,L⦄ ⊢ T0 :[a,h] T1 & ⦃G,L⦄ ⊢ T1 :[a,h] T2 & ⦃G,L⦄ ⊢ ⓝT2.T1 ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 ![a,h].
+#a #h #G #L #X2 #T0 #T1 #H
+elim (cnv_inv_cast … H) -H #X0 #HX2 #H1 #HX20 #H2
+elim (cnv_inv_cast … H1) #X #HT1 #HT0 #HT1X #HT0X
+elim (cpms_inv_cast1 … H2) -H2 [ * || * ]
+[ #U1 #U0 #HTU1 #HTU0 #H destruct
+ elim (cnv_cpms_conf … HT0 … HT0X … HTU0) -HT0 -HT0X -HTU0
+ <minus_n_n #X0 #HX0 #HUX0
+ lapply (cprs_trans … HT1X … HX0) -X #HT1X0
+ /5 width=7 by cnv_cpms_nta, cpcs_cprs_div, cprs_div, cpms_cast, ex4_intro/
+| #HTX0
+ elim (cnv_cpms_conf … HT0 … HT0X … HTX0) -HT0 -HT0X -HTX0
+ <minus_n_n #X1 #HX1 #HX01
+ elim (cnv_nta_sn … HT1) -HT1 #U1 #HTU1
+ lapply (cprs_trans … HT1X … HX1) -X #HTX1
+ lapply (cprs_trans … HX20 … HX01) -X0 #HX21
+ /4 width=5 by cprs_div, cpms_eps, ex4_intro/
+| #n #HT1X0 #H destruct -X -HT0
+ elim (cnv_nta_sn … HT1) -HT1 #U1 #HTU1
+ /4 width=5 by cprs_div, cpms_eps, ex4_intro/
+]
+qed-.
+
(* Forward lemmas based on preservation *************************************)
(* Basic_1: was: ty3_unique *)
(* Basic_2A1: was: nta_inv_bind1 ntaa_inv_bind1 *)
lemma nta_inv_bind_sn
-| ntaa_cast: ∀L,T,U,W. ntaa h L T U → ntaa h L U W → ntaa h L (ⓝU. T) U
-
(* Advanced properties ******************************************************)
+| ntaa_cast: ∀L,T,U,W. ntaa h L T U → ntaa h L U W → ntaa h L (ⓝU. T) U
+
lemma nta_cast_alt: ∀h,L,T,W,U. ⦃h, L⦄ ⊢ T : W → ⦃h, L⦄ ⊢ T : U → ⦃h, L⦄ ⊢ ⓝW.T : U.
#h #L #T #W #U #HTW #HTU
lapply (nta_mono … HTW … HTU) #HWU
elim (nta_fwd_correct … HTU) -HTU /3 width=3/
qed.
-
-lemma nta_inv_cast1: ∀h,L,X,Y,U. ⦃h, L⦄ ⊢ ⓝY.X : U → ⦃h, L⦄ ⊢ X : Y ∧ L ⊢ Y ⬌* U.
-/2 width=3/ qed-.
-
-(* Basic_1: uses: ty3_gen_cast *)
-lemma nta_inv_cast_sn_old (a) (h) (G) (L) (X2):
- ∀T0,T1. ⦃G,L⦄ ⊢ ⓝT1.T0 :[a,h] X2 →
- ∃∃T2. ⦃G,L⦄ ⊢ T0 :[a,h] T1 & ⦃G,L⦄ ⊢ T1 :[a,h] T2 & ⦃G,L⦄ ⊢ ⓝT2.T1 ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 ![a,h].
-
-(* Basic_1: was: ty3_typecheck *)
-lemma nta_typecheck: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∃T0. ⦃h, L⦄ ⊢ ⓝU.T : T0.
-/3 width=2/ qed.
-
-(* Basic_1: was: ty3_correct *)
-(* Basic_2A1: was: ntaa_fwd_correct *)
-lemma nta_fwd_correct: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∃T0. ⦃h, L⦄ ⊢ U : T0.
-#h #L #T #U #H
-elim (ntaa_fwd_correct … (nta_ntaa … H)) -H /3 width=2 by ntaa_nta, ex_intro/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-notation "hvbox( ⦃ h , break L ⦄ ⊢ break term 46 T1 : : * break term 46 T2 )"
- non associative with precedence 45
- for @{ 'NativeTypeStarAlt $h $L $T1 $T2 }.
-
-include "basic_2/dynamic/nta.ma".
-
-(* HIGHER ORDER NATIVE TYPE ASSIGNMENT ON TERMS *****************************)
-
-definition ntas: sh → lenv → relation term ≝
- λh,L. star … (nta h L).
-
-interpretation "higher order native type assignment (term)"
- 'NativeTypeStar h L T U = (ntas h L T U).
-
-(* Basic eliminators ********************************************************)
-(*
-lemma cprs_ind: ∀L,T1. ∀R:predicate term. R T1 →
- (∀T,T2. L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → R T → R T2) →
- ∀T2. L ⊢ T1 ➡* T2 → R T2.
-#L #T1 #R #HT1 #IHT1 #T2 #HT12
-@(TC_star_ind … HT1 IHT1 … HT12) //
-qed-.
-*)
-axiom ntas_ind_dx: ∀h,L,T2. ∀R:predicate term. R T2 →
- (∀T1,T. ⦃h, L⦄ ⊢ T1 : T → ⦃h, L⦄ ⊢ T :* T2 → R T → R T1) →
- ∀T1. ⦃h, L⦄ ⊢ T1 :* T2 → R T1.
-(*
-#h #L #T2 #R #HT2 #IHT2 #T1 #HT12
-@(star_ind_dx … HT2 IHT2 … HT12) //
-qed-.
-*)
-(* Basic properties *********************************************************)
-
-lemma ntas_refl: ∀h,L,T. ⦃h, L⦄ ⊢ T :* T.
-// qed.
-
-lemma ntas_strap1: ∀h,L,T1,T,T2.
- ⦃h, L⦄ ⊢ T1 :* T → ⦃h, L⦄ ⊢ T : T2 → ⦃h, L⦄ ⊢ T1 :* T2.
-/2 width=3/ qed.
-
-lemma ntas_strap2: ∀h,L,T1,T,T2.
- ⦃h, L⦄ ⊢ T1 : T → ⦃h, L⦄ ⊢ T :* T2 → ⦃h, L⦄ ⊢ T1 :* T2.
-/2 width=3/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/dynamic/nta_lift.ma".
-include "basic_2/hod/ntas.ma".
-
-(* HIGHER ORDER NATIVE TYPE ASSIGNMENT ON TERMS *****************************)
-
-(* Advanced properties on native type assignment for terms ******************)
-
-lemma nta_pure_ntas: ∀h,L,U,W,Y. ⦃h, L⦄ ⊢ U :* ⓛW.Y → ∀T. ⦃h, L⦄ ⊢ T : U →
- ∀V. ⦃h, L⦄ ⊢ V : W → ⦃h, L⦄ ⊢ ⓐV.T : ⓐV.U.
-#h #L #U #W #Y #H @(ntas_ind_dx … H) -U /2 width=1/ /3 width=2/
-qed.
-
-axiom pippo: ∀h,L,T,W,Y. ⦃h, L⦄ ⊢ T :* ⓛW.Y → ∀U. ⦃h, L⦄ ⊢ T : U →
- ∃Z. ⦃h, L⦄ ⊢ U :* ⓛW.Z.
-(* REQUIRES SUBJECT CONVERSION
-#h #L #T #W #Y #H @(ntas_ind_dx … H) -T
-[ #U #HYU
- elim (nta_fwd_correct … HYU) #U0 #HU0
- elim (nta_inv_bind1 … HYU) #W0 #Y0 #HW0 #HY0 #HY0U
-*)
-
-(* Advanced inversion lemmas on native type assignment for terms ************)
-
-fact nta_inv_pure1_aux: ∀h,L,Z,U. ⦃h, L⦄ ⊢ Z : U → ∀X,Y. Z = ⓐY.X →
- ∃∃W,V,T. ⦃h, L⦄ ⊢ Y : W & ⦃h, L⦄ ⊢ X : V &
- L ⊢ ⓐY.V ⬌* U & ⦃h, L⦄ ⊢ V :* ⓛW.T.
-#h #L #Z #U #H elim H -L -Z -U
-[ #L #k #X #Y #H destruct
-| #L #K #V #W #U #i #_ #_ #_ #_ #X #Y #H destruct
-| #L #K #W #V #U #i #_ #_ #_ #_ #X #Y #H destruct
-| #I #L #V #W #T #U #_ #_ #_ #_ #X #Y #H destruct
-| #L #V #W #Z #U #HVW #HZU #_ #_ #X #Y #H destruct /2 width=7/
-| #L #V #W #Z #U #HZU #_ #_ #IHUW #X #Y #H destruct
- elim (IHUW U Y ?) -IHUW // /3 width=9/
-| #L #Z #U #_ #_ #X #Y #H destruct
-| #L #Z #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #X #Y #H destruct
- elim (IHTU1 ???) -IHTU1 [4: // |2,3: skip ] #W #V #T #HYW #HXV #HU1 #HVT
- lapply (cpcs_trans … HU1 … HU12) -U1 /2 width=7/
-]
-qed.
-
-(* Basic_1: was only: ty3_gen_appl *)
-lemma nta_inv_pure1: ∀h,L,Y,X,U. ⦃h, L⦄ ⊢ ⓐY.X : U →
- ∃∃W,V,T. ⦃h, L⦄ ⊢ Y : W & ⦃h, L⦄ ⊢ X : V &
- L ⊢ ⓐY.V ⬌* U & ⦃h, L⦄ ⊢ V :* ⓛW.T.
-/2 width=3/ qed-.
-
-axiom nta_inv_appl1: ∀h,L,Z,Y,X,U. ⦃h, L⦄ ⊢ ⓐZ.ⓛY.X : U →
- ∃∃W. ⦃h, L⦄ ⊢ Z : Y & ⦃h, L⦄ ⊢ ⓛY.X : ⓛY.W &
- L ⊢ ⓐZ.ⓛY.W ⬌* U.
-(* REQUIRES SUBJECT REDUCTION
-#h #L #Z #Y #X #U #H
-elim (nta_inv_pure1 … H) -H #W #V #T #HZW #HXV #HVU #HVT
-elim (nta_inv_bind1 … HXV) -HXV #Y0 #X0 #HY0 #HX0 #HX0V
-lapply (cpcs_trans … (ⓐZ.ⓛY.X0) … HVU) -HVU /2 width=1/ -HX0V #HX0U
-@(ex3_1_intro … HX0U) /2 width=2/
-*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/colon_7.ma".
+include "basic_2/notation/relations/colon_6.ma".
+include "basic_2/notation/relations/colonstar_6.ma".
+include "basic_2/dynamic/cnv.ma".
+
+(* ITERATED NATIVE TYPE ASSIGNMENT FOR TERMS ********************************)
+
+definition ntas (a) (h) (n) (G) (L): relation term ≝ λT,U.
+ ∃∃U0. ⦃G,L⦄ ⊢ U ![a,h] & ⦃G,L⦄ ⊢ T ![a,h] & ⦃G,L⦄ ⊢ U ➡*[h] U0 & ⦃G, L⦄ ⊢ T ➡*[n,h] U0.
+
+interpretation "iterated native type assignment (term)"
+ 'Colon a h n G L T U = (ntas a h n G L T U).
+
+interpretation "restricted iterated native type assignment (term)"
+ 'Colon h n G L T U = (ntas true h n G L T U).
+
+interpretation "extended iterated native type assignment (term)"
+ 'ColonStar h n G L T U = (ntas false h n G L T U).
+
+(* Basic properties *********************************************************)
+
+lemma ntas_refl (a) (h) (G) (L):
+ ∀T. ⦃G,L⦄ ⊢ T ![a,h] → ⦃G,L⦄ ⊢ T :[a,h,0] T.
+/2 width=3 by ex4_intro/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+(*
+include "basic_2/dynamic/nta_lift.ma".
+include "basic_2/hod/ntas.ma".
+
+(* HIGHER ORDER NATIVE TYPE ASSIGNMENT ON TERMS *****************************)
+
+(* Advanced properties on native type assignment for terms ******************)
+
+lemma nta_pure_ntas: ∀h,L,U,W,Y. ⦃h, L⦄ ⊢ U :* ⓛW.Y → ∀T. ⦃h, L⦄ ⊢ T : U →
+ ∀V. ⦃h, L⦄ ⊢ V : W → ⦃h, L⦄ ⊢ ⓐV.T : ⓐV.U.
+#h #L #U #W #Y #H @(ntas_ind_dx … H) -U /2 width=1/ /3 width=2/
+qed.
+
+axiom pippo: ∀h,L,T,W,Y. ⦃h, L⦄ ⊢ T :* ⓛW.Y → ∀U. ⦃h, L⦄ ⊢ T : U →
+ ∃Z. ⦃h, L⦄ ⊢ U :* ⓛW.Z.
+(* REQUIRES SUBJECT CONVERSION
+#h #L #T #W #Y #H @(ntas_ind_dx … H) -T
+[ #U #HYU
+ elim (nta_fwd_correct … HYU) #U0 #HU0
+ elim (nta_inv_bind1 … HYU) #W0 #Y0 #HW0 #HY0 #HY0U
+*)
+
+(* Advanced inversion lemmas on native type assignment for terms ************)
+
+fact nta_inv_pure1_aux: ∀h,L,Z,U. ⦃h, L⦄ ⊢ Z : U → ∀X,Y. Z = ⓐY.X →
+ ∃∃W,V,T. ⦃h, L⦄ ⊢ Y : W & ⦃h, L⦄ ⊢ X : V &
+ L ⊢ ⓐY.V ⬌* U & ⦃h, L⦄ ⊢ V :* ⓛW.T.
+#h #L #Z #U #H elim H -L -Z -U
+[ #L #k #X #Y #H destruct
+| #L #K #V #W #U #i #_ #_ #_ #_ #X #Y #H destruct
+| #L #K #W #V #U #i #_ #_ #_ #_ #X #Y #H destruct
+| #I #L #V #W #T #U #_ #_ #_ #_ #X #Y #H destruct
+| #L #V #W #Z #U #HVW #HZU #_ #_ #X #Y #H destruct /2 width=7/
+| #L #V #W #Z #U #HZU #_ #_ #IHUW #X #Y #H destruct
+ elim (IHUW U Y ?) -IHUW // /3 width=9/
+| #L #Z #U #_ #_ #X #Y #H destruct
+| #L #Z #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #X #Y #H destruct
+ elim (IHTU1 ???) -IHTU1 [4: // |2,3: skip ] #W #V #T #HYW #HXV #HU1 #HVT
+ lapply (cpcs_trans … HU1 … HU12) -U1 /2 width=7/
+]
+qed.
+
+(* Basic_1: was only: ty3_gen_appl *)
+lemma nta_inv_pure1: ∀h,L,Y,X,U. ⦃h, L⦄ ⊢ ⓐY.X : U →
+ ∃∃W,V,T. ⦃h, L⦄ ⊢ Y : W & ⦃h, L⦄ ⊢ X : V &
+ L ⊢ ⓐY.V ⬌* U & ⦃h, L⦄ ⊢ V :* ⓛW.T.
+/2 width=3/ qed-.
+
+axiom nta_inv_appl1: ∀h,L,Z,Y,X,U. ⦃h, L⦄ ⊢ ⓐZ.ⓛY.X : U →
+ ∃∃W. ⦃h, L⦄ ⊢ Z : Y & ⦃h, L⦄ ⊢ ⓛY.X : ⓛY.W &
+ L ⊢ ⓐZ.ⓛY.W ⬌* U.
+(* REQUIRES SUBJECT REDUCTION
+#h #L #Z #Y #X #U #H
+elim (nta_inv_pure1 … H) -H #W #V #T #HZW #HXV #HVU #HVT
+elim (nta_inv_bind1 … HXV) -HXV #Y0 #X0 #HY0 #HX0 #HX0V
+lapply (cpcs_trans … (ⓐZ.ⓛY.X0) … HVU) -HVU /2 width=1/ -HX0V #HX0U
+@(ex3_1_intro … HX0U) /2 width=2/
+*)
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/dynamic/nta.ma".
+include "basic_2/i_dynamic/ntas.ma".
+
+(* ITERATED NATIVE TYPE ASSIGNMENT FOR TERMS ********************************)
+
+(*
+
+definition ntas: sh → lenv → relation term ≝
+ λh,L. star … (nta h L).
+
+(* Basic eliminators ********************************************************)
+
+axiom ntas_ind_dx: ∀h,L,T2. ∀R:predicate term. R T2 →
+ (∀T1,T. ⦃h, L⦄ ⊢ T1 : T → ⦃h, L⦄ ⊢ T :* T2 → R T → R T1) →
+ ∀T1. ⦃h, L⦄ ⊢ T1 :* T2 → R T1.
+(*
+#h #L #T2 #R #HT2 #IHT2 #T1 #HT12
+@(star_ind_dx … HT2 IHT2 … HT12) //
+qed-.
+*)
+(* Basic properties *********************************************************)
+
+lemma ntas_strap1: ∀h,L,T1,T,T2.
+ ⦃h, L⦄ ⊢ T1 :* T → ⦃h, L⦄ ⊢ T : T2 → ⦃h, L⦄ ⊢ T1 :* T2.
+/2 width=3/ qed.
+
+lemma ntas_strap2: ∀h,L,T1,T,T2.
+ ⦃h, L⦄ ⊢ T1 : T → ⦃h, L⦄ ⊢ T :* T2 → ⦃h, L⦄ ⊢ T1 :* T2.
+/2 width=3/ qed.
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 :[ break term 46 a, break term 46 h, break term 46 n ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'Colon $a $h $n $G $L $T1 $T2 }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 :*[ break term 46 h, break term 46 n ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'ColonStar $h $n $G $L $T1 $T2 }.
]
}
]
-(*
class "wine"
[ { "iterated dynamic typing" * } {
- [ { "" (* "higher order native type assignment" *) * } {
- [ [ "" ] (* "ntas ( ⦃?,?⦄ ⊢ ? :* ? )" "nta_lift" *) * ]
+ [ { "context-sensitive iterated native type assignment" * } {
+ [ [ "for terms" ] "ntas" + "( ⦃?,?⦄ ⊢ ? :[?,?,?] ? )" + "( ⦃?,?⦄ ⊢ ? :[?,?] ? )" + "( ⦃?,?⦄ ⊢ ? :*[?,?] ? )" * ]
}
]
}
]
-*)
class "magenta"
[ { "dynamic typing" * } {
[ { "context-sensitive native type assignment" * } {