X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fnat.ma;h=7bd51c318fb99f02d2e158395fbd6f56ec7c0a9d;hb=520d4370a540a98f5e5e1d85acfef0c982cc1e04;hp=5b430d7c651fd8521feedf30f8cf8531f0e3677e;hpb=07f64a04ac6061c853d2e60237a7173968c6d759;p=helm.git diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index 5b430d7c6..7bd51c318 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -439,6 +439,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) @@ -505,6 +509,19 @@ 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-. (* More negated equalities **************************************************) @@ -669,6 +686,10 @@ lapply (minus_le x y)