X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Flambda-delta%2Fsubstitution%2Fdrop_main.ma;h=a305b78c7cfce7854e6e36b1d9e7deb79c398b4f;hb=dde568d876a5d2e1b6e554a526c98b09b145d25a;hp=67a1df4c1588c7d5ff6b3ca7f31c716e97ca66be;hpb=d83fa3d6e3604bcc596840219f3998d795630d66;p=helm.git diff --git a/matita/matita/lib/lambda-delta/substitution/drop_main.ma b/matita/matita/lib/lambda-delta/substitution/drop_main.ma index 67a1df4c1..a305b78c7 100644 --- a/matita/matita/lib/lambda-delta/substitution/drop_main.ma +++ b/matita/matita/lib/lambda-delta/substitution/drop_main.ma @@ -18,9 +18,9 @@ include "lambda-delta/substitution/drop_defs.ma". (* the main properties ******************************************************) -lemma drop_conf_ge: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L → - ∀e2,L2. ↑[0, e2] L2 ≡ L → d1 + e1 ≤ e2 → - ↑[0, e2 - e1] L2 ≡ L1. +lemma 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 [ // | #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 @@ -30,19 +30,61 @@ lemma drop_conf_ge: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L → lapply (transitive_le 1 … Hdee2) // #He2 lapply (drop_inv_drop1 … H ?) -H // -He2 #HL2 lapply (transitive_le (1+e) … Hdee2) // #Hee2 - >(plus_minus_m_m (e2-e) 1 ?) [ @drop_drop >minus_minus_comm /3/ | /2/ ] + @drop_drop_lt >minus_minus_comm /3/ (**) (* explicit constructor *) ] qed. -axiom drop_conf_lt: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L → - ∀e2,K2,I,V2. ↑[0, e2] K2. 𝕓{I} V2 ≡ L → +lemma drop_conf_lt: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 → + ∀e2,K2,I,V2. ↓[0, e2] L ≡ K2. 𝕓{I} V2 → e2 < d1 → let d ≝ d1 - e2 - 1 in - ∃∃K1,V1. ↑[0, e2] K1. 𝕓{I} V1 ≡ L1 & - ↑[d, e1] K1 ≡ K2 & ↑[d,e1] V1 ≡ V2. + ∃∃K1,V1. ↓[0, e2] L1 ≡ K1. 𝕓{I} V1 & + ↓[d, e1] K2 ≡ K1 & ↑[d, e1] V1 ≡ V2. +#d1 #e1 #L #L1 #H elim H -H d1 e1 L L1 +[ #L0 #e2 #K2 #I #V2 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d + elim (drop_inv_O1 … H) -H * + [ -IHL12 He2d #H1 #H2 destruct -e2 K2 J V /2 width=5/ + | -HL12 -HV12 #He #HLK + elim (IHL12 … HLK ?) -IHL12 HLK [ le_plus_minus // @IHL12 /2/ (**) (* explicit constructor *) +] +qed. -axiom drop_trans_ge: ∀d1,e1,L1,L. ↑[d1, e1] L ≡ L1 → - ∀e2,L2. ↑[0, e2] L2 ≡ L → d1 ≤ e2 → ↑[0, e1 + e2] L2 ≡ L1. +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.