X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fsubstitution%2Fldrop.ma;h=29759fe08806ebc0a51ca3633218d880686d0d0f;hb=fec1a061eeca5e7e05b4f0c3e299983b163569c3;hp=73a9ff92e373fc9bb0d3aee04849989b25191576;hpb=913512bbc9202f2109d53acd43dc8c0270b17184;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma index 73a9ff92e..29759fe08 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma @@ -13,8 +13,8 @@ (**************************************************************************) include "basic_2/grammar/cl_weight.ma". -include "basic_2/grammar/lsubs.ma". include "basic_2/substitution/lift.ma". +include "basic_2/substitution/lsubs.ma". (* LOCAL ENVIRONMENT SLICING ************************************************) @@ -76,6 +76,12 @@ lemma ldrop_inv_O1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → (0 < e ∧ ⇩[0, e - 1] K ≡ L2). /2 width=3/ qed-. +lemma ldrop_inv_pair1: ∀K,I,V,L2. ⇩[0, 0] K. ⓑ{I} V ≡ L2 → L2 = K. ⓑ{I} V. +#K #I #V #L2 #H +elim (ldrop_inv_O1 … H) -H * // #H destruct +elim (lt_refl_false … H) +qed-. + (* Basic_1: was: drop_gen_drop *) lemma ldrop_inv_ldrop1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → 0 < e → ⇩[0, e - 1] K ≡ L2. @@ -142,13 +148,13 @@ lemma ldrop_O1: ∀L,i. i < |L| → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V. lapply (lt_plus_to_lt_l … H) -H #Hi elim (IHL i ?) // /3 width=4/ ] -qed. +qed. -lemma ldrop_lsubs_ldrop1_abbr: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → - ∀K1,V,i. ⇩[0, i] L1 ≡ K1. ⓓV → +lemma ldrop_lsubs_ldrop2_abbr: ∀L1,L2,d,e. L1 ≼ [d, e] L2 → + ∀K2,V,i. ⇩[0, i] L2 ≡ K2. ⓓV → d ≤ i → i < d + e → - ∃∃K2. K1 [0, d + e - i - 1] ≼ K2 & - ⇩[0, i] L2 ≡ K2. ⓓV. + ∃∃K1. K1 ≼ [0, d + e - i - 1] K2 & + ⇩[0, i] L1 ≡ K1. ⓓV. #L1 #L2 #d #e #H elim H -L1 -L2 -d -e [ #d #e #K1 #V #i #H lapply (ldrop_inv_atom1 … H) -H #H destruct