@(cpms_ind_dx … H) -n -T2 /3 width=6 by cnv_cpm_trans_lpr/
qed-.
+lemma cnv_cpm_trans (a) (h) (G) (L):
+ ∀T1. ⦃G,L⦄ ⊢ T1 ![a,h] →
+ ∀n,T2. ⦃G,L⦄ ⊢ T1 ➡[n,h] T2 → ⦃G,L⦄ ⊢ T2 ![a,h].
+/2 width=6 by cnv_cpm_trans_lpr/ qed-.
+
+(* Note: this is the preservation property *)
lemma cnv_cpms_trans (a) (h) (G) (L):
∀T1. ⦃G,L⦄ ⊢ T1 ![a,h] →
∀n,T2. ⦃G,L⦄ ⊢ T1 ➡*[n,h] T2 → ⦃G,L⦄ ⊢ T2 ![a,h].
/2 width=6 by cnv_cpms_trans_lpr/ qed-.
+
+lemma cnv_lpr_trans (a) (h) (G):
+ ∀L1,T. ⦃G,L1⦄ ⊢ T ![a,h] → ∀L2. ⦃G,L1⦄ ⊢ ➡[h] L2 → ⦃G,L2⦄ ⊢ T ![a,h].
+/2 width=6 by cnv_cpm_trans_lpr/ qed-.
+
+lemma cnv_lprs_trans (a) (h) (G):
+ ∀L1,T. ⦃G,L1⦄ ⊢ T ![a,h] → ∀L2. ⦃G,L1⦄ ⊢ ➡*[h] L2 → ⦃G,L2⦄ ⊢ T ![a,h].
+#a #h #G #L1 #T #HT #L2 #H
+@(lprs_ind_dx … H) -L2 /2 width=3 by cnv_lpr_trans/
+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/cnv_aaa.ma".
+include "basic_2/dynamic/nta.ma".
+
+(* NATIVE TYPE ASSIGNMENT FOR TERMS *****************************************)
+
+(* Forward lemmas with atomic arity assignment for terms ********************)
+
+(* Note: this means that no type is a universe *)
+lemma nta_fwd_aaa (a) (h) (G) (L):
+ ∀T,U. ⦃G,L⦄ ⊢ T :[a,h] U → ∃∃A. ⦃G,L⦄ ⊢ T ⁝ A & ⦃G,L⦄ ⊢ U ⁝ A.
+#a #h #G #L #T #U #H
+elim (cnv_fwd_aaa … H) -H #A #H
+elim (aaa_inv_cast … H) -H #HU #HT
+/2 width=3 by ex2_intro/
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+(* Basic_1: uses: ty3_predicative *)
+lemma nta_abst_predicative (a) (h) (p) (G) (L):
+ ∀W,T. ⦃G,L⦄ ⊢ ⓛ{p}W.T :[a,h] W → ⊥.
+#a #h #p #G #L #W #T #H
+elim (nta_fwd_aaa … H) -a -h #X #H #H1W
+elim (aaa_inv_abst … H) -p #B #A #H2W #_ #H destruct -T
+lapply (aaa_mono … H1W … H2W) -G -L -W #H
+elim (discr_apair_xy_x … H)
+qed-.
+
+(* Basic_2A1: uses: ty3_repellent *)
+theorem nta_abst_repellent (a) (h) (p) (G) (K):
+ ∀W,T,U1. ⦃G,K⦄ ⊢ ⓛ{p}W.T :[a,h] U1 →
+ ∀U2. ⦃G,K.ⓛW⦄ ⊢ T :[a,h] U2 → ⬆*[1] U1 ≘ U2 → ⊥.
+#a #h #p #G #K #W #T #U1 #H1 #U2 #H2 #HU12
+elim (nta_fwd_aaa … H2) -H2 #A2 #H2T #H2U2
+elim (nta_fwd_aaa … H1) -H1 #X1 #H1 #HU1
+elim (aaa_inv_abst … H1) -a -h -p #B #A1 #_ #H1T #H destruct
+lapply (aaa_mono … H1T … H2T) -T #H destruct
+lapply (aaa_inv_lifts … H2U2 (Ⓣ) … K … HU12)
+[ /3 width=1 by drops_refl, drops_drop/ ] -W -U2 #H2U1
+lapply (aaa_mono … HU1 … H2U1) -G -K -U1 #H
+elim (discr_apair_xy_y … H)
+qed-.
(* NATIVE TYPE ASSIGNMENT FOR TERMS *****************************************)
-(* Properties with rt_computation for terms *********************************)
+(* Properties with advanced rt_computation for terms ************************)
(* Basic_2A1: was by definition nta_appl ntaa_appl *)
-lemma nta_beta (a) (h) (p) (G) (K):
+lemma nta_appl_abst (a) (h) (p) (G) (K):
∀V,W. ⦃G,K⦄ ⊢ V :[a,h] W →
∀T,U. ⦃G,K.ⓛW⦄ ⊢ T :[a,h] U → ⦃G,K⦄ ⊢ ⓐV.ⓛ{p}W.T :[a,h] ⓐV.ⓛ{p}W.U.
#a #h #p #G #K #V #W #H1 #T #U #H2
| /2 width=1 by cpms_appl_dx/
]
qed.
+
+(* Inversion lemmas with advanced rt_computation for terms ******************)
+
+lemma nta_inv_abst_bi_cnv (a) (h) (p) (G) (K) (W):
+ ∀T,U. ⦃G,K⦄ ⊢ ⓛ{p}W.T :[a,h] ⓛ{p}W.U →
+ ∧∧ ⦃G,K⦄ ⊢ W ![a,h] & ⦃G,K.ⓛW⦄ ⊢ T :[a,h] U.
+#a #h #p #G #K #W #T #U #H
+elim (cnv_inv_cast … H) -H #X #HWU #HWT #HUX #HTX
+elim (cnv_inv_bind … HWU) -HWU #HW #HU
+elim (cnv_inv_bind … HWT) -HWT #_ #HT
+elim (cpms_inv_abst_sn … HUX) -HUX #W0 #X0 #_ #HUX0 #H destruct
+elim (cpms_inv_abst_bi … HTX) -HTX #_ #HTX0 -W0
+/3 width=3 by cnv_cast, conj/
+qed-.
/3 width=12 by nta_lifts_sn, d_liftable2_sn_bi, lifts_mono/ qed-.
(* Basic_1: was by definition: ty3_abbr *)
-(* Basic_2A1: was by definition: nta_ldef *)
+(* Basic_2A1: was by definition: nta_ldef ntaa_ldef *)
lemma nta_ldef_drops (a) (h) (G) (K) (L) (i):
∀V,W. ⦃G,K⦄ ⊢ V :[a,h] W →
∀U. ⬆*[↑i] W ≘ U → ⬇*[i] L ≘ K.ⓓV → ⦃G,L⦄ ⊢ #i :[a,h] U.
--- /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/cnv_fsb.ma".
+include "basic_2/dynamic/nta.ma".
+
+(* NATIVE TYPE ASSIGNMENT FOR TERMS *****************************************)
+
+(* Forward lemmas with strongly rst-normalizing closures ********************)
+
+(* Note: this is an extended stong normalization property *)
+(* Note: this might use fsb_inv_cast (still to be proved) *)
+(* Basic_1: uses: ty3_sn3 *)
+(* Basic_2A1: uses: nta_fwd_csn *)
+theorem nta_fwd_fsb (a) (h) (o) (G) (L):
+ ∀T,U. ⦃G,L⦄ ⊢ T :[a,h] U →
+ ∧∧ ≥[h,o] 𝐒⦃G,L,T⦄ & ≥[h,o] 𝐒⦃G,L,U⦄.
+#a #h #o #G #L #T #U #H
+elim (cnv_inv_cast … H) #X #HU #HT #_ #_ -X
+/3 width=2 by cnv_fwd_fsb, conj/
+qed-.
lapply (nta_fwd_cnv_dx … HTU) #H
elim (cnv_inv_bind … H) -H #_ #HU
elim (cnv_nta_sn … HU) -HU #X #HUX
-/4 width=2 by nta_beta, nta_fwd_cnv_sn/
+/4 width=2 by nta_appl_abst, nta_fwd_cnv_sn/
qed-.
#H destruct
qed.
+(* Basic_1: uses: ty3_sred_wcpr0_pr0 *)
+lemma nta_cpr_conf_lpr (a) (h) (G):
+ ∀L1,T1,U. ⦃G,L1⦄ ⊢ T1 :[a,h] U → ∀T2. ⦃G,L1⦄ ⊢ T1 ➡[h] T2 →
+ ∀L2. ⦃G,L1⦄ ⊢ ➡[h] L2 → ⦃G,L2⦄ ⊢ T2 :[a,h] U.
+#a #h #G #L1 #T1 #U #H #T2 #HT12 #L2 #HL12
+/3 width=6 by cnv_cpm_trans_lpr, cpm_cast/
+qed-.
+
+(* Basic_1: uses: ty3_sred_pr2 ty3_sred_pr0 *)
+lemma nta_cpr_conf (a) (h) (G) (L):
+ ∀T1,U. ⦃G,L⦄ ⊢ T1 :[a,h] U →
+ ∀T2. ⦃G,L⦄ ⊢ T1 ➡[h] T2 → ⦃G,L⦄ ⊢ T2 :[a,h] U.
+#a #h #G #L #T1 #U #H #T2 #HT12
+/3 width=6 by cnv_cpm_trans, cpm_cast/
+qed-.
+
+(* Note: this is the preservation property *)
+(* Basic_1: uses: ty3_sred_pr3 ty3_sred_pr1 *)
+lemma nta_cprs_conf (a) (h) (G) (L):
+ ∀T1,U. ⦃G,L⦄ ⊢ T1 :[a,h] U →
+ ∀T2. ⦃G,L⦄ ⊢ T1 ➡*[h] T2 → ⦃G,L⦄ ⊢ T2 :[a,h] U.
+#a #h #G #L #T1 #U #H #T2 #HT12
+/3 width=6 by cnv_cpms_trans, cpms_cast/
+qed-.
+
+(* Basic_1: uses: ty3_cred_pr2 *)
+lemma nta_lpr_conf (a) (h) (G):
+ ∀L1,T,U. ⦃G,L1⦄ ⊢ T :[a,h] U →
+ ∀L2. ⦃G,L1⦄ ⊢ ➡[h] L2 → ⦃G,L2⦄ ⊢ T :[a,h] U.
+#a #h #G #L1 #T #U #HTU #L2 #HL12
+/2 width=3 by cnv_lpr_trans/
+qed-.
+
+(* Basic_1: uses: ty3_cred_pr3 *)
+lemma nta_lprs_conf (a) (h) (G):
+ ∀L1,T,U. ⦃G,L1⦄ ⊢ T :[a,h] U →
+ ∀L2. ⦃G,L1⦄ ⊢ ➡*[h] L2 → ⦃G,L2⦄ ⊢ T :[a,h] U.
+#a #h #G #L1 #T #U #HTU #L2 #HL12
+/2 width=3 by cnv_lprs_trans/
+qed-.
+
(* Inversion lemmas based on preservation ***********************************)
lemma nta_inv_ldef_sn (a) (h) (G) (K) (V):
]
qed-.
+(* Basic_1: uses: ty3_gen_lift *)
+(* Note: "⦃G,L⦄ ⊢ U2 ⬌*[h] X2" can be "⦃G,L⦄ ⊢ X2 ➡*[h] U2" *)
+lemma nta_inv_lifts_sn (a) (h) (G):
+ ∀L,T2,X2. ⦃G,L⦄ ⊢ T2 :[a,h] X2 →
+ ∀b,f,K. ⬇*[b,f] L ≘ K → ∀T1. ⬆*[f] T1 ≘ T2 →
+ ∃∃U1,U2. ⦃G,K⦄ ⊢ T1 :[a,h] U1 & ⬆*[f] U1 ≘ U2 & ⦃G,L⦄ ⊢ U2 ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 ![a,h].
+#a #h #G #L #T2 #X2 #H #b #f #K #HLK #T1 #HT12
+elim (cnv_inv_cast … H) -H #U2 #HX2 #HT2 #HXU2 #HTU2
+lapply (cnv_inv_lifts … HT2 … HLK … HT12) -HT2 #HT1
+elim (cpms_inv_lifts_sn … HTU2 … HLK … HT12) -T2 -HLK #U1 #HU12 #HTU1
+/3 width=5 by cnv_cpms_nta, cpcs_cprs_sn, ex4_2_intro/
+qed-.
+
(* Forward lemmas based on preservation *************************************)
(* Basic_1: was: ty3_unique *)
lapply (cnv_cpms_conf_eq … HT … HTX1 … HTX2) -T #HX12
/3 width=3 by cpcs_cprs_div, cpcs_cprs_step_sn/
qed-.
+
+(* Advanced properties ******************************************************)
+
+(* Basic_1: uses: ty3_sconv_pc3 *)
+lemma nta_cpcs_bi (a) (h) (G) (L):
+ ∀T1,U1. ⦃G,L⦄ ⊢ T1 :[a,h] U1 → ∀T2,U2. ⦃G,L⦄ ⊢ T2 :[a,h] U2 →
+ ⦃G,L⦄ ⊢ T1 ⬌*[h] T2 → ⦃G,L⦄ ⊢ U1 ⬌*[h] U2.
+#a #h #G #L #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
+elim (cpcs_inv_cprs … HT12) -HT12 #T0 #HT10 #HT02
+/3 width=6 by nta_mono, nta_cprs_conf/
+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_cpcs.ma".
+include "basic_2/dynamic/nta_preserve.ma".
+
+(* NATIVE TYPE ASSIGNMENT FOR TERMS ****************************************)
+
+(* Properties based on type equivalence and preservation *******************)
+
+(* Basic_1: uses: ty3_tred *)
+lemma nta_cprs_trans (a) (h) (G) (L):
+ ∀T,U1. ⦃G,L⦄ ⊢ T :[a,h] U1 → ∀U2. ⦃G,L⦄ ⊢ U1 ➡*[h] U2 → ⦃G,L⦄ ⊢ T :[a,h] U2.
+#a #h #G #L #T #U1 #H #U2 #HU12
+/4 width=4 by nta_conv_cnv, nta_fwd_cnv_dx, cnv_cpms_trans, cpcs_cprs_dx/
+qed-.
+
+(* Basic_1: uses: ty3_sred_back *)
+lemma cprs_nta_trans (a) (h) (G) (L):
+ ∀T1,U0. ⦃G,L⦄ ⊢ T1 :[a,h] U0 → ∀T2. ⦃G,L⦄ ⊢ T1 ➡*[h] T2 →
+ ∀U. ⦃G,L⦄ ⊢ T2 :[a,h] U → ⦃G,L⦄ ⊢ T1 :[a,h] U.
+#a #h #G #L #T1 #U0 #HT1 #T2 #HT12 #U #H
+lapply (nta_cprs_conf … HT1 … HT12) -HT12 #HT2
+/4 width=6 by nta_mono, nta_conv_cnv, nta_fwd_cnv_dx/
+qed-.
+
+lemma cprs_nta_trans_cnv (a) (h) (G) (L):
+ ∀T1. ⦃G,L⦄ ⊢ T1 ![a,h] → ∀T2. ⦃G,L⦄ ⊢ T1 ➡*[h] T2 →
+ ∀U. ⦃G,L⦄ ⊢ T2 :[a,h] U → ⦃G,L⦄ ⊢ T1 :[a,h] U.
+#a #h #G #L #T1 #HT1 #T2 #HT12 #U #H
+elim (cnv_nta_sn … HT1) -HT1 #U0 #HT1
+/2 width=3 by cprs_nta_trans/
+qed-.
+
+(* Basic_1: uses: ty3_sconv *)
+lemma nta_cpcs_conf (a) (h) (G) (L):
+ ∀T1,U. ⦃G,L⦄ ⊢ T1 :[a,h] U → ∀T2. ⦃G,L⦄ ⊢ T1 ⬌*[h] T2 →
+ ∀U0. ⦃G,L⦄ ⊢ T2 :[a,h] U0 → ⦃G,L⦄ ⊢ T2 :[a,h] U.
+#a #h #G #L #T1 #U #HT1 #T2 #HT12 #U0 #HT2
+elim (cpcs_inv_cprs … HT12) -HT12 #T0 #HT10 #HT02
+/3 width=5 by cprs_nta_trans, nta_cprs_conf/
+qed-.
+
+(* Note: type preservation by valid r-equivalence *)
+lemma nta_cpcs_conf_cnv (a) (h) (G) (L):
+ ∀T1,U. ⦃G,L⦄ ⊢ T1 :[a,h] U →
+ ∀T2. ⦃G,L⦄ ⊢ T1 ⬌*[h] T2 → ⦃G,L⦄ ⊢ T2 ![a,h] → ⦃G,L⦄ ⊢ T2 :[a,h] U.
+#a #h #G #L #T1 #U #HT1 #T2 #HT12 #HT2
+elim (cnv_nta_sn … HT2) -HT2 #U0 #HT2
+/2 width=3 by nta_cpcs_conf/
+qed-.
--- /dev/null
+ty3_gen_cvoid
+ty3_acyclic
+ty3_gen_appl_nf2
+ty3_nf2_inv_all (and following)
+pc3_dec
+pc3_abst_dec
+ty3_inference
+(* Basic_1: was by definition: ty3_abst *)
+(* Basic_2A1: was by definition: nta_ldec ntaa_ldec *)
+lemma nta_ldec_drops
+
(* Basic_1: was by definition: ty3_bind *)
(* Basic_2A1: was by definition: nta_bind ntaa_bind *)
lemma nta_bind
(* Basic_2A1: was: nta_inv_bind1 ntaa_inv_bind1 *)
lemma nta_inv_bind_sn
-(* Basic_1: was by definition: ty3_abst *)
-(* Basic_2A1: was by definition: nta_ldec *)
-lemma nta_ldec_drops
-
(* Basic_1: was: ty3_gen_lref *)
(* Basic_2A1: was: nta_inv_lref1 *)
lemma nta_inv_lref_sn_drops
+(* Basic_1: uses: ty3_gen_abst_abst *)
+lemma nta_inv_abst_bi
+
(* Advanced properties ******************************************************)
| ntaa_cast: ∀L,T,U,W. ntaa h L T U → ntaa h L U W → ntaa h L (ⓝU. T) U
--- /dev/null
+lemma nta_inv_pure_sn_cnv (h) (G) (L) (X2):
+ ∀V,T. ⦃G,L⦄ ⊢ ⓐV.T :*[h] X2 →
+ ∨∨ ∃∃p,W,T0,U0. ⦃G,L⦄ ⊢ V :*[h] W & ⦃G,L.ⓛW⦄ ⊢ T0 :*[h] U0 & ⦃G,L⦄ ⊢ T ➡*[h] ⓛ{p}W.T0 & ⦃G,L⦄ ⊢ ⓐV.ⓛ{p}W.U0 ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 !*[h]
+ | ∃∃U. ⦃G,L⦄ ⊢ T :*[h] U & ⦃G,L⦄ ⊢ ⓐV.U !*[h] & ⦃G,L⦄ ⊢ ⓐV.U ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 !*[h].
+#h #G #L #X2 #V #T #H
+elim (cnv_inv_cast … H) -H #X1 #HX2 #H1 #HX21 #H
+elim (cnv_inv_appl … H1) * [| #n ] #p #W0 #T0 #Hn #HV #HT #HW0 #HT0
+[ lapply (cpms_appl_dx … V V … HT0) // #H0
+ lapply (cnv_cpms_trans … H1 … H0) #H2
+ elim (cnv_cpms_conf … H1 … H … H0) -H1 -H -H0
+ <minus_O_n <minus_n_O #X #HX #HT0X
+ lapply (cprs_trans … HX21 … HX) -X1 #H1X2
+ elim (cnv_cpms_fwd_appl_sn_decompose … H2 … HT0X) -H2 -HT0X #X0 #H #HTX0 #HX0
+ elim (cnv_inv_bind … H) -H #_ #H1T0
+ elim (cpms_inv_abst_sn_cprs … HTX0) -HTX0 #U0 #HTU0 #HUX0
+ @or_introl @(ex5_4_intro … U0 … HT0 … HX2)
+ [ /2 width=1 by cnv_cpms_nta/
+ | /3 width=1 by cnv_cpms_nta/
+ | /4 width=3 by cpcs_cprs_div, cpcs_cprs_step_sn, cpms_appl_dx/
+ ]
+| elim (cnv_cpms_fwd_appl_sn_decompose … H1 … H) -H1 -H #X0 #_ #H #HX01
+ elim (cpms_inv_plus … 1 n … HT0) #U #HTU #HUT0
+ lapply (cnv_cpms_trans … HT … HTU) #HU
+ lapply (cnv_cpms_conf_eq … HT … HTU … H) -H #HUX0
+ @or_intror @(ex4_intro … U … HX2) -HX2
+ [ /2 width=1 by cnv_cpms_nta/
+ | /4 width=7 by cnv_appl, lt_to_le/
+ | /4 width=3 by cpcs_trans, cpcs_cprs_div, cpcs_flat/
+ ]
+]
+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/equivalence/cpcs_ltpss.ma".
+include "basic_2/dynamic/nta_nta.ma".
+
+(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* Properties about parallel unfold *****************************************)
+
+lemma nta_ltpss_tpss_conf: ∀h,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → ⦃h, L2⦄ ⊢ T2 : U.
+#h #L1 #T1 #U #H @(nta_ind_alt … H) -L1 -T1 -U
+[ #L1 #k #L2 #d #e #_ #T2 #H
+ >(tpss_inv_sort1 … H) -H //
+| #L1 #K1 #V1 #W #U #i #HLK1 #_ #HWU #IHV1 #L2 #d #e #HL12 #T2 #H
+ elim (tpss_inv_lref1 … H) -H
+ [ #H destruct
+ elim (lt_or_ge i d) #Hdi
+ [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #V2 #HK12 #HV12 #H destruct
+ /3 width=7/
+ | elim (lt_or_ge i (d + e)) #Hide [ | -Hdi ]
+ [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #V2 #HK12 #HV12 #H destruct
+ /3 width=7/
+ | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=7/
+ ]
+ ]
+ | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2
+ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #HK12 #HV12 #H destruct
+ lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
+ lapply (tpss_trans_eq … HV12 HVW2) -V2 /3 width=9/
+ ]
+| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL12 #T2 #H
+ elim (tpss_inv_lref1 … H) -H [ | -HWV1 -HWU1 -IHWV1 ]
+ [ #H destruct
+ elim (lift_total V1 0 (i+1)) #W #HW
+ elim (lt_or_ge i d) #Hdi [ -HWV1 ]
+ [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #W2 #HK12 #HW12 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
+ lapply (nta_lift h … HLK … HWU1 … HW) /2 width=4/ -HW
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) -HLK -HWU1 // #HU12
+ lapply (cpr_tpss … HU12) -HU12 #HU12
+ @(nta_conv … U2) /2 width=1/ /3 width=6/ (**) (* explicit constructor, /4 width=6/ is too slow *)
+ | elim (lt_or_ge i (d + e)) #Hide [ -HWV1 | -IHWV1 -HW -Hdi ]
+ [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #W2 #HK12 #HW12 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
+ lapply (nta_lift h … HLK … HWU1 … HW) /2 width=4/ -HW
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) -HLK -HWU1 // #HU12
+ lapply (cpr_tpss … HU12) -HU12 #HU12
+ @(nta_conv … U2) /2 width=1/ /3 width=6/ (**) (* explicit constructor, /4 width=6/ is too slow *)
+ | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /2 width=6/
+ ]
+ ]
+ | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #_ #_
+ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #_ #_ #H destruct
+ lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 -HLK2 #H destruct
+ ]
+| #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
+ lapply (cpr_tpss … HV12) #HV
+ lapply (IHTU1 (L2.ⓑ{I}V1) (d+1) e ? T1 ?) // /2 width=1/ #H
+ elim (nta_fwd_correct … H) -H #U2 #HU12
+ @(nta_conv … (ⓑ{I}V2.U1)) /3 width=2/ /3 width=4/ /4 width=4/ (**) (* explicit constructor, /5 width=6/ is too slow *)
+| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_flat1 … H) -H #V2 #Y #HV12 #HY #H destruct
+ elim (tpss_inv_bind1 … HY) -HY #W2 #T2 #HW12 #HT12 #H destruct
+ lapply (cpr_tpss … HV12) #HV
+ lapply (IHTU1 L2 d e ? (ⓛW1.T1) ?) // #H
+ elim (nta_fwd_correct … H) -H #X #H
+ elim (nta_inv_bind1 … H) -H #W #U #HW #HU #_
+ @(nta_conv … (ⓐV2.ⓛW1.U1)) /3 width=2/ /3 width=4/ /4 width=5/ (**) (* explicit constructor, /5 width=5/ is too slow *)
+| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
+ lapply (cpr_tpss … HV12) #HV
+ elim (nta_fwd_correct h L2 (ⓐV1.T1) (ⓐV1.U1) ?) [2: /3 width=4/ ] #U #HU
+ @(nta_conv … (ⓐV2.U1)) // /3 width=1/ /4 width=5/ (**) (* explicit constructor, /5 width=5/ is too slow *)
+| #L1 #T1 #U1 #W1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_flat1 … H) -H #U2 #T2 #HU12 #HT12 #H destruct
+ lapply (cpr_tpss … HU12) /4 width=4/
+| #L1 #T1 #U11 #U12 #U #_ #HU112 #_ #IHTU11 #IHU12 #L2 #d #e #HL12 #T2 #HT12
+ @(nta_conv … U11) /2 width=5/ (**) (* explicot constructor, /3 width=7/ is too slow *)
+]
+qed.
+
+lemma nta_ltpss_tps_conf: ∀h,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶ [d, e] T2 → ⦃h, L2⦄ ⊢ T2 : U.
+/3 width=7/ qed.
+
+lemma nta_ltpss_conf: ∀h,L1,T,U. ⦃h, L1⦄ ⊢ T : U →
+ ∀L2,d,e. L1 ▶* [d, e] L2 → ⦃h, L2⦄ ⊢ T : U.
+/2 width=7/ qed.
+
+lemma nta_tpss_conf: ∀h,L,T1,U. ⦃h, L⦄ ⊢ T1 : U →
+ ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 → ⦃h, L⦄ ⊢ T2 : U.
+/2 width=7/ qed.
+
+lemma nta_tps_conf: ∀h,L,T1,U. ⦃h, L⦄ ⊢ T1 : U →
+ ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ⦃h, L⦄ ⊢ T2 : U.
+/2 width=7/ 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/static/sta.ma".
+include "basic_2/equivalence/cpcs_cpcs.ma".
+include "basic_2/dynamic/nta.ma".
+
+(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* Properties on static type assignment *************************************)
+
+lemma nta_fwd_sta: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U →
+ ∃∃U0. ⦃h, L⦄ ⊢ T • U0 & L ⊢ U0 ⬌* U.
+#h #L #T #U #H elim H -L -T -U
+[ /2 width=3/
+| #L #K #V #W1 #V1 #i #HLK #_ #HWV1 * #W0 #HVW0 #HW01
+ elim (lift_total W0 0 (i+1)) #V0 #HWV0
+ lapply (ldrop_fwd_ldrop2 … HLK) #HLK0
+ lapply (cpcs_lift … HLK0 … HWV0 … HWV1 HW01) -HLK0 -HWV1 -HW01 /3 width=6/
+| #L #K #W #V1 #W1 #i #HLK #HWV1 #HW1 * /3 width=6/
+| #I #L #V #W #T #U #_ #_ * #W0 #_ #_ * /3 width=3/
+| #L #V #W #T #U #_ #_ * #W0 #_ #HW0 * #X #H #HX
+ elim (sta_inv_bind1 … H) -H #U0 #HTU0 #H destruct
+ @(ex2_1_intro … (ⓐV.ⓛW.U0)) /2 width=1/ /3 width=1/
+| #L #V #W #T #U #_ #_ * #U0 #HTU0 #HU0 #_ -W
+ @(ex2_1_intro … (ⓐV.U0)) /2 width=1/
+| #L #T #U #HTU * #U0 #HTU0 #HU0 /3 width=3/
+| #L #T #U1 #U2 #V2 #_ #HU12 #_ * #U0 #HTU0 #HU01 #_
+ lapply (cpcs_trans … HU01 … HU12) -U1 /2 width=3/
+]
+qed-.
+(*
+ Lemma ty3_sty0: (g:?; c:?; u,t1:?) (ty3 g c u t1) ->
+ (t2:?) (sty0 g c u t2) -> (ty3 g c u 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 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/unfold/thin_ldrop.ma".
+include "basic_2/equivalence/cpcs_delift.ma".
+include "basic_2/dynamic/nta_lift.ma".
+
+(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* Properties on basic local environment thinning ***************************)
+
+(* Note: this is known as the substitution lemma *)
+(* Basic_1: was only: ty3_gen_cabbr *)
+lemma nta_thin_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 →
+ ∀L2,d,e. ≽ [d, e] L1 → L1 ▼*[d, e] ≡ L2 →
+ ∃∃T2,U2. ⦃h, L2⦄ ⊢ T2 : U2 &
+ L1 ⊢ T1 ▼*[d, e] ≡ T2 & L1 ⊢ U1 ▼*[d, e] ≡ U2.
+#h #L1 #T1 #U1 #H elim H -L1 -T1 -U1
+[ /2 width=5/
+| #L1 #K1 #V1 #W1 #U1 #i #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL1 #HL12
+ elim (lt_or_ge i d) #Hdi [ -HVW1 ]
+ [ lapply (sfr_ldrop_trans_ge … HLK1 … HL1 ?) -HL1 /2 width=2/ #H
+ lapply (sfr_inv_skip … H ?) -H /2 width=1/ #HK1
+ elim (thin_ldrop_conf_le … HL12 … HLK1 ?) -HL12 /2 width=2/ #X #H #HLK2
+ elim (thin_inv_delift1 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
+ elim (IHVW1 … HK1 HK12) -IHVW1 -HK1 -HK12 #X2 #W2 #HVW2 #H #HW12
+ lapply (delift_mono … H … HV12) -H -HV12 #H destruct
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (ldrop_fwd_ldrop2 … HLK1) -V1 #HLK1
+ lapply (delift_lift_ge … HW12 … HLK1 HWU1 … HWU2) -HW12 -HLK1 -HWU1 //
+ >minus_plus <plus_minus_m_m // /3 width=6/
+ | elim (lt_or_ge i (d+e)) #Hide [ -HVW1 | -Hdi -IHVW1 -HL1 ]
+ [ lapply (sfr_ldrop_trans_be_up … HLK1 … HL1 ? ?) -HL1 // /2 width=2/ <minus_n_O #H
+ elim (sfr_inv_bind … H ?) -H /2 width=1/ #HK1 #_
+ elim (thin_ldrop_conf_be … HL12 … HLK1 ? ?) -HL12 /2 width=2/ #K2 #H #HLK2
+ lapply (thin_inv_thin1 … H ?) -H /2 width=1/ #HK12
+ elim (IHVW1 … HK1 HK12) -IHVW1 -HK1 -HK12 #V2 #W2 #HVW2 #HV12 #HW12
+ elim (lift_total V2 0 d) #T2 #HVT2
+ elim (lift_total W2 0 d) #U2 #HWU2
+ elim (lift_total W2 0 (i+1)) #U #HW2U
+ lapply (nta_lift … HVW2 … HLK2 … HVT2 … HWU2) -HVW2 -HLK2 #HTU2
+ lapply (ldrop_fwd_ldrop2 … HLK1) #HLK0
+ lapply (delift_lift_ge … HW12 … HLK0 HWU1 … HW2U) -HW12 -HLK0 -HWU1 // >minus_plus #HU1
+ lapply (lift_conf_be … HWU2 … HW2U ?) -W2 /2 width=1/ #HU2
+ lapply (delift_lift_div_be … HU1 … HU2 ? ?) -U // /2 width=1/ /3 width=8/
+ | lapply (transitive_le … (i+1) Hide ?) /2 width=1/ #Hdei
+ lapply (thin_ldrop_conf_ge … HL12 … HLK1 ?) -HL12 -HLK1 // #HL2K1
+ elim (lift_split … HWU1 d (i+1-e) ? ? ?) -HWU1 // /2 width=1/ #W
+ <plus_minus in ⊢ (??%??→?); /2 width=2/ #HW1
+ <minus_minus // /2 width=2/ -Hdei >commutative_plus <minus_n_n /3 width=6/
+ ]
+ ]
+| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL1 #HL12
+ elim (lt_or_ge i d) #Hdi [ -HWV1 | -IHWV1 ]
+ [ lapply (sfr_ldrop_trans_ge … HLK1 … HL1 ?) -HL1 /2 width=2/ #H
+ lapply (sfr_inv_skip … H ?) -H /2 width=1/ #HK1
+ elim (thin_ldrop_conf_le … HL12 … HLK1 ?) -HL12 /2 width=2/ #X #H #HLK2
+ elim (thin_inv_delift1 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
+ elim (IHWV1 … HK1 HK12) -IHWV1 -HK1 -HK12 #X2 #V2 #HWV2 #H #_
+ lapply (delift_mono … H … HW12) -H #H destruct
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 #HLK1
+ lapply (delift_lift_ge … HW12 … HLK1 HWU1 … HWU2) -HW12 -HLK1 -HWU1 //
+ >minus_plus <plus_minus_m_m // /3 width=6/
+ | elim (lt_or_ge i (d+e)) #Hide [ -HWV1 -HWU1 -HL12 | -Hdi -HL1 ]
+ [ lapply (sfr_inv_ldrop … HLK1 … HL1 ? ?) -HLK1 -HL1 // -Hdi -Hide #H destruct
+ | lapply (transitive_le … (i+1) Hide ?) /2 width=1/ #Hdei
+ lapply (thin_ldrop_conf_ge … HL12 … HLK1 ?) -HL12 -HLK1 // #HL2K1
+ elim (lift_split … HWU1 d (i+1-e) ? ? ?) -HWU1 // /2 width=1/ #W
+ <plus_minus in ⊢ (??%??→?); /2 width=2/ #HW1
+ <minus_minus // /2 width=2/ -Hdei >commutative_plus <minus_n_n /3 width=6/
+ ]
+ ]
+| #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL1 #HL12
+ elim (IHVW1 … HL1 HL12) -IHVW1 #V2 #W2 #HVW2 #HV12 #_
+ elim (IHTU1 (L2.ⓑ{I}V2) (d+1) e ? ?) -IHTU1 /2 width=1/ -HL1 -HL12 #T2 #U2 #HTU2 #HT12 #HU12
+ lapply (delift_lsubs_trans … HT12 (L1.ⓑ{I}V2) ?) -HT12 /2 width=1/
+ lapply (delift_lsubs_trans … HU12 (L1.ⓑ{I}V2) ?) -HU12 /2 width=1/ /3 width=7/
+| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL1 #HL12
+ elim (IHVW1 … HL1 HL12) -IHVW1 #V2 #W2 #HVW2 #HV12 #HW12
+ elim (IHTU1 … HL1 HL12) -IHTU1 -HL1 -HL12 #X2 #Y2 #HXY2 #HX2 #HY2
+ elim (delift_inv_bind1 … HX2) -HX2 #Z21 #T2 #HZ21 #HT12 #H destruct
+ elim (delift_inv_bind1 … HY2) -HY2 #Z22 #U2 #HZ22 #HU12 #H destruct
+ lapply (delift_mono … HZ21 … HW12) -HZ21 #H destruct
+ lapply (delift_mono … HZ22 … HW12) -HZ22 #H destruct
+ @(ex3_2_intro … (ⓐV2.ⓛW2.T2) (ⓐV2.ⓛW2.U2)) /3 width=1/ (**) (* explict constructor, /4 depth=?/ is too slow *)
+| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL1 #HL12
+ elim (IHTU1 … HL1 HL12) -IHTU1 #T2 #U2 #HTU2 #HT12 #HU12
+ elim (IHUW1 … HL1 HL12) -IHUW1 -HL1 -HL12 #X2 #W2 #HXW2 #H #HW12
+ elim (delift_inv_flat1 … H) -H #V2 #Y2 #HV12 #HY2 #H destruct
+ lapply (delift_mono … HY2 … HU12) -HY2 #H destruct /3 width=7/
+| #L1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL1 #HL12
+ elim (IHTU1 … HL1 HL12) -IHTU1 -HL1 -HL12 /3 width=5/
+| #L1 #T1 #U11 #U12 #V1 #_ #HU112 #_ #IHT1 #IHU12 #L2 #d #e #HL1 #HL12
+ elim (IHT1 … HL1 HL12) -IHT1 #T2 #U21 #HT2 #HT12 #HU121
+ elim (IHU12 … HL1 HL12) -IHU12 -HL1 #U22 #V2 #HU22 #HU122 #_
+ lapply (thin_cpcs_delift_mono … HU112 … HL12 … HU121 … HU122) -HU112 -HL12 -HU121 /3 width=5/
+]
+qed.
+
+lemma nta_ldrop_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 →
+ ∀L2,d,e. ≽ [d, e] L1 → ⇩[d, e] L1 ≡ L2 →
+ ∃∃T2,U2. ⦃h, L2⦄ ⊢ T2 : U2 &
+ L1 ⊢ T1 ▼*[d, e] ≡ T2 & L1 ⊢ U1 ▼*[d, e] ≡ U2.
+/3 width=1/ qed.
+++ /dev/null
-lemma nta_inv_pure_sn_cnv (h) (G) (L) (X2):
- ∀V,T. ⦃G,L⦄ ⊢ ⓐV.T :*[h] X2 →
- ∨∨ ∃∃p,W,T0,U0. ⦃G,L⦄ ⊢ V :*[h] W & ⦃G,L.ⓛW⦄ ⊢ T0 :*[h] U0 & ⦃G,L⦄ ⊢ T ➡*[h] ⓛ{p}W.T0 & ⦃G,L⦄ ⊢ ⓐV.ⓛ{p}W.U0 ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 !*[h]
- | ∃∃U. ⦃G,L⦄ ⊢ T :*[h] U & ⦃G,L⦄ ⊢ ⓐV.U !*[h] & ⦃G,L⦄ ⊢ ⓐV.U ⬌*[h] X2 & ⦃G,L⦄ ⊢ X2 !*[h].
-#h #G #L #X2 #V #T #H
-elim (cnv_inv_cast … H) -H #X1 #HX2 #H1 #HX21 #H
-elim (cnv_inv_appl … H1) * [| #n ] #p #W0 #T0 #Hn #HV #HT #HW0 #HT0
-[ lapply (cpms_appl_dx … V V … HT0) // #H0
- lapply (cnv_cpms_trans … H1 … H0) #H2
- elim (cnv_cpms_conf … H1 … H … H0) -H1 -H -H0
- <minus_O_n <minus_n_O #X #HX #HT0X
- lapply (cprs_trans … HX21 … HX) -X1 #H1X2
- elim (cnv_cpms_fwd_appl_sn_decompose … H2 … HT0X) -H2 -HT0X #X0 #H #HTX0 #HX0
- elim (cnv_inv_bind … H) -H #_ #H1T0
- elim (cpms_inv_abst_sn_cprs … HTX0) -HTX0 #U0 #HTU0 #HUX0
- @or_introl @(ex5_4_intro … U0 … HT0 … HX2)
- [ /2 width=1 by cnv_cpms_nta/
- | /3 width=1 by cnv_cpms_nta/
- | /4 width=3 by cpcs_cprs_div, cpcs_cprs_step_sn, cpms_appl_dx/
- ]
-| elim (cnv_cpms_fwd_appl_sn_decompose … H1 … H) -H1 -H #X0 #_ #H #HX01
- elim (cpms_inv_plus … 1 n … HT0) #U #HTU #HUT0
- lapply (cnv_cpms_trans … HT … HTU) #HU
- lapply (cnv_cpms_conf_eq … HT … HTU … H) -H #HUX0
- @or_intror @(ex4_intro … U … HX2) -HX2
- [ /2 width=1 by cnv_cpms_nta/
- | /4 width=7 by cnv_appl, lt_to_le/
- | /4 width=3 by cpcs_trans, cpcs_cprs_div, cpcs_flat/
- ]
-]
-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 term 46 L1 : ⊑ break term 46 L2 )"
+ non associative with precedence 45
+ for @{ 'CrSubEqN $h $L1 $L2 }.
+
+notation "hvbox( h ⊢ break term 46 L1 : : ⊑ break term 46 L2 )"
+ non associative with precedence 45
+ for @{ 'CrSubEqNAlt $h $L1 $L2 }.
+
+include "basic_2/dynamic/nta.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
+
+(* Note: may not be transitive *)
+inductive lsubn (h:sh): relation lenv ≝
+| lsubn_atom: lsubn h (⋆) (⋆)
+| lsubn_pair: ∀I,L1,L2,W. lsubn h L1 L2 → lsubn h (L1. ⓑ{I} W) (L2. ⓑ{I} W)
+| lsubn_abbr: ∀L1,L2,V,W. ⦃h, L1⦄ ⊢ V : W → ⦃h, L2⦄ ⊢ V : W →
+ lsubn h L1 L2 → lsubn h (L1. ⓓV) (L2. ⓛW)
+.
+
+interpretation
+ "local environment refinement (native type assigment)"
+ 'CrSubEqN h L1 L2 = (lsubn h L1 L2).
+
+(* Basic inversion lemmas ***************************************************)
+
+fact lsubn_inv_atom1_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L1 = ⋆ → L2 = ⋆.
+#h #L1 #L2 * -L1 -L2
+[ //
+| #I #L1 #L2 #V #_ #H destruct
+| #L1 #L2 #V #W #_ #_ #_ #H destruct
+]
+qed.
+
+lemma lsubn_inv_atom1: ∀h,L2. h ⊢ ⋆ :⊑ L2 → L2 = ⋆.
+/2 width=4/ qed-.
+
+fact lsubn_inv_pair1_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀I,K1,V. L1 = K1. ⓑ{I} V →
+ (∃∃K2. h ⊢ K1 :⊑ K2 & L2 = K2. ⓑ{I} V) ∨
+ ∃∃K2,W. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
+ h ⊢ K1 :⊑ K2 & L2 = K2. ⓛW & I = Abbr.
+#h #L1 #L2 * -L1 -L2
+[ #I #K1 #V #H destruct
+| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3/
+| #L1 #L2 #V #W #H1VW #H2VW #HL12 #I #K1 #V1 #H destruct /3 width=7/
+]
+qed.
+
+lemma lsubn_inv_pair1: ∀h,I,K1,L2,V. h ⊢ K1. ⓑ{I} V :⊑ L2 →
+ (∃∃K2. h ⊢ K1 :⊑ K2 & L2 = K2. ⓑ{I} V) ∨
+ ∃∃K2,W. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
+ h ⊢ K1 :⊑ K2 & L2 = K2. ⓛW & I = Abbr.
+/2 width=3/ qed-.
+
+fact lsubn_inv_atom2_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L2 = ⋆ → L1 = ⋆.
+#h #L1 #L2 * -L1 -L2
+[ //
+| #I #L1 #L2 #V #_ #H destruct
+| #L1 #L2 #V #W #_ #_ #_ #H destruct
+]
+qed.
+
+lemma lsubc_inv_atom2: ∀h,L1. h ⊢ L1 :⊑ ⋆ → L1 = ⋆.
+/2 width=4/ qed-.
+
+fact lsubn_inv_pair2_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀I,K2,W. L2 = K2. ⓑ{I} W →
+ (∃∃K1. h ⊢ K1 :⊑ K2 & L1 = K1. ⓑ{I} W) ∨
+ ∃∃K1,V. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
+ h ⊢ K1 :⊑ K2 & L1 = K1. ⓓV & I = Abst.
+#h #L1 #L2 * -L1 -L2
+[ #I #K2 #W #H destruct
+| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3/
+| #L1 #L2 #V #W #H1VW #H2VW #HL12 #I #K2 #W2 #H destruct /3 width=7/
+]
+qed.
+
+(* Basic_1: was: csubt_gen_bind *)
+lemma lsubn_inv_pair2: ∀h,I,L1,K2,W. h ⊢ L1 :⊑ K2. ⓑ{I} W →
+ (∃∃K1. h ⊢ K1 :⊑ K2 & L1 = K1. ⓑ{I} W) ∨
+ ∃∃K1,V. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
+ h ⊢ K1 :⊑ K2 & L1 = K1. ⓓV & I = Abst.
+/2 width=3/ qed-.
+
+(* Basic_forward lemmas *****************************************************)
+
+lemma lsubn_fwd_lsubs1: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L1 ≼[0, |L1|] L2.
+#h #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
+qed-.
+
+lemma lsubn_fwd_lsubs2: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L1 ≼[0, |L2|] L2.
+#h #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: csubt_refl *)
+lemma lsubn_refl: ∀h,L. h ⊢ L :⊑ L.
+#h #L elim L -L // /2 width=1/
+qed.
+
+(* Basic_1: removed theorems 6:
+ csubt_gen_flat csubt_drop_flat csubt_clear_conf
+ csubt_getl_abbr csubt_getl_abst csubt_ty3_ld
+*)
--- /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/equivalence/cpcs_cpcs.ma".
+include "basic_2/dynamic/lsubn.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
+
+(* Properties on context-sensitive parallel equivalence for terms ***********)
+
+(* Basic_1: was: csubt_pr2 *)
+lemma cpr_lsubn_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 →
+ ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡ T2.
+/3 width=4 by lsubn_fwd_lsubs2, cpr_lsubs_trans/ qed.
+
+lemma cprs_lsubn_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 →
+ ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2.
+/3 width=4 by lsubn_fwd_lsubs2, cprs_lsubs_trans/ qed.
+
+(* Basic_1: was: csubt_pc3 *)
+lemma cpcs_lsubn_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 →
+ ∀T1,T2. L2 ⊢ T1 ⬌* T2 → L1 ⊢ T1 ⬌* T2.
+/3 width=4 by lsubn_fwd_lsubs2, cpcs_lsubs_trans/ 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/lsubn.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
+
+(* Properties concerning basic local environment slicing ********************)
+
+(* Note: the constant 0 cannot be generalized *)
+lemma lsubn_ldrop_O1_conf: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀K1,e. ⇩[0, e] L1 ≡ K1 →
+ ∃∃K2. h ⊢ K1 :⊑ K2 & ⇩[0, e] L2 ≡ K2.
+#h #L1 #L2 #H elim H -L1 -L2
+[ /2 width=3/
+| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H
+ elim (ldrop_inv_O1 … H) -H * #He #HLK1
+ [ destruct
+ elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK1) -L1 /3 width=3/
+ ]
+| #L1 #L2 #V #W #H1VW #H2VW #_ #IHL12 #K1 #e #H
+ elim (ldrop_inv_O1 … H) -H * #He #HLK1
+ [ destruct
+ elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK1) -L1 /3 width=3/
+ ]
+]
+qed.
+
+(* Note: the constant 0 cannot be generalized *)
+(* Basic_1: was only: csubt_drop_abbr csubt_drop_abst *)
+lemma lsubn_ldrop_O1_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 →
+ ∃∃K1. h ⊢ K1 :⊑ K2 & ⇩[0, e] L1 ≡ K1.
+#h #L1 #L2 #H elim H -L1 -L2
+[ /2 width=3/
+| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H
+ elim (ldrop_inv_O1 … H) -H * #He #HLK2
+ [ destruct
+ elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK2) -L2 /3 width=3/
+ ]
+| #L1 #L2 #V #W #H1VW #H2VW #_ #IHL12 #K2 #e #H
+ elim (ldrop_inv_O1 … H) -H * #He #HLK2
+ [ destruct
+ elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK2) -L2 /3 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_nta.ma".
+include "basic_2/dynamic/lsubn_ldrop.ma".
+include "basic_2/dynamic/lsubn_cpcs.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
+
+(* Properties concerning atomic arity assignment ****************************)
+
+(* Note: the corresponding confluence property does not hold *)
+(* Basic_1: was: csubt_ty3 *)
+lemma lsubn_nta_trans: ∀h,L2,T,U. ⦃h, L2⦄ ⊢ T : U → ∀L1. h ⊢ L1 :⊑ L2 →
+ ⦃h, L1⦄ ⊢ T : U.
+#h #L2 #T #U #H elim H -L2 -T -U
+[ //
+| #L2 #K2 #V2 #W2 #U2 #i #HLK2 #_ #WU2 #IHVW2 #L1 #HL12
+ elim (lsubn_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubn_inv_pair2 … H) -H * #K1
+ [ #HK12 #H destruct /3 width=6/
+ | #V1 #_ #_ #_ #_ #H destruct
+ ]
+| #L2 #K2 #W2 #V2 #U2 #i #HLK2 #_ #HWU2 #IHWV2 #L1 #HL12
+ elim (lsubn_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubn_inv_pair2 … H) -H * #K1 [ | -IHWV2 ]
+ [ #HK12 #H destruct /3 width=6/
+ | #V1 #H1V1W2 #_ #_ #H #_ destruct /2 width=6/
+ ]
+| /4 width=2/
+| /3 width=1/
+| /3 width=2/
+| /3 width=1/
+| /4 width=6/
+]
+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 term 46 L1 : ⊑ break term 46 L2 )"
- non associative with precedence 45
- for @{ 'CrSubEqN $h $L1 $L2 }.
-
-notation "hvbox( h ⊢ break term 46 L1 : : ⊑ break term 46 L2 )"
- non associative with precedence 45
- for @{ 'CrSubEqNAlt $h $L1 $L2 }.
-
-include "basic_2/dynamic/nta.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
-
-(* Note: may not be transitive *)
-inductive lsubn (h:sh): relation lenv ≝
-| lsubn_atom: lsubn h (⋆) (⋆)
-| lsubn_pair: ∀I,L1,L2,W. lsubn h L1 L2 → lsubn h (L1. ⓑ{I} W) (L2. ⓑ{I} W)
-| lsubn_abbr: ∀L1,L2,V,W. ⦃h, L1⦄ ⊢ V : W → ⦃h, L2⦄ ⊢ V : W →
- lsubn h L1 L2 → lsubn h (L1. ⓓV) (L2. ⓛW)
-.
-
-interpretation
- "local environment refinement (native type assigment)"
- 'CrSubEqN h L1 L2 = (lsubn h L1 L2).
-
-(* Basic inversion lemmas ***************************************************)
-
-fact lsubn_inv_atom1_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L1 = ⋆ → L2 = ⋆.
-#h #L1 #L2 * -L1 -L2
-[ //
-| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #V #W #_ #_ #_ #H destruct
-]
-qed.
-
-lemma lsubn_inv_atom1: ∀h,L2. h ⊢ ⋆ :⊑ L2 → L2 = ⋆.
-/2 width=4/ qed-.
-
-fact lsubn_inv_pair1_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀I,K1,V. L1 = K1. ⓑ{I} V →
- (∃∃K2. h ⊢ K1 :⊑ K2 & L2 = K2. ⓑ{I} V) ∨
- ∃∃K2,W. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
- h ⊢ K1 :⊑ K2 & L2 = K2. ⓛW & I = Abbr.
-#h #L1 #L2 * -L1 -L2
-[ #I #K1 #V #H destruct
-| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3/
-| #L1 #L2 #V #W #H1VW #H2VW #HL12 #I #K1 #V1 #H destruct /3 width=7/
-]
-qed.
-
-lemma lsubn_inv_pair1: ∀h,I,K1,L2,V. h ⊢ K1. ⓑ{I} V :⊑ L2 →
- (∃∃K2. h ⊢ K1 :⊑ K2 & L2 = K2. ⓑ{I} V) ∨
- ∃∃K2,W. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
- h ⊢ K1 :⊑ K2 & L2 = K2. ⓛW & I = Abbr.
-/2 width=3/ qed-.
-
-fact lsubn_inv_atom2_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L2 = ⋆ → L1 = ⋆.
-#h #L1 #L2 * -L1 -L2
-[ //
-| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #V #W #_ #_ #_ #H destruct
-]
-qed.
-
-lemma lsubc_inv_atom2: ∀h,L1. h ⊢ L1 :⊑ ⋆ → L1 = ⋆.
-/2 width=4/ qed-.
-
-fact lsubn_inv_pair2_aux: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀I,K2,W. L2 = K2. ⓑ{I} W →
- (∃∃K1. h ⊢ K1 :⊑ K2 & L1 = K1. ⓑ{I} W) ∨
- ∃∃K1,V. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
- h ⊢ K1 :⊑ K2 & L1 = K1. ⓓV & I = Abst.
-#h #L1 #L2 * -L1 -L2
-[ #I #K2 #W #H destruct
-| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3/
-| #L1 #L2 #V #W #H1VW #H2VW #HL12 #I #K2 #W2 #H destruct /3 width=7/
-]
-qed.
-
-(* Basic_1: was: csubt_gen_bind *)
-lemma lsubn_inv_pair2: ∀h,I,L1,K2,W. h ⊢ L1 :⊑ K2. ⓑ{I} W →
- (∃∃K1. h ⊢ K1 :⊑ K2 & L1 = K1. ⓑ{I} W) ∨
- ∃∃K1,V. ⦃h, K1⦄ ⊢ V : W & ⦃h, K2⦄ ⊢ V : W &
- h ⊢ K1 :⊑ K2 & L1 = K1. ⓓV & I = Abst.
-/2 width=3/ qed-.
-
-(* Basic_forward lemmas *****************************************************)
-
-lemma lsubn_fwd_lsubs1: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L1 ≼[0, |L1|] L2.
-#h #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
-qed-.
-
-lemma lsubn_fwd_lsubs2: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → L1 ≼[0, |L2|] L2.
-#h #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
-qed-.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was: csubt_refl *)
-lemma lsubn_refl: ∀h,L. h ⊢ L :⊑ L.
-#h #L elim L -L // /2 width=1/
-qed.
-
-(* Basic_1: removed theorems 6:
- csubt_gen_flat csubt_drop_flat csubt_clear_conf
- csubt_getl_abbr csubt_getl_abst csubt_ty3_ld
-*)
+++ /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/equivalence/cpcs_cpcs.ma".
-include "basic_2/dynamic/lsubn.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
-
-(* Properties on context-sensitive parallel equivalence for terms ***********)
-
-(* Basic_1: was: csubt_pr2 *)
-lemma cpr_lsubn_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 →
- ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡ T2.
-/3 width=4 by lsubn_fwd_lsubs2, cpr_lsubs_trans/ qed.
-
-lemma cprs_lsubn_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 →
- ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2.
-/3 width=4 by lsubn_fwd_lsubs2, cprs_lsubs_trans/ qed.
-
-(* Basic_1: was: csubt_pc3 *)
-lemma cpcs_lsubn_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 →
- ∀T1,T2. L2 ⊢ T1 ⬌* T2 → L1 ⊢ T1 ⬌* T2.
-/3 width=4 by lsubn_fwd_lsubs2, cpcs_lsubs_trans/ 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/lsubn.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
-
-(* Properties concerning basic local environment slicing ********************)
-
-(* Note: the constant 0 cannot be generalized *)
-lemma lsubn_ldrop_O1_conf: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀K1,e. ⇩[0, e] L1 ≡ K1 →
- ∃∃K2. h ⊢ K1 :⊑ K2 & ⇩[0, e] L2 ≡ K2.
-#h #L1 #L2 #H elim H -L1 -L2
-[ /2 width=3/
-| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H
- elim (ldrop_inv_O1 … H) -H * #He #HLK1
- [ destruct
- elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H
- <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
- | elim (IHL12 … HLK1) -L1 /3 width=3/
- ]
-| #L1 #L2 #V #W #H1VW #H2VW #_ #IHL12 #K1 #e #H
- elim (ldrop_inv_O1 … H) -H * #He #HLK1
- [ destruct
- elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H
- <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
- | elim (IHL12 … HLK1) -L1 /3 width=3/
- ]
-]
-qed.
-
-(* Note: the constant 0 cannot be generalized *)
-(* Basic_1: was only: csubt_drop_abbr csubt_drop_abst *)
-lemma lsubn_ldrop_O1_trans: ∀h,L1,L2. h ⊢ L1 :⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 →
- ∃∃K1. h ⊢ K1 :⊑ K2 & ⇩[0, e] L1 ≡ K1.
-#h #L1 #L2 #H elim H -L1 -L2
-[ /2 width=3/
-| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H
- elim (ldrop_inv_O1 … H) -H * #He #HLK2
- [ destruct
- elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H
- <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
- | elim (IHL12 … HLK2) -L2 /3 width=3/
- ]
-| #L1 #L2 #V #W #H1VW #H2VW #_ #IHL12 #K2 #e #H
- elim (ldrop_inv_O1 … H) -H * #He #HLK2
- [ destruct
- elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H
- <(ldrop_inv_refl … H) in HL12; -H /3 width=3/
- | elim (IHL12 … HLK2) -L2 /3 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_nta.ma".
-include "basic_2/dynamic/lsubn_ldrop.ma".
-include "basic_2/dynamic/lsubn_cpcs.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR NATIVE TYPE ASSIGNMENT ******************)
-
-(* Properties concerning atomic arity assignment ****************************)
-
-(* Note: the corresponding confluence property does not hold *)
-(* Basic_1: was: csubt_ty3 *)
-lemma lsubn_nta_trans: ∀h,L2,T,U. ⦃h, L2⦄ ⊢ T : U → ∀L1. h ⊢ L1 :⊑ L2 →
- ⦃h, L1⦄ ⊢ T : U.
-#h #L2 #T #U #H elim H -L2 -T -U
-[ //
-| #L2 #K2 #V2 #W2 #U2 #i #HLK2 #_ #WU2 #IHVW2 #L1 #HL12
- elim (lsubn_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
- elim (lsubn_inv_pair2 … H) -H * #K1
- [ #HK12 #H destruct /3 width=6/
- | #V1 #_ #_ #_ #_ #H destruct
- ]
-| #L2 #K2 #W2 #V2 #U2 #i #HLK2 #_ #HWU2 #IHWV2 #L1 #HL12
- elim (lsubn_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
- elim (lsubn_inv_pair2 … H) -H * #K1 [ | -IHWV2 ]
- [ #HK12 #H destruct /3 width=6/
- | #V1 #H1V1W2 #_ #_ #H #_ destruct /2 width=6/
- ]
-| /4 width=2/
-| /3 width=1/
-| /3 width=2/
-| /3 width=1/
-| /4 width=6/
-]
-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/computation/csn_aaa.ma".
-include "basic_2/equivalence/lcpcs_aaa.ma".
-include "basic_2/dynamic/nta.ma".
-
-(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Forward lemmas on atomic arity assignment for terms **********************)
-
-lemma nta_fwd_aaa: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∃∃A. L ⊢ T ⁝ A & L ⊢ U ⁝ A.
-#h #L #T #U #H elim H -L -T -U
-[ /2 width=3/
-| #L #K #V #W #U #i #HLK #_ #HWU * #B #HV #HW
- lapply (ldrop_fwd_ldrop2 … HLK) /3 width=9/
-| #L #K #W #V #U #i #HLK #_ #HWU * #B #HW #_ -V
- lapply (ldrop_fwd_ldrop2 … HLK) /3 width=9/
-| * #L #V #W #T #U #_ #_ * #B #HV #HW * #A #HT #HU
- [ /3 width=3/ | /3 width=5/ ]
-| #L #V #W #T #U #_ #_ * #B #HV #HW * #X #H1 #H2
- elim (aaa_inv_abst … H1) -H1 #B1 #A1 #HW1 #HT #H destruct
- elim (aaa_inv_abst … H2) -H2 #B2 #A #_ #HU #H destruct
- lapply (aaa_mono … HW1 … HW) -HW1 #H destruct /4 width=5/
-| #L #V #W #T #U #_ #_ * #X #HT #HUX * #A #H #_ -W
- elim (aaa_inv_appl … H) -H #B #HV #HUA
- lapply (aaa_mono … HUA … HUX) -HUX #H destruct /3 width=5/
-| #L #T #U #_ * #A #HT #HU /3 width=3/
-| #L #T #U1 #U2 #V2 #_ #HU12 #_ * #X #HT #HU1 * #A #HU2 #_
- lapply (aaa_cpcs_mono … HU12 … HU1 … HU2) -U1 #H destruct /2 width=3/
-]
-qed-.
-
-(* Note: this is the stong normalization property *)
-(* Basic_1: was only: ty3_sn3 *)
-theorem nta_fwd_csn: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → L ⊢ ⬇* T ∧ L ⊢ ⬇* U.
-#h #L #T #U #H elim (nta_fwd_aaa … H) -H /3 width=2/
-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/equivalence/cpcs_cpcs.ma".
-include "basic_2/dynamic/nta.ma".
-
-(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* alternative definition of nta *)
-inductive ntaa (h:sh): lenv → relation term ≝
-| ntaa_sort: ∀L,k. ntaa h L (⋆k) (⋆(next h k))
-| ntaa_ldef: ∀L,K,V,W,U,i. ⇩[0, i] L ≡ K. ⓓV → ntaa h K V W →
- ⇧[0, i + 1] W ≡ U → ntaa h L (#i) U
-| ntaa_ldec: ∀L,K,W,V,U,i. ⇩[0, i] L ≡ K. ⓛW → ntaa h K W V →
- ⇧[0, i + 1] W ≡ U → ntaa h L (#i) U
-| ntaa_bind: ∀I,L,V,W,T,U. ntaa h L V W → ntaa h (L. ⓑ{I} V) T U →
- ntaa h L (ⓑ{I}V.T) (ⓑ{I}V.U)
-| ntaa_appl: ∀L,V,W,T,U. ntaa h L V W → ntaa h L (ⓛW.T) (ⓛW.U) →
- ntaa h L (ⓐV.ⓛW.T) (ⓐV.ⓛW.U)
-| ntaa_pure: ∀L,V,W,T,U. ntaa h L T U → ntaa h L (ⓐV.U) W →
- ntaa h L (ⓐV.T) (ⓐV.U)
-| ntaa_cast: ∀L,T,U,W. ntaa h L T U → ntaa h L U W → ntaa h L (ⓝU. T) U
-| ntaa_conv: ∀L,T,U1,U2,V2. ntaa h L T U1 → L ⊢ U1 ⬌* U2 → ntaa h L U2 V2 →
- ntaa h L T U2
-.
-
-interpretation "native type assignment (term) alternative"
- 'NativeTypeAlt h L T U = (ntaa h L T U).
-
-(* Advanced inversion lemmas ************************************************)
-
-fact ntaa_inv_bind1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T :: U → ∀J,X,Y. T = ⓑ{J}Y.X →
- ∃∃Z1,Z2. ⦃h, L⦄ ⊢ Y :: Z1 & ⦃h, L.ⓑ{J}Y⦄ ⊢ X :: Z2 &
- L ⊢ ⓑ{J}Y.Z2 ⬌* U.
-#h #L #T #U #H elim H -L -T -U
-[ #L #k #J #X #Y #H destruct
-| #L #K #V #W #U #i #_ #_ #_ #_ #J #X #Y #H destruct
-| #L #K #W #V #U #i #_ #_ #_ #_ #J #X #Y #H destruct
-| #I #L #V #W #T #U #HVW #HTU #_ #_ #J #X #Y #H destruct /2 width=3/
-| #L #V #W #T #U #_ #_ #_ #_ #J #X #Y #H destruct
-| #L #V #W #T #U #_ #_ #_ #_ #J #X #Y #H destruct
-| #L #T #U #W #_ #_ #_ #_ #J #X #Y #H destruct
-| #L #T #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #J #X #Y #H destruct
- elim (IHTU1 ????) -IHTU1 [5: // |2,3,4: skip ] #Z1 #Z2 #HZ1 #HZ2 #HU1
- lapply (cpcs_trans … HU1 … HU12) -U1 /2 width=3/
-]
-qed.
-
-lemma ntaa_inv_bind1: ∀h,J,L,Y,X,U. ⦃h, L⦄ ⊢ ⓑ{J}Y.X :: U →
- ∃∃Z1,Z2. ⦃h, L⦄ ⊢ Y :: Z1 & ⦃h, L.ⓑ{J}Y⦄ ⊢ X :: Z2 &
- L ⊢ ⓑ{J}Y.Z2 ⬌* U.
-/2 width=3/ qed-.
-
-lemma ntaa_nta: ∀h,L,T,U. ⦃h, L⦄ ⊢ T :: U → ⦃h, L⦄ ⊢ T : U.
-#h #L #T #U #H elim H -L -T -U
-// /2 width=1/ /2 width=2/ /2 width=3/ /2 width=6/
-qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma ntaa_lift: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 :: U1 → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 →
- ∀T2. ⇧[d, e] T1 ≡ T2 → ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 :: U2.
-#h #L1 #T1 #U1 #H elim H -L1 -T1 -U1
-[ #L1 #k #L2 #d #e #HL21 #X1 #H1 #X2 #H2
- >(lift_inv_sort1 … H1) -X1
- >(lift_inv_sort1 … H2) -X2 //
-| #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // #W2 #HW12 #HWU2
- elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
- /3 width=8/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
- lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
- ]
-| #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // <minus_plus #W #HW1 #HWU2
- elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
- lapply (lift_mono … HW1 … HW12) -HW1 #H destruct
- elim (lift_total V1 (d-i-1) e) /3 width=8/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
- lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
- ]
-| #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
- elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
- elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
- lapply (lift_mono … H1 … HV12) -H1 #H destruct
- elim (lift_total W1 d e) /4 width=6/
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
- elim (lift_inv_flat1 … H1) -H1 #V2 #X #HV12 #H1 #H destruct
- elim (lift_inv_bind1 … H1) -H1 #W2 #T2 #HW12 #HT12 #H destruct
- elim (lift_inv_flat1 … H2) -H2 #Y2 #X #HY #H2 #H destruct
- elim (lift_inv_bind1 … H2) -H2 #X2 #U2 #HX #HU12 #H destruct
- lapply (lift_mono … HY … HV12) -HY #H destruct
- lapply (lift_mono … HX … HW12) -HX #H destruct /4 width=6/
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
- elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
- elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
- lapply (lift_mono … H1 … HV12) -H1 #H destruct
- elim (lift_total W1 d e) /4 width=6/
-| #L1 #T1 #U1 #W1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL21 #X #H #U2 #HU12
- elim (lift_inv_flat1 … H) -H #X2 #T2 #HUX2 #HT12 #H destruct
- lapply (lift_mono … HUX2 … HU12) -HUX2 #H destruct
- elim (lift_total W1 d e) /3 width=6/
-| #L1 #T1 #U11 #U12 #V12 #_ #HU112 #_ #IHTU11 #IHUV12 #L2 #d #e #HL21 #U1 #HTU1 #U2 #HU12
- elim (lift_total U11 d e) #U #HU11
- elim (lift_total V12 d e) #V22 #HV122
- lapply (cpcs_lift … HL21 … HU11 … HU12 HU112) -HU112 /3 width=6/
-]
-qed.
-
-(* Advanced forvard lemmas **************************************************)
-
-lemma ntaa_fwd_correct: ∀h,L,T,U. ⦃h, L⦄ ⊢ T :: U → ∃T0. ⦃h, L⦄ ⊢ U :: T0.
-#h #L #T #U #H elim H -L -T -U
-[ /2 width=2/
-| #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
- lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
- elim (lift_total V0 0 (i+1)) /3 width=10/
-| #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
- lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
- elim (lift_total V 0 (i+1)) /3 width=10/
-| #I #L #V #W #T #U #HVW #_ #_ * /3 width=2/
-| #L #V #W #T #U #HVW #_ #_ * #X #H
- elim (ntaa_inv_bind1 … H) -H /4 width=2/
-| #L #V #W #T #U #_ #HUW * #T0 #HUT0 /3 width=2/
-| /2 width=2/
-| /2 width=2/
-]
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma nta_ntaa: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ⦃h, L⦄ ⊢ T :: U.
-#h #L #T #U #H elim H -L -T -U
-// /2 width=1/ /2 width=2/ /2 width=3/ /2 width=6/
-#L #T #U #_ #HTU
-elim (ntaa_fwd_correct … HTU) /2 width=2/
-qed.
-
-(* Advanced eliminators *****************************************************)
-
-lemma nta_ind_alt: ∀h. ∀R:lenv→relation term.
- (∀L,k. R L ⋆k ⋆(next h k)) →
- (∀L,K,V,W,U,i.
- ⇩[O, i] L ≡ K.ⓓV → ⦃h, K⦄ ⊢ V : W → ⇧[O, i + 1] W ≡ U →
- R K V W → R L (#i) U
- ) →
- (∀L,K,W,V,U,i.
- ⇩[O, i] L ≡ K.ⓛW → ⦃h, K⦄ ⊢ W : V → ⇧[O, i + 1] W ≡ U →
- R K W V → R L (#i) U
- ) →
- (∀I,L,V,W,T,U.
- ⦃h, L⦄ ⊢ V : W → ⦃h, L.ⓑ{I}V⦄ ⊢ T : U →
- R L V W → R (L.ⓑ{I}V) T U → R L (ⓑ{I}V.T) (ⓑ{I}V.U)
- ) →
- (∀L,V,W,T,U.
- ⦃h, L⦄ ⊢ V : W → ⦃h, L⦄ ⊢ (ⓛW.T):(ⓛW.U) →
- R L V W →R L (ⓛW.T) (ⓛW.U) →R L (ⓐV.ⓛW.T) (ⓐV.ⓛW.U)
- ) →
- (∀L,V,W,T,U.
- ⦃h, L⦄ ⊢ T : U → ⦃h, L⦄ ⊢ (ⓐV.U) : W →
- R L T U → R L (ⓐV.U) W → R L (ⓐV.T) (ⓐV.U)
- ) →
- (∀L,T,U,W.
- ⦃h, L⦄ ⊢ T : U → ⦃h, L⦄ ⊢ U : W →
- R L T U → R L U W → R L (ⓝU.T) U
- ) →
- (∀L,T,U1,U2,V2.
- ⦃h, L⦄ ⊢ T : U1 → L ⊢ U1 ⬌* U2 → ⦃h, L⦄ ⊢ U2 : V2 →
- R L T U1 →R L U2 V2 →R L T U2
- ) →
- ∀L,T,U. ⦃h, L⦄ ⊢ T : U → R L T U.
-#h #R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #L #T #U #H elim (nta_ntaa … H) -L -T -U
-// /3 width=1 by ntaa_nta/ /3 width=3 by ntaa_nta/ /3 width=4 by ntaa_nta/
-/3 width=7 by ntaa_nta/
-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/equivalence/cpcs_delift.ma".
-include "basic_2/dynamic/nta.ma".
-(*
-lemma pippo: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∀X,Y. L ⊢ T ➡* ⓛX.Y →
- ∃Z. L ⊢ U ➡* ⓛX.Z.
-#h #L #T #U #H elim H -L -T -U
-[
-|
-|
-|
-| #L #V #W #T #U #_ #_ #IHVW #IHTU #X #Y #H
-| #L #V #W #T #U #_ #HUW #IHTU #IHUW #X #Y #HTY
- elim (cprs_inv_appl_abst … HTY) -HTY #W1 #T1 #W2 #T2 #HT1 #HT12 #HYT2
- elim (IHTU … HT1) -IHTU -HT1 #U1 #HU1
-
-
-
- *
- [ #V0 #T0 #_ #_ #H destruct
- | #V0 #W0 #T0 #HV0 #HT0 #HTY
- elim (IHTU … HT0) -IHTU -HT0 #Z #HUZ
- elim (cprs_inv_abbr1 … HTY) -HTY *
- [ #V1 #T1 #_ #_ #H destruct #X0
-
-*)
-
-(*
-
-include "basic_2/computation/cprs_lcprs.ma".
-
-
-
-
-include "basic_2/dynamic/nta_ltpss.ma".
-include "basic_2/dynamic/nta_thin.ma".
-include "basic_2/dynamic/lsubn_nta.ma".
-
-include "basic_2/hod/ntas_lift.ma".
-
-
- elim (nta_inv_pure1 … HUW) -HUW #V0 #U0 #U1 #HV0 #HU0 #HU0W #HU01
- @(ex2_2_intro … HYW)
- [2:
-
-
-axiom pippo_aux: ∀h,L,Z,U. ⦃h, L⦄ ⊢ Z : U → ∀Y,X. Z = ⓐY.X →
- ∀W,T. L ⊢ X ➡* ⓛW.T → ⦃h, L⦄ ⊢ Y : W.
-#h #L #Z #U #H elim H -L -Z -U
-[
-|
-|
-|
-|
-| #L #V #W #T #U #HTU #_ #_ #IHUW #Y #X #H #W0 #T0 #HX destruct
- lapply (IHUW Y U ? ?) -IHUW -W // #T
- @(ex2_2_intro … HYW)
- [2:
-
-axiom pippo: ∀h,L,V,X,U. ⦃h, L⦄ ⊢ ⓐV.X : U →
- ∃∃W,T. L ⊢ X ➡* ⓛW.T & ⦃h, L⦄ ⊢ V : W.
-#h #L #V #X #Y #H
-
-*)
-(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Properties on context-free parallel reduction for local environments ******)
-(*
-axiom nta_ltpr_cprs_conf: ∀h,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U → ∀L2. L1 ➡ L2 →
- ∀T2. L2 ⊢ T1 ➡* T2 → ⦃h, L2⦄ ⊢ T2 : U.
-#h #L1 #T1 #U #H @(nta_ind_alt … H) -L1 -T1 -U
-[ #L1 #k #L2 #_ #T2 #H
- >(cprs_inv_sort1 … H) -H //
-|
-|
-|
-|
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #HL12 #T2 #H
- elim (cprs_inv_appl1 … H) -H *
- [ #V2 #T0 #HV12 #HT10 #H destruct
- elim (nta_fwd_correct h L2 (ⓐV1.T1) (ⓐV1.U1) ?) [2: /3 width=2/ ] #U
- @(nta_conv … (ⓐV2.U1)) (* /2 width=1/*) [ /4 width=2/] (**) (* explicit constructor, /5 width=5/ is too slow *)
- | #V2 #W2 #T0 #HV12 #HT10 #HT02
- lapply (IHTU1 … HL12 (ⓛW2.T0) ?) -IHTU1 /2 width=1/ -HT10 #H
- elim (nta_inv_bind1 … H) -H #W #U0 #HW2 #HTU0 #HU01
- elim (cpcs_inv_abst1 … HU01) -HU01 #W #U #HU1 #HU0
- lapply (IHUW1 … HL12 (ⓐV2.ⓛW.U) ?) -IHUW1 -HL12 /2 width=1/ -HV12 #H
-
-
-
- elim (nta_fwd_pure1 … H) -H #W0 #U2 #HVU2 #H #HW01
- elim (nta_inv_bind1 … H) -H #W3 #U3 #HW3 #HU3 #H
- elim (cpcs_inv_abst1 … H) -H #W4 #U4
-*)
-(*
-axiom nta_ltpr_tpr_conf: ∀h,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U → ∀L2. L1 ➡ L2 →
- ∀T2. T1 ➡ T2 → ⦃h, L2⦄ ⊢ T2 : U.
-#h #L1 #T1 #U #H @(nta_ind_alt … H) -L1 -T1 -U
-[ #L1 #k #L2 #_ #T2 #H
- >(tpr_inv_atom1 … H) -H //
-| #L1 #K1 #V1 #W #U #i #HLK1 #_ #HWU #IHV1 #L2 #HL12 #T2 #H
- >(tpr_inv_atom1 … H) -T2
- elim (ltpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #HLK2 #H
- elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct /3 width=6/
-| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #HL12 #T2 #H
- >(tpr_inv_atom1 … H) -T2
- elim (ltpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #HLK2 #H
- elim (ltpr_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
- elim (lift_total V1 0 (i+1)) #W #HW
- lapply (nta_lift h … HLK … HWU1 … HW) /2 width=1/ -HLK -HW
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpr_lift … HW12 … HWU1 … HWU2) -HWU1 #HU12
- @(nta_conv … U2) /2 width=1/ /3 width=6/ (**) (* explicit constructor, /3 width=6/ is too slow *)
-| #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #HL12 #X #H
- elim (tpr_inv_bind1 … H) -H *
- [ #V2 #T0 #T2 #HV12 #HT10 #HT02 #H destruct
- lapply (IHVW1 … HL12 … HV12) #HV2W1
- lapply (IHVW1 L2 … V1 ?) // -IHVW1 #HWV1
- lapply (IHTU1 (L2.ⓑ{I}V2) … HT10) -HT10 /2 width=1/ #HT0U1
- lapply (IHTU1 (L2.ⓑ{I}V1) ? T1 ?) -IHTU1 // /2 width=1/ -HL12 #H
- lapply (tps_lsubs_trans … HT02 (L2.ⓑ{I}V2) ?) -HT02 /2 width=1/ #HT02
- lapply (nta_tps_conf … HT0U1 … HT02) -T0 #HT2U1
- elim (nta_fwd_correct … H) -H #U2 #HU12
- @(nta_conv … (ⓑ{I}V2.U1)) /2 width=2/ /3 width=1/ (**) (* explicit constructor, /4 width=6/ is too slow *)
- | #T #HT1 #HTX #H destruct
- lapply (IHVW1 … HL12 V1 ?) -IHVW1 // #HVW1
- elim (lift_total X 0 1) #Y #HXY
- lapply (tpr_lift … HTX … HT1 … HXY) -T #H
- lapply (IHTU1 (L2.ⓓV1) … H) -T1 /2 width=1/ -L1 #H
- elim (nta_fwd_correct … H) #T1 #HUT1
- elim (nta_thin_conf … H L2 0 (0+1) ? ?) -H /2 width=1/ /3 width=1/ #T #U #HTU #H
- normalize in ⊢ (??%??? → ?); #HU1
- lapply (delift_inv_lift1_eq … H L2 … HXY) -Y /2 width=1/ #H destruct
- @(nta_conv … U) // /2 width=2/
- ]
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #HL12 #X #H
- elim (tpr_inv_appl1 … H) -H *
- [ #V2 #Y #HV12 #HY #H destruct
- elim (tpr_inv_abst1 … HY) -HY #W2 #T2 #HW12 #HT12 #H destruct
- lapply (IHTU1 L2 ? (ⓛW1.T1) ?) // #H
- elim (nta_fwd_correct … H) -H #X #H
- elim (nta_inv_bind1 … H) -H #W #U #HW #HU #_
- @(nta_conv … (ⓐV2.ⓛW1.U1)) /4 width=2/ (**) (* explicit constructor, /5 width=5/ is too slow *)
- | #V2 #W2 #T0 #T2 #HV12 #HT02 #H1 #H2 destruct
- lapply (IHVW1 … HL12 … HV12) #HVW2
- lapply (IHVW1 … HL12 V1 ?) -IHVW1 // #HV1W2
- lapply (IHTU1 … HL12 (ⓛW2.T2) ?) -IHTU1 -HL12 /2 width=1/ -HT02 #H1
- elim (nta_fwd_correct … H1) #T #H2
- elim (nta_inv_bind1 … H1) -H1 #W #U2 #HW2 #HTU2 #H
- elim (cpcs_inv_abst … H Abst W2) -H #_ #HU21
- elim (nta_inv_bind1 … H2) -H2 #W0 #U0 #_ #H #_ -T -W0
- lapply (lsubn_nta_trans … HTU2 (L2.ⓓV2) ?) -HTU2 /2 width=1/ #HTU2
- @(nta_conv … (ⓓV2.U2)) /2 width=2/ /3 width=2/ (**) (* explicit constructor, /4 width=5/ is too slow *)
- | #V0 #V2 #W0 #W2 #T0 #T2 #_ #_ #_ #_ #H destruct
- ]
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #HL12 #X #H
- elim (tpr_inv_appl1 … H) -H *
- [ #V2 #T2 #HV12 #HT12 #H destruct
- elim (nta_fwd_correct h L2 (ⓐV1.T1) (ⓐV1.U1) ?) [2: /3 width=2/ ] #U
- @(nta_conv … (ⓐV2.U1)) /2 width=1/ /4 width=2/ (**) (* explicit constructor, /5 width=5/ is too slow *)
- | #V2 #W2 #T0 #T2 #HV12 #HT02 #H1 #H2 destruct
- lapply (IHTU1 … HL12 (ⓛW2.T2) ?) -IHTU1 /2 width=1/ -T0 #H
- elim (nta_inv_bind1 … H) -H #W #U2 #HW2 #HTU2 #HU21
- lapply (IHUW1 … HL12 (ⓐV2.U1) ?) -IHUW1 -HL12 /2 width=1/ #H
- elim (nta_inv_pure1 … H) -H #V0 #U0 #U #HV20 #HU10 #HU0W1 #HU0
- @(nta_conv … (ⓓV2.U2))
- [2: @nta_bind //
- @(lsubn_nta_trans … HTU2) @lsubn_abbr //
-(*
- lapply (IH … HV1 … HL12 … HV12) -HV1 -HV12 /width=5/ #HB
- lapply (IH … HB0 … HL12 W2 ?) -HB0 /width=5/ #HB0
- lapply (IH … HA0 … (L2.ⓛW2) … HT02) -IH -HA0 -HT02 /width=5/ -T0 /2 width=1/ -L1 -V1 /4 width=7/
-*)
-*)
-(*
-axiom pippo: ⦃h, L⦄ ⊢ ⓐV.X : Y →
- ∃∃W,T. L ⊢ X ➡* ⓛW.T & ⦃h, L⦄ ⊢ ⓐV : W.
-
-*)
-(* SEGMENT 2
-| #L1 #T1 #U1 #W1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_flat1 … H) -H #U2 #T2 #HU12 #HT12 #H destruct
- lapply (cpr_tpss … HU12) /4 width=4/
-| #L1 #T1 #U11 #U12 #U #_ #HU112 #_ #IHTU11 #IHU12 #L2 #d #e #HL12 #T2 #HT12
- @(nta_conv … U11) /2 width=5/ (**) (* explicot constructor, /3 width=7/ is too slow *)
-]
-qed.
-*)
-
-(* SEGMENT 3
-fact nta_ltpr_tpr_conf_aux: ∀h,L,T,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U → L = L1 → T = T1 →
- ∀L2. L1 ➡ L2 → ∀T2. T1 ➡ T2 → ⦃h, L2⦄ ⊢ T2 : U.
-
-
- | #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HW02 #HT02 #HV02 #H1 #H2 destruct
- elim (nta_inv_abbr … HT1) -HT1 #B0 #HW0 #HT0
- lapply (IH … HW0 … HL12 … HW02) -HW0 /width=5/ #HW2
- lapply (IH … HV1 … HL12 … HV10) -HV1 -HV10 /width=5/ #HV0
- lapply (IH … HT0 … (L2.ⓓW2) … HT02) -IH -HT0 -HT02 /width=5/ -V1 -T0 /2 width=1/ -L1 -W0 #HT2
- @(nta_abbr … HW2) -HW2
- @(nta_appl … HT2) -HT2 /3 width=7/ (**) (* explict constructors, /5 width=7/ is too slow *)
- ]
-| #L1 #V1 #T1 #A #HV1 #HT1 #H1 #H2 #L2 #HL12 #X #H destruct
- elim (tpr_inv_cast1 … H) -H
- [ * #V2 #T2 #HV12 #HT12 #H destruct
- lapply (IH … HV1 … HL12 … HV12) -HV1 -HV12 /width=5/ #HV2
- lapply (IH … HT1 … HL12 … HT12) -IH -HT1 -HL12 -HT12 /width=5/ -L1 -V1 -T1 /2 width=1/
- | -HV1 #HT1X
- lapply (IH … HT1 … HL12 … HT1X) -IH -HT1 -HL12 -HT1X /width=5/
- ]
-]
-qed.
-
-/2 width=9/ qed.
-
-axiom nta_ltpr_conf: ∀L1,T,A. L1 ⊢ T : A → ∀L2. L1 ➡ L2 → L2 ⊢ T : A.
-/2 width=5/ qed.
-
-axiom nta_tpr_conf: ∀L,T1,A. L ⊢ T1 : A → ∀T2. T1 ➡ T2 → L ⊢ T2 : A.
-/2 width=5/ 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/equivalence/cpcs_ltpss.ma".
-include "basic_2/dynamic/nta_nta.ma".
-
-(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Properties about parallel unfold *****************************************)
-
-lemma nta_ltpss_tpss_conf: ∀h,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → ⦃h, L2⦄ ⊢ T2 : U.
-#h #L1 #T1 #U #H @(nta_ind_alt … H) -L1 -T1 -U
-[ #L1 #k #L2 #d #e #_ #T2 #H
- >(tpss_inv_sort1 … H) -H //
-| #L1 #K1 #V1 #W #U #i #HLK1 #_ #HWU #IHV1 #L2 #d #e #HL12 #T2 #H
- elim (tpss_inv_lref1 … H) -H
- [ #H destruct
- elim (lt_or_ge i d) #Hdi
- [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #V2 #HK12 #HV12 #H destruct
- /3 width=7/
- | elim (lt_or_ge i (d + e)) #Hide [ | -Hdi ]
- [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #V2 #HK12 #HV12 #H destruct
- /3 width=7/
- | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=7/
- ]
- ]
- | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2
- elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #HK12 #HV12 #H destruct
- lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
- lapply (tpss_trans_eq … HV12 HVW2) -V2 /3 width=9/
- ]
-| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL12 #T2 #H
- elim (tpss_inv_lref1 … H) -H [ | -HWV1 -HWU1 -IHWV1 ]
- [ #H destruct
- elim (lift_total V1 0 (i+1)) #W #HW
- elim (lt_or_ge i d) #Hdi [ -HWV1 ]
- [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #W2 #HK12 #HW12 #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
- lapply (nta_lift h … HLK … HWU1 … HW) /2 width=4/ -HW
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) -HLK -HWU1 // #HU12
- lapply (cpr_tpss … HU12) -HU12 #HU12
- @(nta_conv … U2) /2 width=1/ /3 width=6/ (**) (* explicit constructor, /4 width=6/ is too slow *)
- | elim (lt_or_ge i (d + e)) #Hide [ -HWV1 | -IHWV1 -HW -Hdi ]
- [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #W2 #HK12 #HW12 #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
- lapply (nta_lift h … HLK … HWU1 … HW) /2 width=4/ -HW
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) -HLK -HWU1 // #HU12
- lapply (cpr_tpss … HU12) -HU12 #HU12
- @(nta_conv … U2) /2 width=1/ /3 width=6/ (**) (* explicit constructor, /4 width=6/ is too slow *)
- | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /2 width=6/
- ]
- ]
- | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #_ #_
- elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #_ #_ #H destruct
- lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 -HLK2 #H destruct
- ]
-| #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- lapply (cpr_tpss … HV12) #HV
- lapply (IHTU1 (L2.ⓑ{I}V1) (d+1) e ? T1 ?) // /2 width=1/ #H
- elim (nta_fwd_correct … H) -H #U2 #HU12
- @(nta_conv … (ⓑ{I}V2.U1)) /3 width=2/ /3 width=4/ /4 width=4/ (**) (* explicit constructor, /5 width=6/ is too slow *)
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_flat1 … H) -H #V2 #Y #HV12 #HY #H destruct
- elim (tpss_inv_bind1 … HY) -HY #W2 #T2 #HW12 #HT12 #H destruct
- lapply (cpr_tpss … HV12) #HV
- lapply (IHTU1 L2 d e ? (ⓛW1.T1) ?) // #H
- elim (nta_fwd_correct … H) -H #X #H
- elim (nta_inv_bind1 … H) -H #W #U #HW #HU #_
- @(nta_conv … (ⓐV2.ⓛW1.U1)) /3 width=2/ /3 width=4/ /4 width=5/ (**) (* explicit constructor, /5 width=5/ is too slow *)
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- lapply (cpr_tpss … HV12) #HV
- elim (nta_fwd_correct h L2 (ⓐV1.T1) (ⓐV1.U1) ?) [2: /3 width=4/ ] #U #HU
- @(nta_conv … (ⓐV2.U1)) // /3 width=1/ /4 width=5/ (**) (* explicit constructor, /5 width=5/ is too slow *)
-| #L1 #T1 #U1 #W1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_flat1 … H) -H #U2 #T2 #HU12 #HT12 #H destruct
- lapply (cpr_tpss … HU12) /4 width=4/
-| #L1 #T1 #U11 #U12 #U #_ #HU112 #_ #IHTU11 #IHU12 #L2 #d #e #HL12 #T2 #HT12
- @(nta_conv … U11) /2 width=5/ (**) (* explicot constructor, /3 width=7/ is too slow *)
-]
-qed.
-
-lemma nta_ltpss_tps_conf: ∀h,L1,T1,U. ⦃h, L1⦄ ⊢ T1 : U →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶ [d, e] T2 → ⦃h, L2⦄ ⊢ T2 : U.
-/3 width=7/ qed.
-
-lemma nta_ltpss_conf: ∀h,L1,T,U. ⦃h, L1⦄ ⊢ T : U →
- ∀L2,d,e. L1 ▶* [d, e] L2 → ⦃h, L2⦄ ⊢ T : U.
-/2 width=7/ qed.
-
-lemma nta_tpss_conf: ∀h,L,T1,U. ⦃h, L⦄ ⊢ T1 : U →
- ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 → ⦃h, L⦄ ⊢ T2 : U.
-/2 width=7/ qed.
-
-lemma nta_tps_conf: ∀h,L,T1,U. ⦃h, L⦄ ⊢ T1 : U →
- ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ⦃h, L⦄ ⊢ T2 : U.
-/2 width=7/ 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/static/sta.ma".
-include "basic_2/equivalence/cpcs_cpcs.ma".
-include "basic_2/dynamic/nta.ma".
-
-(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Properties on static type assignment *************************************)
-
-lemma nta_fwd_sta: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U →
- ∃∃U0. ⦃h, L⦄ ⊢ T • U0 & L ⊢ U0 ⬌* U.
-#h #L #T #U #H elim H -L -T -U
-[ /2 width=3/
-| #L #K #V #W1 #V1 #i #HLK #_ #HWV1 * #W0 #HVW0 #HW01
- elim (lift_total W0 0 (i+1)) #V0 #HWV0
- lapply (ldrop_fwd_ldrop2 … HLK) #HLK0
- lapply (cpcs_lift … HLK0 … HWV0 … HWV1 HW01) -HLK0 -HWV1 -HW01 /3 width=6/
-| #L #K #W #V1 #W1 #i #HLK #HWV1 #HW1 * /3 width=6/
-| #I #L #V #W #T #U #_ #_ * #W0 #_ #_ * /3 width=3/
-| #L #V #W #T #U #_ #_ * #W0 #_ #HW0 * #X #H #HX
- elim (sta_inv_bind1 … H) -H #U0 #HTU0 #H destruct
- @(ex2_1_intro … (ⓐV.ⓛW.U0)) /2 width=1/ /3 width=1/
-| #L #V #W #T #U #_ #_ * #U0 #HTU0 #HU0 #_ -W
- @(ex2_1_intro … (ⓐV.U0)) /2 width=1/
-| #L #T #U #HTU * #U0 #HTU0 #HU0 /3 width=3/
-| #L #T #U1 #U2 #V2 #_ #HU12 #_ * #U0 #HTU0 #HU01 #_
- lapply (cpcs_trans … HU01 … HU12) -U1 /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/unfold/thin_ldrop.ma".
-include "basic_2/equivalence/cpcs_delift.ma".
-include "basic_2/dynamic/nta_lift.ma".
-
-(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Properties on basic local environment thinning ***************************)
-
-(* Note: this is known as the substitution lemma *)
-(* Basic_1: was only: ty3_gen_cabbr *)
-lemma nta_thin_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 →
- ∀L2,d,e. ≽ [d, e] L1 → L1 ▼*[d, e] ≡ L2 →
- ∃∃T2,U2. ⦃h, L2⦄ ⊢ T2 : U2 &
- L1 ⊢ T1 ▼*[d, e] ≡ T2 & L1 ⊢ U1 ▼*[d, e] ≡ U2.
-#h #L1 #T1 #U1 #H elim H -L1 -T1 -U1
-[ /2 width=5/
-| #L1 #K1 #V1 #W1 #U1 #i #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL1 #HL12
- elim (lt_or_ge i d) #Hdi [ -HVW1 ]
- [ lapply (sfr_ldrop_trans_ge … HLK1 … HL1 ?) -HL1 /2 width=2/ #H
- lapply (sfr_inv_skip … H ?) -H /2 width=1/ #HK1
- elim (thin_ldrop_conf_le … HL12 … HLK1 ?) -HL12 /2 width=2/ #X #H #HLK2
- elim (thin_inv_delift1 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
- elim (IHVW1 … HK1 HK12) -IHVW1 -HK1 -HK12 #X2 #W2 #HVW2 #H #HW12
- lapply (delift_mono … H … HV12) -H -HV12 #H destruct
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (ldrop_fwd_ldrop2 … HLK1) -V1 #HLK1
- lapply (delift_lift_ge … HW12 … HLK1 HWU1 … HWU2) -HW12 -HLK1 -HWU1 //
- >minus_plus <plus_minus_m_m // /3 width=6/
- | elim (lt_or_ge i (d+e)) #Hide [ -HVW1 | -Hdi -IHVW1 -HL1 ]
- [ lapply (sfr_ldrop_trans_be_up … HLK1 … HL1 ? ?) -HL1 // /2 width=2/ <minus_n_O #H
- elim (sfr_inv_bind … H ?) -H /2 width=1/ #HK1 #_
- elim (thin_ldrop_conf_be … HL12 … HLK1 ? ?) -HL12 /2 width=2/ #K2 #H #HLK2
- lapply (thin_inv_thin1 … H ?) -H /2 width=1/ #HK12
- elim (IHVW1 … HK1 HK12) -IHVW1 -HK1 -HK12 #V2 #W2 #HVW2 #HV12 #HW12
- elim (lift_total V2 0 d) #T2 #HVT2
- elim (lift_total W2 0 d) #U2 #HWU2
- elim (lift_total W2 0 (i+1)) #U #HW2U
- lapply (nta_lift … HVW2 … HLK2 … HVT2 … HWU2) -HVW2 -HLK2 #HTU2
- lapply (ldrop_fwd_ldrop2 … HLK1) #HLK0
- lapply (delift_lift_ge … HW12 … HLK0 HWU1 … HW2U) -HW12 -HLK0 -HWU1 // >minus_plus #HU1
- lapply (lift_conf_be … HWU2 … HW2U ?) -W2 /2 width=1/ #HU2
- lapply (delift_lift_div_be … HU1 … HU2 ? ?) -U // /2 width=1/ /3 width=8/
- | lapply (transitive_le … (i+1) Hide ?) /2 width=1/ #Hdei
- lapply (thin_ldrop_conf_ge … HL12 … HLK1 ?) -HL12 -HLK1 // #HL2K1
- elim (lift_split … HWU1 d (i+1-e) ? ? ?) -HWU1 // /2 width=1/ #W
- <plus_minus in ⊢ (??%??→?); /2 width=2/ #HW1
- <minus_minus // /2 width=2/ -Hdei >commutative_plus <minus_n_n /3 width=6/
- ]
- ]
-| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL1 #HL12
- elim (lt_or_ge i d) #Hdi [ -HWV1 | -IHWV1 ]
- [ lapply (sfr_ldrop_trans_ge … HLK1 … HL1 ?) -HL1 /2 width=2/ #H
- lapply (sfr_inv_skip … H ?) -H /2 width=1/ #HK1
- elim (thin_ldrop_conf_le … HL12 … HLK1 ?) -HL12 /2 width=2/ #X #H #HLK2
- elim (thin_inv_delift1 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
- elim (IHWV1 … HK1 HK12) -IHWV1 -HK1 -HK12 #X2 #V2 #HWV2 #H #_
- lapply (delift_mono … H … HW12) -H #H destruct
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 #HLK1
- lapply (delift_lift_ge … HW12 … HLK1 HWU1 … HWU2) -HW12 -HLK1 -HWU1 //
- >minus_plus <plus_minus_m_m // /3 width=6/
- | elim (lt_or_ge i (d+e)) #Hide [ -HWV1 -HWU1 -HL12 | -Hdi -HL1 ]
- [ lapply (sfr_inv_ldrop … HLK1 … HL1 ? ?) -HLK1 -HL1 // -Hdi -Hide #H destruct
- | lapply (transitive_le … (i+1) Hide ?) /2 width=1/ #Hdei
- lapply (thin_ldrop_conf_ge … HL12 … HLK1 ?) -HL12 -HLK1 // #HL2K1
- elim (lift_split … HWU1 d (i+1-e) ? ? ?) -HWU1 // /2 width=1/ #W
- <plus_minus in ⊢ (??%??→?); /2 width=2/ #HW1
- <minus_minus // /2 width=2/ -Hdei >commutative_plus <minus_n_n /3 width=6/
- ]
- ]
-| #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL1 #HL12
- elim (IHVW1 … HL1 HL12) -IHVW1 #V2 #W2 #HVW2 #HV12 #_
- elim (IHTU1 (L2.ⓑ{I}V2) (d+1) e ? ?) -IHTU1 /2 width=1/ -HL1 -HL12 #T2 #U2 #HTU2 #HT12 #HU12
- lapply (delift_lsubs_trans … HT12 (L1.ⓑ{I}V2) ?) -HT12 /2 width=1/
- lapply (delift_lsubs_trans … HU12 (L1.ⓑ{I}V2) ?) -HU12 /2 width=1/ /3 width=7/
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL1 #HL12
- elim (IHVW1 … HL1 HL12) -IHVW1 #V2 #W2 #HVW2 #HV12 #HW12
- elim (IHTU1 … HL1 HL12) -IHTU1 -HL1 -HL12 #X2 #Y2 #HXY2 #HX2 #HY2
- elim (delift_inv_bind1 … HX2) -HX2 #Z21 #T2 #HZ21 #HT12 #H destruct
- elim (delift_inv_bind1 … HY2) -HY2 #Z22 #U2 #HZ22 #HU12 #H destruct
- lapply (delift_mono … HZ21 … HW12) -HZ21 #H destruct
- lapply (delift_mono … HZ22 … HW12) -HZ22 #H destruct
- @(ex3_2_intro … (ⓐV2.ⓛW2.T2) (ⓐV2.ⓛW2.U2)) /3 width=1/ (**) (* explict constructor, /4 depth=?/ is too slow *)
-| #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL1 #HL12
- elim (IHTU1 … HL1 HL12) -IHTU1 #T2 #U2 #HTU2 #HT12 #HU12
- elim (IHUW1 … HL1 HL12) -IHUW1 -HL1 -HL12 #X2 #W2 #HXW2 #H #HW12
- elim (delift_inv_flat1 … H) -H #V2 #Y2 #HV12 #HY2 #H destruct
- lapply (delift_mono … HY2 … HU12) -HY2 #H destruct /3 width=7/
-| #L1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL1 #HL12
- elim (IHTU1 … HL1 HL12) -IHTU1 -HL1 -HL12 /3 width=5/
-| #L1 #T1 #U11 #U12 #V1 #_ #HU112 #_ #IHT1 #IHU12 #L2 #d #e #HL1 #HL12
- elim (IHT1 … HL1 HL12) -IHT1 #T2 #U21 #HT2 #HT12 #HU121
- elim (IHU12 … HL1 HL12) -IHU12 -HL1 #U22 #V2 #HU22 #HU122 #_
- lapply (thin_cpcs_delift_mono … HU112 … HL12 … HU121 … HU122) -HU112 -HL12 -HU121 /3 width=5/
-]
-qed.
-
-lemma nta_ldrop_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 →
- ∀L2,d,e. ≽ [d, e] L1 → ⇩[d, e] L1 ≡ L2 →
- ∃∃T2,U2. ⦃h, L2⦄ ⊢ T2 : U2 &
- L1 ⊢ T1 ▼*[d, e] ≡ T2 & L1 ⊢ U1 ▼*[d, e] ≡ U2.
-/3 width=1/ qed.
--- /dev/null
+lemma da_cpcs: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ T1 ¡[h, g] →
+ ∀T2. ⦃G, L⦄ ⊢ T2 ¡[h, g] →
+ ∀l1. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ∀l2. ⦃G, L⦄ ⊢ T2 ▪[h, g] l2 →
+ ⦃G, L⦄ ⊢ T1 ⬌* T2 → l1 = l2.
+#h #g #G #L #T1 #HT1 #T2 #HT2 #l1 #Hl1 #l2 #Hl2 #H
+elim (cpcs_inv_cprs … H) -H /3 width=12 by da_cprs_lpr, da_mono/
+qed-.
+
+lemma sta_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l+1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •[h] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •[h] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+#h #g #G #L1 #T1 #HT1 #l #Hl #U1 #HTU1 #T2 #HT12 #L2 #HL12
+elim (lstas_cpr_lpr … 1 … Hl U1 … HT12 … HL12) -Hl -HT12 -HL12
+/3 width=3 by lstas_inv_SO, sta_lstas, ex2_intro/
+qed-.
+
+lemma snv_sta: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
+ ∀l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 →
+ ∀U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ U ¡[h, g].
+/3 width=7 by lstas_inv_SO, sta_lstas, snv_lstas/ qed-.
+
+lemma lstas_cpds: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ T1 ¡[h, g] →
+ ∀l1,l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
+ ∀U1. ⦃G, L⦄ ⊢ T1 •*[h, l2] U1 → ∀T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 →
+ ∃∃U2,l. l ≤ l2 & ⦃G, L⦄ ⊢ T2 •*[h, l] U2 & ⦃G, L⦄ ⊢ U1 •*⬌*[h, g] U2.
+#h #g #G #L #T1 #HT1 #l1 #l2 #Hl21 #Hl1 #U1 #HTU1 #T2 * #T #l0 #l #Hl0 #H #HT1T #HTT2
+lapply (da_mono … H … Hl1) -H #H destruct
+lapply (lstas_da_conf … HTU1 … Hl1) #Hl12
+elim (le_or_ge l2 l) #Hl2
+[ lapply (lstas_conf_le … HTU1 … HT1T) -HT1T //
+ /5 width=11 by cpds_cpes_dx, monotonic_le_minus_l, ex3_2_intro, ex4_3_intro/
+| lapply (lstas_da_conf … HT1T … Hl1) #Hl1l
+ lapply (lstas_conf_le … HT1T … HTU1) -HTU1 // #HTU1
+ elim (lstas_cprs_lpr … Hl1l … HTU1 … HTT2 L) -Hl1l -HTU1 -HTT2
+ /3 width=7 by snv_lstas, cpcs_cpes, monotonic_le_minus_l, ex3_2_intro/
+]
+qed-.
+
+lemma cpds_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
+#h #g #G #L1 #T1 #HT1 #U1 * #W1 #l1 #l2 #Hl21 #Hl1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
+elim (lstas_cpr_lpr … Hl1 … HTW1 … HT12 … HL12) // #W2 #HTW2 #HW12
+lapply (da_cpr_lpr … Hl1 … HT12 … HL12) // -T1
+lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
+lapply (cpcs_canc_sn … HW12 HWU1) -W1 #H
+elim (cpcs_inv_cprs … H) -H /3 width=7 by ex4_3_intro, ex2_intro/
+qed-.
+
+(* Note: missing da_scpds_lpr, da_scpes *)
+
+lemma scpds_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀U1,l. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, l] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g, l] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
+#h #g #G #L1 #T1 #HT1 #U1 #l2 * #W1 #l1 #Hl21 #HTl1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
+elim (lstas_cpr_lpr … HTl1 … HTW1 … HT12 … HL12) // #W2 #HTW2 #HW12
+lapply (da_cpr_lpr … HTl1 … HT12 … HL12) // -T1
+lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
+lapply (cpcs_canc_sn … HW12 HWU1) -W1 #H
+elim (cpcs_inv_cprs … H) -H /3 width=6 by ex4_2_intro, ex2_intro/
+qed-.
+
+lemma scpes_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀T2. ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
+ ∀l1,l2. ⦃G, L1⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 ➡ U1 → ∀U2. ⦃G, L1⦄ ⊢ T2 ➡ U2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L2⦄ ⊢ U1 •*⬌*[h, g, l1, l2] U2.
+#h #g #G #L1 #T1 #HT1 #T2 #HT2 #l1 #l2 * #T0 #HT10 #HT20 #U1 #HTU1 #U2 #HTU2 #L2 #HL12
+elim (scpds_cpr_lpr … HT10 … HTU1 … HL12) -HT10 -HTU1 // #X1 #HUX1 #H1
+elim (scpds_cpr_lpr … HT20 … HTU2 … HL12) -HT20 -HTU2 // #X2 #HUX2 #H2
+elim (cprs_conf … H1 … H2) -T0 /3 width=5 by scpds_div, scpds_cprs_trans/
+qed-.
+
+(* Note: missing lstas_scpds, scpes_le *)
+++ /dev/null
-lemma da_cpcs: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ T1 ¡[h, g] →
- ∀T2. ⦃G, L⦄ ⊢ T2 ¡[h, g] →
- ∀l1. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ∀l2. ⦃G, L⦄ ⊢ T2 ▪[h, g] l2 →
- ⦃G, L⦄ ⊢ T1 ⬌* T2 → l1 = l2.
-#h #g #G #L #T1 #HT1 #T2 #HT2 #l1 #Hl1 #l2 #Hl2 #H
-elim (cpcs_inv_cprs … H) -H /3 width=12 by da_cprs_lpr, da_mono/
-qed-.
-
-lemma sta_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l+1 →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •[h] U1 →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •[h] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L1 #T1 #HT1 #l #Hl #U1 #HTU1 #T2 #HT12 #L2 #HL12
-elim (lstas_cpr_lpr … 1 … Hl U1 … HT12 … HL12) -Hl -HT12 -HL12
-/3 width=3 by lstas_inv_SO, sta_lstas, ex2_intro/
-qed-.
-
-lemma snv_sta: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 →
- ∀U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ U ¡[h, g].
-/3 width=7 by lstas_inv_SO, sta_lstas, snv_lstas/ qed-.
-
-lemma lstas_cpds: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ T1 ¡[h, g] →
- ∀l1,l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
- ∀U1. ⦃G, L⦄ ⊢ T1 •*[h, l2] U1 → ∀T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 →
- ∃∃U2,l. l ≤ l2 & ⦃G, L⦄ ⊢ T2 •*[h, l] U2 & ⦃G, L⦄ ⊢ U1 •*⬌*[h, g] U2.
-#h #g #G #L #T1 #HT1 #l1 #l2 #Hl21 #Hl1 #U1 #HTU1 #T2 * #T #l0 #l #Hl0 #H #HT1T #HTT2
-lapply (da_mono … H … Hl1) -H #H destruct
-lapply (lstas_da_conf … HTU1 … Hl1) #Hl12
-elim (le_or_ge l2 l) #Hl2
-[ lapply (lstas_conf_le … HTU1 … HT1T) -HT1T //
- /5 width=11 by cpds_cpes_dx, monotonic_le_minus_l, ex3_2_intro, ex4_3_intro/
-| lapply (lstas_da_conf … HT1T … Hl1) #Hl1l
- lapply (lstas_conf_le … HT1T … HTU1) -HTU1 // #HTU1
- elim (lstas_cprs_lpr … Hl1l … HTU1 … HTT2 L) -Hl1l -HTU1 -HTT2
- /3 width=7 by snv_lstas, cpcs_cpes, monotonic_le_minus_l, ex3_2_intro/
-]
-qed-.
-
-lemma cpds_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g] U1 →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
-#h #g #G #L1 #T1 #HT1 #U1 * #W1 #l1 #l2 #Hl21 #Hl1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
-elim (lstas_cpr_lpr … Hl1 … HTW1 … HT12 … HL12) // #W2 #HTW2 #HW12
-lapply (da_cpr_lpr … Hl1 … HT12 … HL12) // -T1
-lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
-lapply (cpcs_canc_sn … HW12 HWU1) -W1 #H
-elim (cpcs_inv_cprs … H) -H /3 width=7 by ex4_3_intro, ex2_intro/
-qed-.
-
-(* Note: missing da_scpds_lpr, da_scpes *)
-
-lemma scpds_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1,l. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, l] U1 →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g, l] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
-#h #g #G #L1 #T1 #HT1 #U1 #l2 * #W1 #l1 #Hl21 #HTl1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
-elim (lstas_cpr_lpr … HTl1 … HTW1 … HT12 … HL12) // #W2 #HTW2 #HW12
-lapply (da_cpr_lpr … HTl1 … HT12 … HL12) // -T1
-lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
-lapply (cpcs_canc_sn … HW12 HWU1) -W1 #H
-elim (cpcs_inv_cprs … H) -H /3 width=6 by ex4_2_intro, ex2_intro/
-qed-.
-
-lemma scpes_cpr_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀T2. ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
- ∀l1,l2. ⦃G, L1⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2 →
- ∀U1. ⦃G, L1⦄ ⊢ T1 ➡ U1 → ∀U2. ⦃G, L1⦄ ⊢ T2 ➡ U2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ U1 •*⬌*[h, g, l1, l2] U2.
-#h #g #G #L1 #T1 #HT1 #T2 #HT2 #l1 #l2 * #T0 #HT10 #HT20 #U1 #HTU1 #U2 #HTU2 #L2 #HL12
-elim (scpds_cpr_lpr … HT10 … HTU1 … HL12) -HT10 -HTU1 // #X1 #HUX1 #H1
-elim (scpds_cpr_lpr … HT20 … HTU2 … HL12) -HT20 -HTU2 // #X2 #HUX2 #H2
-elim (cprs_conf … H1 … H2) -T0 /3 width=5 by scpds_div, scpds_cprs_trans/
-qed-.
-
-(* Note: missing lstas_scpds, scpes_le *)
class "magenta"
[ { "dynamic typing" * } {
[ { "context-sensitive native type assignment" * } {
- [ [ "for terms" ] "nta" + "( ⦃?,?⦄ ⊢ ? :[?,?] ? )" + "( ⦃?,?⦄ ⊢ ? :[?] ? )" + "( ⦃?,?⦄ ⊢ ? :*[?] ? )" "nta_drops" + "nta_cpms" + "nta_cpcs" + "nta_preserve" + "nta_ind" * ]
+ [ [ "for terms" ] "nta" + "( ⦃?,?⦄ ⊢ ? :[?,?] ? )" + "( ⦃?,?⦄ ⊢ ? :[?] ? )" + "( ⦃?,?⦄ ⊢ ? :*[?] ? )" "nta_drops" + "nta_aaa" + "nta_fsb" + "nta_cpms" + "nta_cpcs" + "nta_preserve" + "nta_preserve_cpcs" + "nta_ind" * ]
}
]
[ { "context-sensitive native validity" * } {