X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda-delta%2FBasic-2%2Fsubstitution%2Fdrop_drop.ma;h=9343748537ebf0b9c1a685e0e46da83b93ee544f;hb=55dc00c1c44cc21c7ae179cb9df03e7446002c46;hp=03a8e31ddaf9a7a1cf856f609c7bec2bff3359dd;hpb=b24e4faf4501e54da29dc70940101eeb160e9c9f;p=helm.git diff --git a/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop_drop.ma b/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop_drop.ma index 03a8e31dd..934374853 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop_drop.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop_drop.ma @@ -19,13 +19,13 @@ include "Basic-2/substitution/drop.ma". (* Main properties **********************************************************) +(* Basic-1: was: drop_mono *) theorem drop_mono: ∀d,e,L,L1. ↓[d, e] L ≡ L1 → ∀L2. ↓[d, e] L ≡ L2 → L1 = L2. #d #e #L #L1 #H elim H -H d e L L1 [ #d #e #L2 #H - >(drop_inv_sort1 … H) -H L2 // -| #K1 #K2 #I #V #HK12 #_ #L2 #HL12 - <(drop_inv_refl … HK12) -HK12 K2 + >(drop_inv_atom1 … H) -H L2 // +| #K #I #V #L2 #HL12 <(drop_inv_refl … HL12) -HL12 L2 // | #L #K #I #V #e #_ #IHLK #L2 #H lapply (drop_inv_drop1 … H ?) -H /2/ @@ -36,14 +36,14 @@ theorem drop_mono: ∀d,e,L,L1. ↓[d, e] L ≡ L1 → ] qed. +(* Basic-1: was: drop_conf_ge *) theorem drop_conf_ge: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 → ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 + e1 ≤ e2 → ↓[0, e2 - e1] L1 ≡ L2. #d1 #e1 #L #L1 #H elim H -H d1 e1 L L1 [ #d #e #e2 #L2 #H - >(drop_inv_sort1 … H) -H L2 // -| #K1 #K2 #I #V #HK12 #_ #e2 #L2 #H #_ (drop_inv_atom1 … H) -H L2 // +| // | #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 lapply (drop_inv_drop1 … H ?) -H /2/ #HL2 (drop_inv_sort1 … H) -H L2 /2/ -| #K1 #K2 #I #V #HK12 #_ #e2 #L2 #HL2 #H - >(drop_inv_refl … HK12) -HK12 K1; + >(drop_inv_atom1 … H) -H L2 /2/ +| #K #I #V #e2 #L2 #HL2 #H lapply (le_O_to_eq_O … H) -H #H destruct -e2 /2/ | #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H lapply (le_O_to_eq_O … H) -H #H destruct -e2; @@ -99,13 +100,13 @@ theorem drop_trans_le: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L → ] qed. +(* Basic-1: was: drop_trans_ge *) theorem drop_trans_ge: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L → ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 ≤ e2 → ↓[0, e1 + e2] L1 ≡ L2. #d1 #e1 #L1 #L #H elim H -H d1 e1 L1 L [ #d #e #e2 #L2 #H - >(drop_inv_sort1 … H) -H L2 // -| #K1 #K2 #I #V #HK12 #_ #e2 #L2 #H #_ normalize - >(drop_inv_refl … HK12) -HK12 K1 // + >(drop_inv_atom1 … H) -H L2 // +| // | /3/ | #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2 lapply (lt_to_le_to_lt 0 … Hde2) // #He2 @@ -121,5 +122,6 @@ theorem drop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L. #e1 #e1 #e2 >commutative_plus /2 width=5/ qed. +(* Basic-1: was: drop_conf_rev *) axiom drop_div: ∀e1,L1,L. ↓[0, e1] L1 ≡ L → ∀e2,L2. ↓[0, e2] L2 ≡ L → ∃∃L0. ↓[0, e1] L0 ≡ L2 & ↓[e1, e2] L0 ≡ L1.