X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=3eec591c692269b16eae8bff92bb6efbfb7bc4cf;hb=3325b784763ae9e6bac4307463071bb38e5641c9;hp=92b07e1f2b9525a4bbd743f498ede54a2860e057;hpb=a0d25627e80a3a2fe68da954b68f6d541c6dbc34;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma index 92b07e1f2..3eec591c6 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -36,7 +36,7 @@ lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b. 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 width=1/ +#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1 by arith_b1/ qed. lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. @@ -44,7 +44,7 @@ lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b → a1 - c1 + a2 - (b - c1) = a1 + a2 - b. -#a1 #a2 #b #c1 #H1 #H2 >plus_minus // /2 width=1/ +#a1 #a2 #b #c1 #H1 #H2 >plus_minus /2 width=1 by arith_b2/ qed. lemma arith_i: ∀x,y,z. y < x → x+z-y-1 = x-y-1+z. @@ -91,21 +91,21 @@ axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). 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/ +#m #n elim (lt_or_ge m n) /2 width=1 by or3_intro0/ +#H elim H -m /2 width=1 by or3_intro1/ +#m #Hm * /3 width=1 by not_le_to_lt, le_S_S, or3_intro2/ qed-. lemma lt_refl_false: ∀n. n < n → ⊥. -#n #H elim (lt_to_not_eq … H) -H /2 width=1/ +#n #H elim (lt_to_not_eq … H) -H /2 width=1 by/ qed-. lemma lt_zero_false: ∀n. n < 0 → ⊥. -#n #H elim (lt_to_not_le … H) -H /2 width=1/ +#n #H elim (lt_to_not_le … H) -H /2 width=1 by/ qed-. lemma false_lt_to_le: ∀x,y. (x < y → ⊥) → y ≤ x. -#x #y #H elim (decidable_lt x y) /2 width=1/ +#x #y #H elim (decidable_lt x y) /2 width=1 by not_lt_to_le/ #Hxy elim (H Hxy) qed-. @@ -113,12 +113,13 @@ lemma pred_inv_refl: ∀m. pred m = m → m = 0. * // normalize #m #H elim (lt_refl_false m) // qed-. +lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥. +#x #y #H lapply (le_n_O_to_eq … H) -H