X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fnat.ma;h=249bd274e74622e59d69ab7d77782a86c713bc4d;hb=80178d6cf86b78bb9fc47f397f4bcfb1fd15a24f;hp=67b7de50b7dfb5711e9679b502720862e49c6d03;hpb=6b87a3e9d6dd7c3abb922750587444ac3fd08e16;p=helm.git diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index 67b7de50b..249bd274e 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -164,7 +164,7 @@ lemma times_times: ∀x,y,z. x*(y*z) = y*(x*z). // qed. theorem times_n_1 : ∀n:nat. n = n * 1. -#n // qed. +// qed. theorem minus_S_S: ∀n,m:nat.S n - S m = n -m. // qed. @@ -187,6 +187,13 @@ theorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m). lemma plus_plus_comm_23: ∀x,y,z. x + y + z = x + z + y. // qed. +lemma discr_plus_xy_minus_xz: ∀x,z,y. x + y = x - z → y = 0. +#x elim x -x // #x #IHx * normalize +[ #y #H @(IHx 0) plus_n_Sm #H lapply (IHx … H) -x -z #H destruct +] +qed-. + (* Negated equalities *******************************************************) theorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. @@ -439,6 +446,10 @@ theorem decidable_lt: ∀n,m. decidable (n < m). theorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m. #n #m #lenm (elim lenm) /3/ qed. +theorem eq_or_gt: ∀n. 0 = n ∨ 0 < n. +#n elim (le_to_or_lt_eq 0 n ?) // /2 width=1/ +qed-. + theorem increasing_to_le2: ∀f:nat → nat. increasing f → ∀m:nat. f 0 ≤ m → ∃i. f i ≤ m ∧ m < f (S i). #f #incr #m #lem (elim lem) @@ -495,13 +506,31 @@ cut (∀q:nat. q ≤ n → P q) /2/ ] qed. +fact f_ind_aux: ∀A. ∀f:A→ℕ. ∀P:predicate A. + (∀n. (∀a. f a < n → P a) → ∀a. f a = n → P a) → + ∀n,a. f a = n → P a. +#A #f #P #H #n @(nat_elim1 … n) -n #n /3 width=3/ (**) (* auto slow (34s) without #n *) +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 *) +@(f_ind_aux … H) -H [2: // | skip ] qed-. +fact f2_ind_aux: ∀A1,A2. ∀f:A1→A2→ℕ. ∀P:relation2 A1 A2. + (∀n. (∀a1,a2. f a1 a2 < n → P a1 a2) → ∀a1,a2. f a1 a2 = n → P a1 a2) → + ∀n,a1,a2. f a1 a2 = n → P a1 a2. +#A1 #A2 #f #P #H #n @(nat_elim1 … n) -n #n /3 width=3/ (**) (* auto slow (34s) without #n *) +qed-. + +lemma f2_ind: ∀A1,A2. ∀f:A1→A2→ℕ. ∀P:relation2 A1 A2. + (∀n. (∀a1,a2. f a1 a2 < n → P a1 a2) → ∀a1,a2. f a1 a2 = n → P a1 a2) → + ∀a1,a2. P a1 a2. +#A1 #A2 #f #P #H #a1 #a2 +@(f2_ind_aux … H) -H [2: // | skip ] +qed-. + (* More negated equalities **************************************************) theorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m. @@ -664,6 +693,10 @@ lapply (minus_le x y)