(**************************************************************************) (* ___ *) (* ||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 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 commutative_plus minus_plus commutative_plus