X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=inline;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FGround_2%2Farith.ma;h=39d28c959182d17c31fa4a851751928c3fc12b8d;hb=16bbb2d6b16d5647d944f18f0fd6d4dd3df431fe;hp=37fac204b70452b7cb98881b30f63b303a844633;hpb=6ed2537d49a307259db46481469fa44b2cfc56e6;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 37fac204b..39d28c959 100644 --- a/matita/matita/contribs/lambda_delta/Ground_2/arith.ma +++ b/matita/matita/contribs/lambda_delta/Ground_2/arith.ma @@ -17,10 +17,7 @@ include "Ground_2/star.ma". (* ARITHMETICAL PROPERTIES **************************************************) -(* equations ****************************************************************) - -lemma plus_plus_comm_23: ∀x,y,z. x + y + z = x + z + y. -// qed. +(* Equations ****************************************************************) lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p). /2 by plus_minus/ qed. @@ -28,13 +25,6 @@ lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p). 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 #c #a #H >(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_minus_comm >minus_le_minus_minus_comm // qed. @@ -57,25 +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 #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_inv_plus_l: ∀x,y,z. x + y < z → x < z ∧ y < z - x. -/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,13 +66,11 @@ lemma false_lt_to_le: ∀x,y. (x < y → False) → y ≤ x. #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