X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=9fedf60c0c180fd43f5e2ce90139254ba59dbfca;hb=f6b452b9c9be141740d4058dfbcf81a4b75fd00b;hp=d70ea1bfa74cc134edde18bd68c9ffdcc6a9dead;hpb=ff1cd6f29b3aaef01e4674544d399f44949c5738;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 d70ea1bfa..9fedf60c0 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -35,6 +35,9 @@ normalize // qed. lemma pred_S: ∀m. pred (S m) = m. // qed. +lemma plus_S1: ∀x,y. ⫯(x+y) = (⫯x) + y. +// qed. + lemma max_O1: ∀n. n = (0 ∨ n). // qed. @@ -48,6 +51,9 @@ qed. (* Equations ****************************************************************) +lemma plus_SO: ∀n. n + 1 = ⫯n. +// qed. + lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m. // qed-. @@ -152,9 +158,15 @@ qed. (* Inversion & forward lemmas ***********************************************) +lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0. +/2 width=1 by plus_le_0/ qed-. + lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0. // qed-. +lemma discr_plus_x_xy: ∀x,y. x = x + y → y = 0. +/2 width=2 by le_plus_minus_comm/ qed-. + lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y. /2 width=1 by monotonic_pred/ qed-. @@ -169,6 +181,10 @@ qed-. lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥. /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-. +lemma succ_inv_refl_sn: ∀x. ⫯x = x → ⊥. +#x #H @(lt_le_false x (⫯x)) // +qed-. + lemma lt_inv_O1: ∀n. 0 < n → ∃m. ⫯m = n. * /2 width=2 by ex_intro/ #H cases (lt_le_false … H) -H //