X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fnat.ma;h=67b7de50b7dfb5711e9679b502720862e49c6d03;hb=2199f327081f49b21bdcd23d702b5e07ea4f58ce;hp=649108f47c27110a30a770489ce0d9f82bd221d3;hpb=7163f2ff5eb4960162d2fcf135f031ae2a9f8b56;p=helm.git diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index 649108f47..67b7de50b 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -135,9 +135,6 @@ theorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). theorem injective_plus_l: ∀n:nat.injective nat nat (λm.m+n). /2/ qed. -theorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. -/2/ qed. - theorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m. // qed. @@ -192,6 +189,9 @@ lemma plus_plus_comm_23: ∀x,y,z. x + y + z = x + z + y. (* Negated equalities *******************************************************) +theorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. +/2/ qed. + theorem not_eq_O_S : ∀n:nat. 0 ≠ S n. #n @nmk #eqOS (change with (not_zero O)) >eqOS // qed. @@ -299,6 +299,16 @@ theorem le_plus_to_minus_r: ∀a,b,c. a + b ≤ c → a ≤ c -b. #a #b #c #H @(le_plus_to_le_r … b) /2/ qed. +lemma lt_to_le: ∀x,y. x < y → x ≤ y. +/2 width=2/ qed. + +lemma inv_eq_minus_O: ∀x,y. x - y = 0 → x ≤ y. +// qed-. + +lemma le_x_times_x: ∀x. x ≤ x * x. +#x elim x -x // +qed. + (* lt *) theorem transitive_lt: transitive nat lt. @@ -485,6 +495,13 @@ cut (∀q:nat. q ≤ n → P q) /2/ ] qed. +lemma f_ind: ∀A. ∀f:A→ℕ. ∀P:predicate A. + (∀n. (∀a. f a < n → P a) → ∀a. f a = n → P a) → ∀a. P a. +#A #f #P #H #a +cut (∀n,a. f a = n → P a) /2 width=3/ -a +#n @(nat_elim1 … n) -n #n /3 width=3/ (**) (* auto very slow (274s) without #n *) +qed-. + (* More negated equalities **************************************************) theorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m. @@ -600,6 +617,9 @@ lemma le_plus_compatible: ∀x1,x2,y1,y2. x1 ≤ y1 → x2 ≤ y2 → x1 + x2 #x1 #y1 #x2 #y2 #H1 #H2 /2/ @le_plus // /2/ /3 by le_minus_to_plus, monotonic_le_plus_r, transitive_le/ qed. *) +lemma minus_le: ∀x,y. x - y ≤ x. +/2 width=1/ qed. + (* lt *) theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n