X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FGround_2%2Farith.ma;h=418a176d5024ab7e8190c2c3220d7e45c1b3d0b2;hb=35653f628dc3a3e665fee01acc19c660c9d555e3;hp=80e23e6956a2ea78bccda3a080348a853774207b;hpb=c4ac63d7ae22b2adcc7fe7b54286a0226296eabc;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 80e23e695..418a176d5 100644 --- a/matita/matita/contribs/lambda_delta/Ground_2/arith.ma +++ b/matita/matita/contribs/lambda_delta/Ground_2/arith.ma @@ -19,30 +19,14 @@ include "Ground_2/star.ma". (* equations ****************************************************************) -lemma plus_plus_comm_23: ∀x,y,z. x + y + z = x + z + y. -// qed. - lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p). /2 by plus_minus/ qed. lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n. /2 by plus_minus/ qed. -lemma minus_minus_comm: ∀a,b,c. a - b - c = a - c - b. - /3 by monotonic_le_minus_l, le_to_le_to_eq/ qed. - -lemma minus_le_minus_minus_comm: ∀b,c,a. c ≤ b → a - (b - c) = a + c - b. -#b elim b -b -[ #c #a #H >(le_n_O_to_eq … H) -H // -| #b #IHb #c elim c -c // - #c #_ #a #Hcb - lapply (le_S_S_to_le … Hcb) -Hcb #Hcb - minus_plus /3 width=1/ +#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. @@ -63,24 +47,12 @@ axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). -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 lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m. -#m elim m -m -[ * /2 width=1/ -| #m #IHm * /2 width=1/ - #n elim (IHm n) -IHm #H - [ @or3_intro0 | @or3_intro1 destruct | @or3_intro2 ] // /2 width=1/ (**) (* /3 width=1/ is slow *) +#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 le_inv_plus_plus_r: ∀x,y,z. x + z ≤ y + z → x ≤ y. -/2 by le_plus_to_le/ qed-. - -lemma le_inv_plus_l: ∀x,y,z. x + y ≤ z → x ≤ z - y ∧ y ≤ z. -/3 width=2/ qed-. - lemma lt_refl_false: ∀n. n < n → False. #n #H elim (lt_to_not_eq … H) -H /2 width=1/ qed-. @@ -89,24 +61,12 @@ lemma lt_zero_false: ∀n. n < 0 → False. #n #H elim (lt_to_not_le … H) -H /2 width=1/ qed-. -lemma plus_S_eq_O_false: ∀n,m. n + S m = 0 → False. -#n #m