X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fsubstitution%2Ftps.ma;h=0bde244171dedc979a8d303b1785f43b45ed3ac2;hb=99c7be7031e506c7ad4a6c5e3f12ad5ae542b049;hp=db116f77e63b526eb2dc408ec50dfb645a0b9327;hpb=439b6ec33d749ba4e6ae0938e973a85bc23e306e;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma index db116f77e..0bde24417 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/substitution/ldrop.ma". +include "basic_2/substitution/ldrop_append.ma". (* PARALLEL SUBSTITUTION ON TERMS *******************************************) @@ -144,6 +144,15 @@ lemma tps_split_down: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ] qed. +lemma tps_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶ [d, e] T2 → + ∀L. L @@ K ⊢ T1 ▶ [d, e] T2. +#K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e // /2 width=1/ +#K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L +lapply (ldrop_fwd_ldrop2_length … HK0) #H +@(tps_subst … (L@@K0) … HVW) // (**) (* /3/ does not work *) +@(ldrop_O1_append_sn_le … HK0) /2 width=2/ +qed. + (* Basic inversion lemmas ***************************************************) fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀I. T1 = ⓪{I} →