--- /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/rt_transition/cpr_drops_basic.ma".
+include "basic_2/dynamic/nta_ind.ma".
+
+lemma lpr_cpr_conf_cpc_sym (h) (G) (L1):
+ ∀T1,T2:term. ⦃G,L1⦄ ⊢ T1 ➡[h] T2 →
+ ∀L2. ⦃G,L1⦄ ⊢ ➡[h] L2 → ⦃G,L2⦄ ⊢ T2 ⬌*[h] T1.
+#h #G #L1 #T1 #T2 #HT12 #L2 #HL12
+elim (lpr_cpr_conf_dx … HT12 … HL12) -L1 #T #HT1 #HT2
+/3 width=3 by cprs_div, cpm_cpms/
+qed.
+
+(* NATIVE TYPE ASSIGNMENT FOR TERMS *****************************************)
+
+definition IH_nta (h) (G) (L1) (T1) ≝
+ ∀U. ⦃G,L1⦄ ⊢ T1 :[h,𝟐] U → ∀T2. ⦃G,L1⦄ ⊢ T1 ➡[h] T2 →
+ ∀L2. ⦃G,L1⦄ ⊢ ➡[h] L2 → ⦃G,L2⦄ ⊢ T2 :[h,𝟐] U.
+
+definition IH_cnv (h) (G) (L1) (T1) ≝
+ (⦃G,L1⦄ ⊢ T1 ![h,𝟐]) → ∀T2. ⦃G,L1⦄ ⊢ T1 ➡[h] T2 →
+ ∀L2. ⦃G,L1⦄ ⊢ ➡[h] L2 → ⦃G,L2⦄ ⊢ T2 ![h,𝟐].
+
+(* Specific proofs for restricted applicability *****************************)
+
+lemma nta_to_cnv (h) (G) (L1) (T1):
+ IH_nta h G L1 T1 → IH_cnv h G L1 T1.
+#h #G #L1 #T1 #H #HT1 #T2 #HT12 #L2 #HL12
+elim (cnv_nta_sn … HT1) -HT1 #U1 #HTU1
+/3 width=2 by nta_fwd_cnv_sn/
+qed-.
+
+(* Note preservation of type without big-tree theorem *)
+lemma nta_cpr_conf_lpr (h) (G) (L1) (T1): IH_nta h G L1 T1.
+#h #G #L1 #T1 @(fqup_wf_ind (Ⓣ) … G L1 T1) -G -L1 -T1
+#G0 #L0 #T0 #IH #U
+@(insert_eq_0 … T0) #T1
+@(insert_eq_0 … L0) #L1
+@(insert_eq_0 … G0) #G
+#H @(nta_ind_rest_cnv … H) -G -L1 -T1 -U
+[ #G #L1 #s #HG0 #HL0 #HT0 #X #HX #L2 #HL12 destruct -IH
+ lapply (cpr_inv_sort1 … HX) -HX #H destruct //
+| #G #K1 #V1 #W1 #U #HVW1 #HW1U #_ #HG0 #HL0 #HT0 #X #HX #Y #HY destruct
+ elim (lpr_inv_pair_sn … HY) -HY #K2 #V2 #HK12 #HV12 #H destruct
+ elim (cpr_inv_zero1 … HX) -HX
+ [ #H destruct
+ /4 width=6 by nta_ldef, fqu_fqup, fqu_lref_O/
+ | * #Y1 #X1 #X2 #HX12 #HX2 #H destruct -HV12
+ @(nta_lifts_bi … (Ⓣ) … HX2 … HW1U) -X -U
+ /4 width=6 by drops_refl, drops_drop, fqu_fqup, fqu_lref_O/
+ ]
+| #G #K1 #W1 #V1 #HW1 #HWV1 #HG0 #HL0 #HT0 #X #HX #Y #HY destruct
+ elim (lpr_inv_pair_sn … HY) -HY #K2 #W2 #HK12 #HW12 #H destruct
+ elim (cpr_inv_zero1 … HX) -HX
+ [ #H destruct
+ lapply (nta_to_cnv … HW1 … W1 … HK12)
+ [1,2: /3 width=1 by fqu_fqup, fqu_lref_O/ ] #H
+ lapply (cnv_lifts … H (Ⓣ) … (K2.ⓛW2) … HWV1) -H
+ [ /3 width=1 by drops_refl, drops_drop/ ] #HV1
+ elim (cpm_lifts_sn … HW12 (Ⓣ) … (K1.ⓛW2) … HWV1) -HWV1
+ [2: /3 width=1 by drops_refl, drops_drop/ ] #V2 #HWV2 #HV12
+ lapply (nta_to_cnv … HW1 … HW12 … HK12) -HW1 -HW12
+ [ /3 width=1 by fqu_fqup, fqu_lref_O/ ] -IH #HW2
+ @(nta_conv_cnv … V2)
+ /3 width=3 by nta_ldec_cnv, lpr_cpr_conf_cpc_sym, lpr_bind_refl_dx/
+ | * #Y1 #X1 #X2 #_ #_ #H destruct
+ ]
+| #I1 #G #K1 #T1 #U1 #i #Hi #HTU1 #_ #HG0 #HL0 #HT0 #X #HX #Y #HY destruct
+ elim (lpr_inv_bind_sn … HY) -HY #I2 #K2 #HK12 #_ #H destruct
+ elim (cpr_inv_lref1 … HX) -HX
+ [ #H destruct /4 width=6 by nta_lref, fqu_fqup/
+ | * #Z1 #Y1 #X1 #HiX1 #HX1 #H destruct
+ @(nta_lifts_bi … (Ⓣ) … HX1 … HTU1) -X -U1
+ /4 width=6 by drops_refl, drops_drop, fqu_fqup/
+ ]
+| #p #I #G #L1 #V1 #T1 #U #HV1 #HT1 #_ #HG0 #HL0 #HT0 #X #HX #L2 #HL12 destruct
+ elim (cpm_inv_bind1 … HX) -HX *
+ [ #V2 #T2 #HV12 #HT12 #H destruct
+ lapply (IH … HT1 T1 … (L2.ⓑ{I}V1) ?)
+ [4:|*: /2 width=1 by lpr_bind_refl_dx/ ] #H
+ lapply (nta_fwd_cnv_dx … H) -H #HU
+ @(nta_conv_cnv … (ⓑ{p,I}V2.U))
+ [ @nta_bind_cnv [ /3 width=6 by nta_to_cnv/ | /4 width=8 by lpr_bind, ext2_pair/ ]
+ | /3 width=3 by lpr_cpr_conf_cpc_sym, cpr_pair_sn/
+ | /4 width=6 by nta_to_cnv, cnv_bind/
+ ]
+ | #U1 #HUT1 #HU1X #H1 #H2 destruct
+(*
+ elim (cpr_subst h G (L1.ⓓV1) U 0)
+ [|*: /3 width=3 by drops_refl/ ] #U0 #T0 #HU0 #HTU0
+ lapply (nta_conv_cnv … HT1 … U0 ??)
+ [ @(cnv_lifts … HTU0)
+
+ @(IH … HU1X … HL12) -L2 -X
+ @(nta_conv_cnv … T0)
+ [
+ | @cpcs_cprs_sn @(cprs_step_sn … (+ⓓV1.U0))
+ /2 width=1 by cpm_bind/ /3 width=3 by cpm_cpms, cpm_zeta/
+
+
+ elim (cpm_lifts_sn … HU1X (Ⓣ) … (L1.ⓓV1) … HUT1) -U1
+ [| /3 width=1 by drops_refl, drops_drop/ ] #T2 #HXT2 #HT12
+*)
--- /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/rt_equivalence/cpcs_cprs.ma".
+include "basic_2/i_dynamic/ntas.ma".
+
+(* ITERATED NATIVE TYPE ASSIGNMENT FOR TERMS ********************************)
+
+(* Properties with r-equivalence for terms **********************************)
+
+lemma ntas_zero (h) (a) (G) (L):
+ ∀T1,T2. ⦃G,L⦄ ⊢ T1 ![h,a] → ⦃G,L⦄ ⊢ T2 ![h,a] → ⦃G,L⦄ ⊢ T1 ⬌*[h] T2 → ⦃G,L⦄ ⊢ T1 :*[h,a,0] T2.
+#h #a #G #L #T1 #T2 #HT1 #HT2 #H
+elim (cpcs_inv_cprs … H) -H #T0 #HT10 #HT20
+/2 width=3 by ntas_intro/
+qed.
+
+(* Inversion lemmas with r-equivalence for terms ****************************)
+
+lemma ntas_inv_zero (h) (a) (G) (L):
+ ∀T1,T2. ⦃G,L⦄ ⊢ T1 :*[h,a,0] T2 →
+ ∧∧ ⦃G,L⦄ ⊢ T1 ![h,a] & ⦃G,L⦄ ⊢ T2 ![h,a] & ⦃G,L⦄ ⊢ T1 ⬌*[h] T2.
+#h #a #G #L #T1 #T2 * #T0 #HT1 #HT2 #HT20 #HT10
+/3 width=3 by cprs_div, and3_intro/
+qed-.
(* ITERATED NATIVE TYPE ASSIGNMENT FOR TERMS ********************************)
+(* Advanced properties of native validity for terms *************************)
+
+lemma cnv_appl_ntas_ge (h) (a) (G) (L):
+ ∀m,n. m ≤ n → ad a n → ∀V,W. ⦃G,L⦄ ⊢ V :[h,a] W →
+ ∀p,T,U. ⦃G,L⦄ ⊢ T :*[h,a,m] ⓛ{p}W.U → ⦃G,L⦄ ⊢ ⓐV.T ![h,a].
+#h #a #G #L #m #n #Hmn #Hn #V #W #HVW #p #T #U #HTU
+elim (cnv_inv_cast … HVW) -HVW #W0 #HW #HV #HW0 #HVW0
+elim HTU -HTU #U0 #H #HT #HU0 #HTU0
+elim (cnv_cpms_conf … H … HU0 0 (ⓛ{p}W0.U)) -H -HU0
+[| /2 width=1 by cpms_bind/ ] -HW0 <minus_n_O #X0 #HUX0 #H
+elim (cpms_inv_abst_sn … H) -H #W1 #U1 #HW01 #_ #H destruct
+/3 width=13 by cnv_appl_ge, cpms_cprs_trans/
+qed-.
+
+(* Advanced inversion lemmas of native validity for terms *******************)
+
+lemma cnv_inv_appl_ntas (h) (a) (G) (L):
+ ∀V,T. ⦃G,L⦄ ⊢ ⓐV.T ![h,a] →
+ ∃∃n,p,W,T,U. ad a n & ⦃G,L⦄ ⊢ V :[h,a] W & ⦃G,L⦄ ⊢ T :*[h,a,n] ⓛ{p}W.U.
+#h #a #G #L #V #T #H
+elim (cnv_inv_appl … H) -H #n #p #W #U #Hn #HV #HT #HVW #HTU
+/3 width=9 by cnv_cpms_nta, cnv_cpms_ntas, ex3_5_intro/
+qed-.
+
(* Properties with native type assignment for terms *************************)
lemma nta_ntas (h) (a) (G) (L):
elim (cnv_inv_cast … H) -H /2 width=3 by ntas_intro/
qed-.
+lemma ntas_step_sn (h) (a) (G) (L):
+ ∀T1,T. ⦃G,L⦄ ⊢ T1 :[h,a] T →
+ ∀n,T2. ⦃G,L⦄ ⊢ T :*[h,a,n] T2 → ⦃G,L⦄ ⊢ T1 :*[h,a,↑n] T2.
+#h #a #G #L #T1 #T #H #n #T2 * #X2 #HT2 #HT #H1TX2 #H2TX2
+elim (cnv_inv_cast … H) -H #X1 #_ #HT1 #H1TX1 #H2TX1
+elim (cnv_cpms_conf … HT … H1TX1 … H2TX2) -T <minus_n_O <minus_O_n <plus_SO_sn #X #HX1 #HX2
+/3 width=5 by ntas_intro, cprs_trans, cpms_trans/
+qed-.
+
+lemma ntas_step_dx (h) (a) (G) (L):
+ ∀n,T1,T. ⦃G,L⦄ ⊢ T1 :*[h,a,n] T →
+ ∀T2. ⦃G,L⦄ ⊢ T :[h,a] T2 → ⦃G,L⦄ ⊢ T1 :*[h,a,↑n] T2.
+#h #a #G #L #n #T1 #T * #X1 #HT #HT1 #H1TX1 #H2TX1 #T2 #H
+elim (cnv_inv_cast … H) -H #X2 #HT2 #_ #H1TX2 #H2TX2
+elim (cnv_cpms_conf … HT … H1TX1 … H2TX2) -T <minus_n_O <minus_O_n <plus_SO_dx #X #HX1 #HX2
+/3 width=5 by ntas_intro, cprs_trans, cpms_trans/
+qed-.
+
+lemma nta_appl_ntas_zero (h) (a) (G) (L): ad a 0 →
+ ∀V,W. ⦃G,L⦄ ⊢ V :[h,a] W → ∀p,T,U0. ⦃G,L⦄ ⊢ T :*[h,a,0] ⓛ{p}W.U0 →
+ ∀U. ⦃G,L.ⓛW⦄ ⊢ U0 :[h,a] U → ⦃G,L⦄ ⊢ ⓐV.T :[h,a] ⓐV.ⓛ{p}W.U.
+#h #a #G #L #Ha #V #W #HVW #p #T #U0 #HTU0 #U #HU0
+lapply (nta_fwd_cnv_dx … HVW) #HW
+lapply (nta_bind_cnv … HW … HU0 p) -HW -HU0 #HU0
+elim (ntas_step_dx … HTU0 … HU0) -HU0 #X #HU #_ #HUX #HTX
+/4 width=9 by cnv_appl_ntas_ge, ntas_refl, cnv_cast, cpms_appl_dx/
+qed.
+
+lemma nta_appl_ntas_pos (h) (a) (n) (G) (L): ad a (↑n) →
+ ∀V,W. ⦃G,L⦄ ⊢ V :[h,a] W → ∀T,U. ⦃G,L⦄ ⊢ T :[h,a] U →
+ ∀p,U0. ⦃G,L⦄ ⊢ U :*[h,a,n] ⓛ{p}W.U0 → ⦃G,L⦄ ⊢ ⓐV.T :[h,a] ⓐV.U.
+#h #a #n #G #L #Ha #V #W #HVW #T #U #HTU #p #U0 #HU0
+elim (cnv_inv_cast … HTU) #X #_ #_ #HUX #HTX
+/4 width=11 by ntas_step_sn, cnv_appl_ntas_ge, cnv_cast, cpms_appl_dx/
+qed-.
+
(* Inversion lemmas with native type assignment for terms *******************)
lemma ntas_inv_nta (h) (a) (G) (L):
/5 width=10 by cnv_cpms_ntas, cnv_cpms_nta, cnv_cpms_trans, ex6_5_intro, or_intror/
]
qed-.
-
-(*
-
-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 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/dynamic/nta_drops.ma".
+include "basic_2/dynamic/nta_cpcs.ma".
+include "basic_2/i_dynamic/ntas_cpcs.ma".
+include "basic_2/i_dynamic/ntas_nta.ma".
+include "basic_2/i_dynamic/ntas_ntas.ma".
+
+(* ITERATED NATIVE TYPE ASSIGNMENT FOR TERMS ********************************)
+
+(* Advanced eliminators for native type assignment **************************)
+
+lemma ntas_ind_bi_nta (h) (a) (G) (L) (Q:relation3 …):
+ (∀T1,T2. ⦃G,L⦄ ⊢ T1 ![h,a] → ⦃G,L⦄ ⊢ T2 ![h,a] → ⦃G,L⦄ ⊢ T1 ⬌*[h] T2 →
+ Q 0 T1 T2
+ ) →
+ (∀T1,T2. ⦃G,L⦄ ⊢ T1 :[h,a] T2 → Q 1 T1 T2
+ ) →
+ (∀n1,n2,T1,T2,T0. ⦃G,L⦄ ⊢ T1 :*[h,a,n1] T0 → ⦃G,L⦄ ⊢ T0 :*[h,a,n2] T2 →
+ Q n1 T1 T0 → Q n2 T0 T2 → Q (n1+n2) T1 T2
+ ) →
+ ∀n,T1,T2. ⦃G,L⦄ ⊢ T1 :*[h,a,n] T2 → Q n T1 T2.
+#h #a #G #L #Q #IH1 #IH2 #IH3 #n
+@(nat_elim1 n) -n * [| * ]
+[ #_ #T1 #T2 #H
+ elim (ntas_inv_zero … H) -H #HT1 #HT2 #HT12
+ /2 width=1 by/
+| #_ #T1 #T2 #H
+ /3 width=1 by ntas_inv_nta/
+| #n #IH #T1 #T2 <plus_SO_sn #H
+ elim (ntas_inv_plus … H) -H #T0 #HT10 #HT02
+ /3 width=5 by/
+]
+qed-.
+
+lemma nta_ind_cnv (h) (a) (Q:relation4 …):
+ (∀G,L,s. Q G L (⋆s) (⋆(⫯[h]s))) →
+ (∀G,K,V,W,U.
+ ⦃G,K⦄ ⊢ V :[h,a] W → ⬆*[1] W ≘ U →
+ Q G K V W → Q G (K.ⓓV) (#0) U
+ ) →
+ (∀G,K,W,U. ⦃G,K⦄ ⊢ W ![h,a] → ⬆*[1] W ≘ U → Q G (K.ⓛW) (#0) U) →
+ (∀I,G,K,W,U,i.
+ ⦃G,K⦄ ⊢ #i :[h,a] W → ⬆*[1] W ≘ U →
+ Q G K (#i) W → Q G (K.ⓘ{I}) (#↑i) U
+ ) →
+ (∀p,I,G,K,V,T,U.
+ ⦃G,K⦄ ⊢ V ![h,a] → ⦃G,K.ⓑ{I}V⦄ ⊢ T :[h,a] U →
+ Q G (K.ⓑ{I}V) T U → Q G K (ⓑ{p,I}V.T) (ⓑ{p,I}V.U)
+ ) →
+ (∀p,G,L,V,W,T,U0,U. (**) (* one IH is missing *)
+ ad a 0 → ⦃G,L⦄ ⊢ V :[h,a] W → ⦃G,L⦄ ⊢ T :*[h,a,0] ⓛ{p}W.U0 → ⦃G,L.ⓛW⦄ ⊢ U0 :[h,a] U →
+ Q G L V W (* → Q G (L.ⓛW) U0 U *) → Q G L (ⓐV.T) (ⓐV.ⓛ{p}W.U)
+ ) →
+ (∀n,p,G,L,V,W,T,U,U0.
+ ad a (↑n) → ⦃G,L⦄ ⊢ V :[h,a] W → ⦃G,L⦄ ⊢ T :[h,a] U → ⦃G,L⦄ ⊢ U :*[h,a,n] ⓛ{p}W.U0 →
+ Q G L V W → Q G L T U → Q G L (ⓐV.T) (ⓐV.U)
+ ) →
+ (∀G,L,T,U. ⦃G,L⦄ ⊢ T :[h,a] U → Q G L T U → Q G L (ⓝU.T) U
+ ) →
+ (∀G,L,T,U1,U2.
+ ⦃G,L⦄ ⊢ T :[h,a] U1 → ⦃G,L⦄ ⊢ U1 ⬌*[h] U2 → ⦃G,L⦄ ⊢ U2 ![h,a] →
+ Q G L T U1 → Q G L T U2
+ ) →
+ ∀G,L,T,U. ⦃G,L⦄ ⊢ T :[h,a] U → Q G L T U.
+#h #a #Q #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #IH8 #H9 #G #L #T
+@(fqup_wf_ind_eq (Ⓣ) … G L T) -G -L -T #G0 #L0 #T0 #IH #G #L * * [|||| * ]
+[ #s #HG #HL #HT #X #H destruct -IH
+ elim (nta_inv_sort_sn … H) -H #HUX #HX
+ /2 width=4 by/
+| * [| #i ] #HG #HL #HT #X #H destruct
+ [ elim (nta_inv_lref_sn_drops_cnv … H) -H *
+ [ #K #V #W #U #H #HVW #HWU #HUX #HX
+ lapply (drops_fwd_isid … H ?) -H [ // ] #H destruct
+ /5 width=7 by nta_ldef, fqu_fqup, fqu_lref_O/
+ | #K #W #U #H #HW #HWU #HUX #HX
+ lapply (drops_fwd_isid … H ?) -H [ // ] #H destruct
+ /3 width=4 by nta_ldec_cnv/
+ ]
+ | elim (nta_inv_lref_sn … H) -H #I #K #T #U #HT #HTU #HUX #HX #H destruct
+ /5 width=7 by nta_lref, fqu_fqup/
+ ]
+| #l #HG #HL #HT #U #H destruct -IH
+ elim (nta_inv_gref_sn … H)
+| #p #I #V #T #HG #HL #HT #X #H destruct
+ elim (nta_inv_bind_sn_cnv … H) -H #U #HV #HTU #HUX #HX
+ /4 width=5 by nta_bind_cnv/
+| #V #T #HG #HL #HT #X #H destruct
+ elim (nta_inv_appl_sn_ntas … H) -H *
+ [ #p #W #U #U0 #Ha #HVW #HTU0 #HU0 #HUX #HX -IH7
+ /4 width=10 by nta_appl_ntas_zero/
+ | #n #p #W #U #U0 #Ha #HVW #HTU #HU0 #HUX #HX -IH6
+ /4 width=13 by nta_appl_ntas_pos/
+ ]
+| #U #T #HG #HL #HT #X #H destruct
+ elim (nta_inv_cast_sn … H) -H #HTU #HUX #HX
+ /4 width=4 by nta_cast/
+]
+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/i_dynamic/ntas_preserve.ma".
+
+(* ITERATED NATIVE TYPE ASSIGNMENT FOR TERMS ********************************)
+
+(* Main properties **********************************************************)
+
+theorem ntas_trans (h) (a) (G) (L) (T0):
+ ∀n1,T1. ⦃G,L⦄ ⊢ T1 :*[h,a,n1] T0 →
+ ∀n2,T2. ⦃G,L⦄ ⊢ T0 :*[h,a,n2] T2 → ⦃G,L⦄ ⊢ T1 :*[h,a,n1+n2] T2.
+#h #a #G #L #T0 #n1 #T1 * #X1 #HT0 #HT1 #H01 #H11 #n2 #T2 * #X2 #HT2 #_ #H22 #H02
+elim (cnv_cpms_conf … HT0 … H01 … H02) -T0 <minus_O_n <minus_n_O #X0 #H10 #H20
+/3 width=5 by ntas_intro, cpms_trans/
+qed-.
(* Inversion lemmas based on preservation ***********************************)
+lemma ntas_inv_plus (h) (a) (n1) (n2) (G) (L):
+ ∀T1,T2. ⦃G,L⦄ ⊢ T1 :*[h,a,n1+n2] T2 →
+ ∃∃T0. ⦃G,L⦄ ⊢ T1 :*[h,a,n1] T0 & ⦃G,L⦄ ⊢ T0 :*[h,a,n2] T2.
+#h #a #n1 #n2 #G #L #T1 #T2 * #X0 #HT2 #HT1 #H20 #H10
+elim (cpms_inv_plus … H10) -H10 #T0 #H10 #H00
+lapply (cnv_cpms_trans … HT1 … H10) #HT0
+/3 width=6 by cnv_cpms_ntas, ntas_intro, ex2_intro/
+qed-.
+
lemma ntas_inv_appl_sn (h) (a) (m) (G) (L) (V) (T):
∀X. ⦃G,L⦄ ⊢ ⓐV.T :*[h,a,m] X →
∨∨ ∃∃n,p,W,U,U0. n ≤ m & ad a n & ⦃G,L⦄ ⊢ V :*[h,a,1] W & ⦃G,L⦄ ⊢ T :*[h,a,n] ⓛ{p}W.U0 & ⦃G,L.ⓛW⦄ ⊢ U0 :*[h,a,m-n] U & ⦃G,L⦄ ⊢ ⓐV.ⓛ{p}W.U ⬌*[h] X & ⦃G,L⦄ ⊢ X ![h,a]
/5 width=11 by cnv_cpms_ntas, cnv_cpms_trans, ex7_5_intro, or_intror/
]
qed-.
-
-(*
-(* 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/
-*)
-*)
lemma cpms_sort (h) (G) (L) (n):
∀s. ⦃G,L⦄ ⊢ ⋆s ➡*[n,h] ⋆((next h)^n s).
#h #G #L #n elim n -n [ // ]
-#n #IH #s <plus_SO
+#n #IH #s <plus_SO_dx
/3 width=3 by cpms_step_dx, cpm_sort/
qed.
∀T. ⦃G,L⦄ ⊢ T ⁝ A → ∃U. ⦃G,L⦄ ⊢ T ➡*[n,h] U.
#h #G #L #n elim n -n
[ /2 width=3 by ex_intro/
-| #n #IH #A #T1 #HT1 <plus_SO
+| #n #IH #A #T1 #HT1 <plus_SO_dx
elim (IH … HT1) -IH #T0 #HT10
lapply (cpms_aaa_conf … HT1 … HT10) -HT1 #HT0
elim (aaa_cpm_SO h … HT0) -HT0 #T2 #HT02
class "wine"
[ { "iterated dynamic typing" * } {
[ { "context-sensitive iterated native type assignment" * } {
- [ [ "for terms" ] "ntas" + "( ⦃?,?⦄ ⊢ ? :*[?,?,?] ? )" "ntas_nta" + "ntas_preserve" * ]
+ [ [ "for terms" ] "ntas" + "( ⦃?,?⦄ ⊢ ? :*[?,?,?] ? )" "ntas_cpcs" + "ntas_nta" + "ntas_nta_ind" + " ntas_ntas" + "ntas_preserve" * ]
}
]
}
(* Equalities ***************************************************************)
-lemma plus_SO: ∀n. n + 1 = ↑n.
+lemma plus_SO_sn (n): 1 + n = ↑n.
+// qed-.
+
+lemma plus_SO_dx (n): n + 1 = ↑n.
// qed.
lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.