29752
|
1 |
(* Title: FOL/ex/Nat_Class.thy
|
1246
|
2 |
Author: Markus Wenzel, TU Muenchen
|
|
3 |
*)
|
|
4 |
|
29752
|
5 |
theory Nat_Class
|
17245
|
6 |
imports FOL
|
|
7 |
begin
|
|
8 |
|
|
9 |
text {*
|
29753
|
10 |
This is an abstract version of theory @{text Nat}. Instead of
|
17245
|
11 |
axiomatizing a single type @{text nat} we define the class of all
|
|
12 |
these types (up to isomorphism).
|
|
13 |
|
|
14 |
Note: The @{text rec} operator had to be made \emph{monomorphic},
|
|
15 |
because class axioms may not contain more than one type variable.
|
|
16 |
*}
|
1246
|
17 |
|
29751
|
18 |
class nat =
|
|
19 |
fixes Zero :: 'a ("0")
|
29753
|
20 |
and Suc :: "'a \<Rightarrow> 'a"
|
29751
|
21 |
and rec :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a"
|
|
22 |
assumes induct: "P(0) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> P(Suc(x))) \<Longrightarrow> P(n)"
|
|
23 |
and Suc_inject: "Suc(m) = Suc(n) \<Longrightarrow> m = n"
|
|
24 |
and Suc_neq_Zero: "Suc(m) = 0 \<Longrightarrow> R"
|
|
25 |
and rec_Zero: "rec(0, a, f) = a"
|
|
26 |
and rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))"
|
|
27 |
begin
|
1246
|
28 |
|
19819
|
29 |
definition
|
29751
|
30 |
add :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "+" 60) where
|
|
31 |
"m + n = rec(m, n, \<lambda>x y. Suc(y))"
|
19819
|
32 |
|
29753
|
33 |
lemma Suc_n_not_n: "Suc(k) \<noteq> (k::'a)"
|
29751
|
34 |
apply (rule_tac n = k in induct)
|
|
35 |
apply (rule notI)
|
|
36 |
apply (erule Suc_neq_Zero)
|
|
37 |
apply (rule notI)
|
|
38 |
apply (erule notE)
|
|
39 |
apply (erule Suc_inject)
|
|
40 |
done
|
19819
|
41 |
|
29751
|
42 |
lemma "(k + m) + n = k + (m + n)"
|
|
43 |
apply (rule induct)
|
|
44 |
back
|
|
45 |
back
|
|
46 |
back
|
|
47 |
back
|
|
48 |
back
|
|
49 |
oops
|
19819
|
50 |
|
29751
|
51 |
lemma add_Zero [simp]: "0 + n = n"
|
|
52 |
apply (unfold add_def)
|
|
53 |
apply (rule rec_Zero)
|
|
54 |
done
|
1246
|
55 |
|
29751
|
56 |
lemma add_Suc [simp]: "Suc(m) + n = Suc(m + n)"
|
|
57 |
apply (unfold add_def)
|
|
58 |
apply (rule rec_Suc)
|
|
59 |
done
|
19819
|
60 |
|
29751
|
61 |
lemma add_assoc: "(k + m) + n = k + (m + n)"
|
|
62 |
apply (rule_tac n = k in induct)
|
|
63 |
apply simp
|
|
64 |
apply simp
|
|
65 |
done
|
19819
|
66 |
|
29751
|
67 |
lemma add_Zero_right: "m + 0 = m"
|
|
68 |
apply (rule_tac n = m in induct)
|
|
69 |
apply simp
|
|
70 |
apply simp
|
|
71 |
done
|
19819
|
72 |
|
29751
|
73 |
lemma add_Suc_right: "m + Suc(n) = Suc(m + n)"
|
|
74 |
apply (rule_tac n = m in induct)
|
|
75 |
apply simp_all
|
|
76 |
done
|
19819
|
77 |
|
|
78 |
lemma
|
29751
|
79 |
assumes prem: "\<And>n. f(Suc(n)) = Suc(f(n))"
|
|
80 |
shows "f(i + j) = i + f(j)"
|
|
81 |
apply (rule_tac n = i in induct)
|
|
82 |
apply simp
|
|
83 |
apply (simp add: prem)
|
|
84 |
done
|
1246
|
85 |
|
|
86 |
end
|
29751
|
87 |
|
|
88 |
end
|