| author | haftmann | 
| Tue, 07 Oct 2008 16:07:33 +0200 | |
| changeset 28522 | eacb54d9e78d | 
| parent 26720 | 8d1925ad0dac | 
| child 31974 | e81979a703a4 | 
| permissions | -rw-r--r-- | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 1 | (* Title: FOL/ex/Natural_Numbers.thy | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Munich | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 4 | *) | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 5 | |
| 12371 | 6 | header {* Natural numbers *}
 | 
| 7 | ||
| 16417 | 8 | theory Natural_Numbers imports FOL begin | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 9 | |
| 12371 | 10 | text {*
 | 
| 11 | Theory of the natural numbers: Peano's axioms, primitive recursion. | |
| 12 | (Modernized version of Larry Paulson's theory "Nat".) \medskip | |
| 13 | *} | |
| 14 | ||
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 15 | typedecl nat | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 16 | arities nat :: "term" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 17 | |
| 26720 | 18 | axiomatization | 
| 19 |   Zero :: nat    ("0") and
 | |
| 20 | Suc :: "nat => nat" and | |
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 21 | rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a" | 
| 26720 | 22 | where | 
| 11789 | 23 | induct [case_names 0 Suc, induct type: nat]: | 
| 26720 | 24 | "P(0) ==> (!!x. P(x) ==> P(Suc(x))) ==> P(n)" and | 
| 25 | Suc_inject: "Suc(m) = Suc(n) ==> m = n" and | |
| 26 | Suc_neq_0: "Suc(m) = 0 ==> R" and | |
| 27 | rec_0: "rec(0, a, f) = a" and | |
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 28 | rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 29 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 30 | lemma Suc_n_not_n: "Suc(k) \<noteq> k" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 31 | proof (induct k) | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 32 | show "Suc(0) \<noteq> 0" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 33 | proof | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 34 | assume "Suc(0) = 0" | 
| 25989 | 35 | then show False by (rule Suc_neq_0) | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 36 | qed | 
| 26720 | 37 | next | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 38 | fix n assume hyp: "Suc(n) \<noteq> n" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 39 | show "Suc(Suc(n)) \<noteq> Suc(n)" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 40 | proof | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 41 | assume "Suc(Suc(n)) = Suc(n)" | 
| 25989 | 42 | then have "Suc(n) = n" by (rule Suc_inject) | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 43 | with hyp show False by contradiction | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 44 | qed | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 45 | qed | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 46 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 47 | |
| 26720 | 48 | definition | 
| 49 | add :: "[nat, nat] => nat" (infixl "+" 60) where | |
| 50 | "m + n = rec(m, n, \<lambda>x y. Suc(y))" | |
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 51 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 52 | lemma add_0 [simp]: "0 + n = n" | 
| 26720 | 53 | unfolding add_def by (rule rec_0) | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 54 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 55 | lemma add_Suc [simp]: "Suc(m) + n = Suc(m + n)" | 
| 26720 | 56 | unfolding add_def by (rule rec_Suc) | 
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 57 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 58 | lemma add_assoc: "(k + m) + n = k + (m + n)" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 59 | by (induct k) simp_all | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 60 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 61 | lemma add_0_right: "m + 0 = m" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 62 | by (induct m) simp_all | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 63 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 64 | lemma add_Suc_right: "m + Suc(n) = Suc(m + n)" | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 65 | by (induct m) simp_all | 
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 66 | |
| 26720 | 67 | lemma | 
| 68 | assumes "!!n. f(Suc(n)) = Suc(f(n))" | |
| 69 | shows "f(i + j) = i + f(j)" | |
| 70 | using assms by (induct i) simp_all | |
| 11673 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 71 | |
| 
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
 wenzelm parents: diff
changeset | 72 | end |