X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fnat.ma;h=cca4782806c2264ce0bc294a84bee55881d4a18c;hb=cd2f5b59215ea771ac137b9a7b115a05175f45d5;hp=cc6b19debaf147377cf02e9bfb39c26eea1201a5;hpb=bdc0a7a8c1de693a40f116742f5d2d3f3290c381;p=helm.git diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index cc6b19deb..cca478280 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. @@ -167,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. @@ -190,8 +187,18 @@ 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. +/2/ qed. + theorem not_eq_O_S : ∀n:nat. 0 ≠ S n. #n @nmk #eqOS (change with (not_zero O)) >eqOS // qed. @@ -242,7 +249,7 @@ theorem monotonic_le_plus_l: theorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 → m1 ≤ m2 → n1 + m1 ≤ n2 + m2. #n1 #n2 #m1 #m2 #len #lem @(transitive_le ? (n1+m2)) -/2/ qed-. +/2/ qed. theorem le_plus_n :∀n,m:nat. m ≤ n + m. /2/ qed. @@ -299,6 +306,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. @@ -429,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) @@ -454,6 +475,11 @@ lemma le_or_ge: ∀m,n. m ≤ n ∨ n ≤ m. #m #n elim (decidable_le m n) /2/ /4/ qed-. +lemma le_inv_S1: ∀x,y. S x ≤ y → ∃∃z. x ≤ z & y = S z. +#x #y #H elim H -y /2 width=3 by ex2_intro/ +#y #_ * #n #Hxn #H destruct /3 width=3 by le_S, ex2_intro/ +qed-. + (* More general conclusion **************************************************) theorem nat_ind_plus: ∀R:predicate nat. @@ -485,6 +511,44 @@ 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 +@(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-. + +fact f3_ind_aux: ∀A1,A2,A3. ∀f:A1→A2→A3→ℕ. ∀P:relation3 A1 A2 A3. + (∀n. (∀a1,a2,a3. f a1 a2 a3 < n → P a1 a2 a3) → ∀a1,a2,a3. f a1 a2 a3 = n → P a1 a2 a3) → + ∀n,a1,a2,a3. f a1 a2 a3 = n → P a1 a2 a3. +#A1 #A2 #A3 #f #P #H #n @(nat_elim1 … n) -n #n /3 width=3/ (**) (* auto slow (34s) without #n *) +qed-. + +lemma f3_ind: ∀A1,A2,A3. ∀f:A1→A2→A3→ℕ. ∀P:relation3 A1 A2 A3. + (∀n. (∀a1,a2,a3. f a1 a2 a3 < n → P a1 a2 a3) → ∀a1,a2,a3. f a1 a2 a3 = n → P a1 a2 a3) → + ∀a1,a2,a3. P a1 a2 a3. +#A1 #A2 #A3 #f #P #H #a1 #a2 #a3 +@(f3_ind_aux … H) -H [2: // | skip ] +qed-. + (* More negated equalities **************************************************) theorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m. @@ -550,7 +614,7 @@ pred n - pred m = n - m. #n #m #posn #posm @(lt_O_n_elim n posn) @(lt_O_n_elim m posm) //. qed. -theorem plus_minus_commutative: ∀x,y,z. z ≤ y → x + (y - z) = x + y - z. +theorem plus_minus_associative: ∀x,y,z. z ≤ y → x + (y - z) = x + y - z. /2 by plus_minus/ qed. (* More atomic conclusion ***************************************************) @@ -595,8 +659,13 @@ theorem increasing_to_le: ∀f:nat → nat. @(ex_intro ?? (S a)) /2/ qed. +(* thus is le_plus lemma le_plus_compatible: ∀x1,x2,y1,y2. x1 ≤ y1 → x2 ≤ y2 → x1 + x2 ≤ y1 + y2. -/3 by le_minus_to_plus, monotonic_le_plus_r, transitive_le/ qed. +#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 *) @@ -634,6 +703,18 @@ theorem monotonic_lt_minus_l: ∀p,q,n. n ≤ q → q < p → q - n < p - n. @lt_plus_to_minus_r