8906
|
1 |
|
9146
|
2 |
header {* Defining natural numbers in FOL \label{sec:ex-natclass} *}
|
8906
|
3 |
|
9146
|
4 |
theory NatClass = FOL:
|
8890
|
5 |
|
8906
|
6 |
text {*
|
|
7 |
\medskip\noindent Axiomatic type classes abstract over exactly one
|
|
8 |
type argument. Thus, any \emph{axiomatic} theory extension where each
|
|
9 |
axiom refers to at most one type variable, may be trivially turned
|
|
10 |
into a \emph{definitional} one.
|
|
11 |
|
|
12 |
We illustrate this with the natural numbers in
|
|
13 |
Isabelle/FOL.\footnote{See also
|
|
14 |
\url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}}
|
9146
|
15 |
*}
|
8906
|
16 |
|
8890
|
17 |
consts
|
|
18 |
zero :: 'a ("0")
|
|
19 |
Suc :: "'a \\<Rightarrow> 'a"
|
9146
|
20 |
rec :: "'a \\<Rightarrow> 'a \\<Rightarrow> ('a \\<Rightarrow> 'a \\<Rightarrow> 'a) \\<Rightarrow> 'a"
|
8890
|
21 |
|
|
22 |
axclass
|
|
23 |
nat < "term"
|
|
24 |
induct: "P(0) \\<Longrightarrow> (\\<And>x. P(x) \\<Longrightarrow> P(Suc(x))) \\<Longrightarrow> P(n)"
|
|
25 |
Suc_inject: "Suc(m) = Suc(n) \\<Longrightarrow> m = n"
|
|
26 |
Suc_neq_0: "Suc(m) = 0 \\<Longrightarrow> R"
|
|
27 |
rec_0: "rec(0, a, f) = a"
|
9146
|
28 |
rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))"
|
8890
|
29 |
|
|
30 |
constdefs
|
|
31 |
add :: "'a::nat \\<Rightarrow> 'a \\<Rightarrow> 'a" (infixl "+" 60)
|
9146
|
32 |
"m + n \\<equiv> rec(m, n, \\<lambda>x y. Suc(y))"
|
8890
|
33 |
|
8906
|
34 |
text {*
|
|
35 |
This is an abstract version of the plain $Nat$ theory in
|
|
36 |
FOL.\footnote{See
|
8907
|
37 |
\url{http://isabelle.in.tum.de/library/FOL/ex/Nat.html}} Basically,
|
|
38 |
we have just replaced all occurrences of type $nat$ by $\alpha$ and
|
|
39 |
used the natural number axioms to define class $nat$. There is only
|
|
40 |
a minor snag, that the original recursion operator $rec$ had to be
|
|
41 |
made monomorphic.
|
8906
|
42 |
|
8907
|
43 |
Thus class $nat$ contains exactly those types $\tau$ that are
|
|
44 |
isomorphic to ``the'' natural numbers (with signature $0$, $Suc$,
|
|
45 |
$rec$).
|
8906
|
46 |
|
|
47 |
\medskip What we have done here can be also viewed as \emph{type
|
|
48 |
specification}. Of course, it still remains open if there is some
|
|
49 |
type at all that meets the class axioms. Now a very nice property of
|
8907
|
50 |
axiomatic type classes is that abstract reasoning is always possible
|
8906
|
51 |
--- independent of satisfiability. The meta-logic won't break, even
|
8907
|
52 |
if some classes (or general sorts) turns out to be empty later ---
|
|
53 |
``inconsistent'' class definitions may be useless, but do not cause
|
|
54 |
any harm.
|
8906
|
55 |
|
|
56 |
Theorems of the abstract natural numbers may be derived in the same
|
|
57 |
way as for the concrete version. The original proof scripts may be
|
8907
|
58 |
re-used with some trivial changes only (mostly adding some type
|
8906
|
59 |
constraints).
|
9146
|
60 |
*}
|
8906
|
61 |
|
9146
|
62 |
end |