X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fnat.ma;h=81d6c9310247dc42c5f2d0d76e66a986ae1d45fc;hb=3fc94a73952678239bed11c605e180163f924c10;hp=d806697b8257d3c3922d478404593c0bfca1f80f;hpb=1efc4c2c7be1e4aff0ccccabf905d45795b3865f;p=helm.git diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index d806697b8..81d6c9310 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -450,8 +450,16 @@ 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 le_or_ge: ∀m,n. m ≤ n ∨ n ≤ m. +#m #n elim (decidable_le m n) /2 width=1/ /4 width=2/ +qed-. + (* More general conclusion **************************************************) +theorem nat_ind_plus: ∀R:predicate nat. + R 0 → (∀n. R n → R (n + 1)) → ∀n. R n. +/3 width=1 by nat_ind/ qed-. + theorem lt_O_n_elim: ∀n:nat. 0 < n → ∀P:nat → Prop.(∀m:nat.P (S m)) → P n. #n (elim n) // #abs @False_ind /2/ @absurd