X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsubstitution%2Fldrop.ma;h=87d83b3472819ccaea16949e8b61ab8afb4352d1;hb=cdfd45ca5a2b52601b7bde732a7811de55a52fed;hp=9511648aab81cad4032fb903dc4ef76ebd340216;hpb=e8998d29ab83e7b6aa495a079193705b2f6743d3;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/ldrop.ma index 9511648aa..87d83b347 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/ldrop.ma @@ -81,7 +81,7 @@ lemma ldrop_inv_atom1: ∀d,e,L2. ⇩[d, e] ⋆ ≡ L2 → L2 = ⋆. /2 width=5/ qed-. fact ldrop_inv_O1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → - ∀K,I,V. L1 = K. ⓑ{I} V → + ∀K,I,V. L1 = K. ⓑ{I} V → (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ (0 < e ∧ ⇩[d, e - 1] K ≡ L2). #d #e #L1 #L2 * -d -e -L1 -L2 @@ -114,7 +114,7 @@ qed-. fact ldrop_inv_skip1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & - ⇧[d - 1, e] V2 ≡ V1 & + ⇧[d - 1, e] V2 ≡ V1 & L2 = K2. ⓑ{I} V2. #d #e #L1 #L2 * -d -e -L1 -L2 [ #d #e #_ #I #K #V #H destruct @@ -127,14 +127,14 @@ qed. (* Basic_1: was: drop_gen_skip_l *) lemma ldrop_inv_skip1: ∀d,e,I,K1,V1,L2. ⇩[d, e] K1. ⓑ{I} V1 ≡ L2 → 0 < d → ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & - ⇧[d - 1, e] V2 ≡ V1 & + ⇧[d - 1, e] V2 ≡ V1 & L2 = K2. ⓑ{I} V2. /2 width=3/ qed-. fact ldrop_inv_skip2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & - ⇧[d - 1, e] V2 ≡ V1 & + ⇧[d - 1, e] V2 ≡ V1 & L1 = K1. ⓑ{I} V1. #d #e #L1 #L2 * -d -e -L1 -L2 [ #d #e #_ #I #K #V #H destruct @@ -171,7 +171,7 @@ qed. lemma ldrop_O1_le: ∀i,L. i ≤ |L| → ∃K. ⇩[0, i] L ≡ K. #i @(nat_ind_plus … i) -i /2 width=2/ #i #IHi * -[ #H lapply (le_n_O_to_eq … H) -H >commutative_plus normalize #H destruct +[ #H lapply (le_n_O_to_eq … H) -H >commutative_plus normalize #H destruct | #L #I #V normalize #H elim (IHi L ?) -IHi /2 width=1/ -H /3 width=2/ ] @@ -258,16 +258,16 @@ lemma ldrop_fwd_ldrop2: ∀L1,I2,K2,V2,e. ⇩[O, e] L1 ≡ K2. ⓑ{I2} V2 → ] qed-. -lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → #{L2} ≤ #{L1}. +lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}. #L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize [ /2 width=3/ | #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 >(tw_lift … HV21) -HV21 /2 width=1/ ] -qed-. +qed-. lemma ldrop_pair2_fwd_fw: ∀I,L,K,V,d,e. ⇩[d, e] L ≡ K. ⓑ{I} V → - ∀T. #{K, V} < #{L, T}. + ∀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/ @@ -298,7 +298,7 @@ lemma ldrop_fwd_O1_length: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → |L2| = |L1| - e. qed-. (* Basic_1: removed theorems 50: - drop_ctail drop_skip_flat + 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