author | paulson |
Sat, 09 Sep 2023 17:18:52 +0100 | |
changeset 78655 | 0d065af1a99c |
parent 69590 | e65314985426 |
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 |
69590 | 17 |
instance nat :: \<open>term\<close> .. |
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
18 |
|
26720 | 19 |
axiomatization |
69590 | 20 |
Zero :: \<open>nat\<close> (\<open>0\<close>) and |
21 |
Suc :: \<open>nat => nat\<close> and |
|
22 |
rec :: \<open>[nat, 'a, [nat, 'a] => 'a] => 'a\<close> |
|
26720 | 23 |
where |
11789 | 24 |
induct [case_names 0 Suc, induct type: nat]: |
69590 | 25 |
\<open>P(0) ==> (!!x. P(x) ==> P(Suc(x))) ==> P(n)\<close> and |
26 |
Suc_inject: \<open>Suc(m) = Suc(n) ==> m = n\<close> and |
|
27 |
Suc_neq_0: \<open>Suc(m) = 0 ==> R\<close> and |
|
28 |
rec_0: \<open>rec(0, a, f) = a\<close> and |
|
29 |
rec_Suc: \<open>rec(Suc(m), a, f) = f(m, rec(m, a, f))\<close> |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
30 |
|
69590 | 31 |
lemma Suc_n_not_n: \<open>Suc(k) \<noteq> k\<close> |
32 |
proof (induct \<open>k\<close>) |
|
33 |
show \<open>Suc(0) \<noteq> 0\<close> |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
34 |
proof |
69590 | 35 |
assume \<open>Suc(0) = 0\<close> |
36 |
then show \<open>False\<close> 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 |
69590 | 39 |
fix n assume hyp: \<open>Suc(n) \<noteq> n\<close> |
40 |
show \<open>Suc(Suc(n)) \<noteq> Suc(n)\<close> |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
41 |
proof |
69590 | 42 |
assume \<open>Suc(Suc(n)) = Suc(n)\<close> |
43 |
then have \<open>Suc(n) = n\<close> by (rule Suc_inject) |
|
44 |
with hyp show \<open>False\<close> by contradiction |
|
11673
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 |
|
69590 | 49 |
definition add :: \<open>nat => nat => nat\<close> (infixl \<open>+\<close> 60) |
50 |
where \<open>m + n = rec(m, n, \<lambda>x y. Suc(y))\<close> |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
51 |
|
69590 | 52 |
lemma add_0 [simp]: \<open>0 + n = n\<close> |
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 |
|
69590 | 55 |
lemma add_Suc [simp]: \<open>Suc(m) + n = Suc(m + n)\<close> |
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 |
|
69590 | 58 |
lemma add_assoc: \<open>(k + m) + n = k + (m + n)\<close> |
59 |
by (induct \<open>k\<close>) simp_all |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
60 |
|
69590 | 61 |
lemma add_0_right: \<open>m + 0 = m\<close> |
62 |
by (induct \<open>m\<close>) simp_all |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
63 |
|
69590 | 64 |
lemma add_Suc_right: \<open>m + Suc(n) = Suc(m + n)\<close> |
65 |
by (induct \<open>m\<close>) simp_all |
|
11673
79e5536af6c4
Theory of the natural numbers: Peano's axioms, primitive recursion.
wenzelm
parents:
diff
changeset
|
66 |
|
26720 | 67 |
lemma |
69590 | 68 |
assumes \<open>!!n. f(Suc(n)) = Suc(f(n))\<close> |
69 |
shows \<open>f(i + j) = i + f(j)\<close> |
|
70 |
using assms by (induct \<open>i\<close>) 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 |