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