X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FGround_2%2Farith.ma;h=37fac204b70452b7cb98881b30f63b303a844633;hb=6ed2537d49a307259db46481469fa44b2cfc56e6;hp=a18844e9ba805359bd53b7d772c597e099167e02;hpb=035e3f52f8da3cb3cdb493aa20568ad673cc2cf5;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/Ground_2/arith.ma b/matita/matita/contribs/lambda_delta/Ground_2/arith.ma index a18844e9b..37fac204b 100644 --- a/matita/matita/contribs/lambda_delta/Ground_2/arith.ma +++ b/matita/matita/contribs/lambda_delta/Ground_2/arith.ma @@ -13,179 +13,89 @@ (**************************************************************************) include "arithmetics/nat.ma". -include "Ground_2/xoa_props.ma". +include "Ground_2/star.ma". (* ARITHMETICAL PROPERTIES **************************************************) -lemma plus_S_eq_O_false: ∀n,m. n + S m = 0 → False. -#n #m (commutative_plus p) (le_O_to_eq_O … H) -H // -| #b #IHb #c elim c -c // - #c #_ #a #Hcb - lapply (le_S_S_to_le … Hcb) -Hcb #Hcb - (plus_minus_m_m b c) in ⊢ (? ? ?%); // +qed. lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b. -#a #b #c1 #H >minus_plus @eq_f2 /2/ +#a #b #c1 #H >minus_minus_comm >minus_le_minus_minus_comm // qed. lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b. -#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2/ +#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1/ qed. -lemma arith_c1: ∀a,b,c1. a + c1 - (b + c1) = a - b. -// qed. - lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. -#x #a #b #c1 >plus_plus_comm_23 // -qed. - -lemma arith_d1: ∀a,b,c1. c1 ≤ b → a + c1 + (b - c1) = a + b. -/2/ qed. - -lemma arith_e2: ∀a,c1,c2. a ≤ c1 → c1 + c2 - (c1 - a + c2) = a. -/3/ qed. - -lemma arith_f1: ∀a,b,c1. a + b ≤ c1 → c1 - (c1 - a - b) = a + b. -#a #b #c1 #H >minus_plus plus_minus // /2 width=1/ qed. -lemma arith_i2: ∀a,c1,c2. c1 + c2 ≤ a → c1 + c2 + (a - c1 - c2) = a. -/2/ qed. +(* inversion & forward lemmas ***********************************************) -lemma arith_z1: ∀a,b,c1. a + c1 - b - c1 = a - b. -// qed. +axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). -(* unstable *****************************************************************) +axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). -lemma arith1: ∀n,h,m,p. n + h + m ≤ p + h → n + m ≤ p. -/2/ qed. +lemma lt_or_ge: ∀m,n. m < n ∨ n ≤ m. +#m #n elim (decidable_lt m n) /2 width=1/ /3 width=1/ +qed-. -lemma arith2: ∀j,i,e,d. d + e ≤ i → d ≤ i - e + j. -#j #i #e #d #H lapply (plus_le_minus … H) -H /2/ -qed. +lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m. +#m #n elim (lt_or_ge m n) /2 width=1/ +#H elim H -m /2 width=1/ +#m #Hm * #H /2 width=1/ /3 width=1/ +qed-. -lemma arith3: ∀a1,a2,b,c1. a1 + a2 ≤ b → a1 + c1 + a2 ≤ b + c1. -/2/ qed. +lemma le_inv_plus_plus_r: ∀x,y,z. x + z ≤ y + z → x ≤ y. +/2 by le_plus_to_le/ qed-. -lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2. -/2/ qed. +lemma le_inv_plus_l: ∀x,y,z. x + y ≤ z → x ≤ z - y ∧ y ≤ z. +/3 width=2/ qed-. -lemma arith5: ∀a,b1,b2,c1. c1 ≤ b1 → c1 ≤ a → a < b1 + b2 → a - c1 < b1 - c1 + b2. -#a #b1 #b2 #c1 #H1 #H2 #H3 -minus_plus >minus_plus @monotonic_le_plus_r @monotonic_le_minus_l // -qed. +lemma false_lt_to_le: ∀x,y. (x < y → False) → y ≤ x. +#x #y #H elim (decidable_lt x y) /2 width=1/ +#Hxy elim (H Hxy) +qed-. + +lemma le_fwd_plus_plus_ge: ∀m1,m2. m2 ≤ m1 → ∀n1,n2. m1 + n1 ≤ m2 + n2 → n1 ≤ n2. +#m1 #m2 #H #n1 #n2 >commutative_plus +#H elim (le_inv_plus_l … H) -H >commutative_plus