author | wenzelm |
Mon, 06 Nov 2017 17:21:32 +0100 | |
changeset 67016 | 57d58c3cf16b |
parent 60770 | 240563fbf41d |
child 69587 | 53982d5ec0bb |
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 |
Author: Markus Wenzel, TU Munich |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
3 |
*) |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
4 |
|
60770 | 5 |
section \<open>Natural numbers\<close> |
12371 | 6 |
|
31974 | 7 |
theory Natural_Numbers |
8 |
imports FOL |
|
9 |
begin |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
10 |
|
60770 | 11 |
text \<open> |
12371 | 12 |
Theory of the natural numbers: Peano's axioms, primitive recursion. |
13 |
(Modernized version of Larry Paulson's theory "Nat".) \medskip |
|
60770 | 14 |
\<close> |
12371 | 15 |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
16 |
typedecl nat |
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
44605
diff
changeset
|
17 |
instance nat :: "term" .. |
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
18 |
|
26720 | 19 |
axiomatization |
20 |
Zero :: nat ("0") and |
|
21 |
Suc :: "nat => nat" and |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
22 |
rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a" |
26720 | 23 |
where |
11789 | 24 |
induct [case_names 0 Suc, induct type: nat]: |
26720 | 25 |
"P(0) ==> (!!x. P(x) ==> P(Suc(x))) ==> P(n)" and |
26 |
Suc_inject: "Suc(m) = Suc(n) ==> m = n" and |
|
27 |
Suc_neq_0: "Suc(m) = 0 ==> R" and |
|
28 |
rec_0: "rec(0, a, f) = a" and |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
29 |
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
|
30 |
|
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
31 |
lemma Suc_n_not_n: "Suc(k) \<noteq> k" |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
32 |
proof (induct k) |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
33 |
show "Suc(0) \<noteq> 0" |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
34 |
proof |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
35 |
assume "Suc(0) = 0" |
25989 | 36 |
then show False by (rule Suc_neq_0) |
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
37 |
qed |
26720 | 38 |
next |
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
39 |
fix n assume hyp: "Suc(n) \<noteq> n" |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
40 |
show "Suc(Suc(n)) \<noteq> Suc(n)" |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
41 |
proof |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
42 |
assume "Suc(Suc(n)) = Suc(n)" |
25989 | 43 |
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
|
44 |
with hyp show False by contradiction |
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 |
qed |
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
47 |
|
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
48 |
|
44605 | 49 |
definition add :: "nat => nat => nat" (infixl "+" 60) |
50 |
where "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 |