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=e667c7cde14f83e07ea49bcf135663d1d324e621;hpb=a8c166f1e1baeeae04553058bd179420ada8bbe7;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 e667c7cde..29759fe08 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma @@ -12,9 +12,9 @@ (* *) (**************************************************************************) -include "basic_2/grammar/lenv_weight.ma". -include "basic_2/grammar/lsubs.ma". +include "basic_2/grammar/cl_weight.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. @@ -135,11 +141,20 @@ lemma ldrop_ldrop_lt: ∀L1,L2,I,V,e. #L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) // /2 width=1/ qed. -lemma ldrop_lsubs_ldrop1_abbr: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → - ∀K1,V,i. ⇩[0, i] L1 ≡ K1. ⓓV → +lemma ldrop_O1: ∀L,i. i < |L| → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V. +#L elim L -L +[ #i #H elim (lt_zero_false … H) +| #L #I #V #IHL #i @(nat_ind_plus … i) -i /2 width=4/ #i #_ #H + lapply (lt_plus_to_lt_l … H) -H #Hi + elim (IHL i ?) // /3 width=4/ +] +qed. + +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 @@ -187,6 +202,13 @@ lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → #[L2] ≤ #[L1]. ] qed-. +lemma ldrop_pair2_fwd_cw: ∀I,L,K,V,d,e. ⇩[d, e] L ≡ K. ⓑ{I} V → + ∀T. #[K, V] < #[L, T]. +#I #L #K #V #d #e #H #T +lapply (ldrop_fwd_lw … H) -H #H +@(le_to_lt_to_lt … H) -H /3 width=1/ +qed-. + lemma ldrop_fwd_ldrop2_length: ∀L1,I2,K2,V2,e. ⇩[0, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|. #L1 elim L1 -L1 @@ -211,8 +233,8 @@ lemma ldrop_fwd_O1_length: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → |L2| = |L1| - e. ] qed-. -(* Basic_1: removed theorems 49: - drop_skip_flat +(* Basic_1: removed theorems 50: + drop_ctail drop_skip_flat cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf drop_clear drop_clear_O drop_clear_S clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r