9767
|
1 |
%
|
|
2 |
\begin{isabellebody}%
|
9921
|
3 |
\def\isabellecontext{Product}%
|
8903
|
4 |
%
|
|
5 |
\isamarkupheader{Syntactic classes}
|
9672
|
6 |
\isacommand{theory}\ Product\ {\isacharequal}\ Main{\isacharcolon}%
|
8903
|
7 |
\begin{isamarkuptext}%
|
|
8 |
\medskip\noindent There is still a feature of Isabelle's type system
|
8907
|
9 |
left that we have not yet discussed. When declaring polymorphic
|
|
10 |
constants $c :: \sigma$, the type variables occurring in $\sigma$ may
|
|
11 |
be constrained by type classes (or even general sorts) in an
|
|
12 |
arbitrary way. Note that by default, in Isabelle/HOL the declaration
|
|
13 |
$\TIMES :: \alpha \To \alpha \To \alpha$ is actually an abbreviation
|
|
14 |
for $\TIMES :: (\alpha::term) \To \alpha \To \alpha$. Since class
|
|
15 |
$term$ is the universal class of HOL, this is not really a constraint
|
|
16 |
at all.
|
8903
|
17 |
|
|
18 |
The $product$ class below provides a less degenerate example of
|
|
19 |
syntactic type classes.%
|
|
20 |
\end{isamarkuptext}%
|
8890
|
21 |
\isacommand{axclass}\isanewline
|
9665
|
22 |
\ \ product\ {\isacharless}\ {\isachardoublequote}term{\isachardoublequote}\isanewline
|
8890
|
23 |
\isacommand{consts}\isanewline
|
9672
|
24 |
\ \ product\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isacharcolon}{\isacharcolon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymOtimes}{\isachardoublequote}\ \isadigit{7}\isadigit{0}{\isacharparenright}%
|
8903
|
25 |
\begin{isamarkuptext}%
|
|
26 |
Here class $product$ is defined as subclass of $term$ without any
|
|
27 |
additional axioms. This effects in logical equivalence of $product$
|
|
28 |
and $term$, as is reflected by the trivial introduction rule
|
|
29 |
generated for this definition.
|
|
30 |
|
|
31 |
\medskip So what is the difference of declaring $\TIMES :: (\alpha ::
|
|
32 |
product) \To \alpha \To \alpha$ vs.\ declaring $\TIMES :: (\alpha ::
|
|
33 |
term) \To \alpha \To \alpha$ anyway? In this particular case where
|
|
34 |
$product \equiv term$, it should be obvious that both declarations
|
|
35 |
are the same from the logic's point of view. It even makes the most
|
|
36 |
sense to remove sort constraints from constant declarations, as far
|
|
37 |
as the purely logical meaning is concerned \cite{Wenzel:1997:TPHOL}.
|
|
38 |
|
|
39 |
On the other hand there are syntactic differences, of course.
|
|
40 |
Constants $\TIMES^\tau$ are rejected by the type-checker, unless the
|
|
41 |
arity $\tau :: product$ is part of the type signature. In our
|
|
42 |
example, this arity may be always added when required by means of an
|
|
43 |
$\isarkeyword{instance}$ with the trivial proof $\BY{intro_classes}$.
|
|
44 |
|
|
45 |
\medskip Thus, we may observe the following discipline of using
|
|
46 |
syntactic classes. Overloaded polymorphic constants have their type
|
|
47 |
arguments restricted to an associated (logically trivial) class $c$.
|
|
48 |
Only immediately before \emph{specifying} these constants on a
|
|
49 |
certain type $\tau$ do we instantiate $\tau :: c$.
|
|
50 |
|
|
51 |
This is done for class $product$ and type $bool$ as follows.%
|
|
52 |
\end{isamarkuptext}%
|
9672
|
53 |
\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ product\isanewline
|
9665
|
54 |
\ \ \isacommand{by}\ intro{\isacharunderscore}classes\isanewline
|
|
55 |
\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline
|
9672
|
56 |
\ \ product{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x\ {\isasymOtimes}\ y\ {\isasymequiv}\ x\ {\isasymand}\ y{\isachardoublequote}%
|
8903
|
57 |
\begin{isamarkuptext}%
|
|
58 |
The definition $prod_bool_def$ becomes syntactically well-formed only
|
|
59 |
after the arity $bool :: product$ is made known to the type checker.
|
|
60 |
|
|
61 |
\medskip It is very important to see that above $\DEFS$ are not
|
|
62 |
directly connected with $\isarkeyword{instance}$ at all! We were
|
|
63 |
just following our convention to specify $\TIMES$ on $bool$ after
|
|
64 |
having instantiated $bool :: product$. Isabelle does not require
|
|
65 |
these definitions, which is in contrast to programming languages like
|
|
66 |
Haskell \cite{haskell-report}.
|
|
67 |
|
|
68 |
\medskip While Isabelle type classes and those of Haskell are almost
|
|
69 |
the same as far as type-checking and type inference are concerned,
|
8907
|
70 |
there are important semantic differences. Haskell classes require
|
|
71 |
their instances to \emph{provide operations} of certain \emph{names}.
|
8903
|
72 |
Therefore, its \texttt{instance} has a \texttt{where} part that tells
|
|
73 |
the system what these ``member functions'' should be.
|
|
74 |
|
8907
|
75 |
This style of \texttt{instance} won't make much sense in Isabelle's
|
|
76 |
meta-logic, because there is no internal notion of ``providing
|
|
77 |
operations'' or even ``names of functions''.%
|
8903
|
78 |
\end{isamarkuptext}%
|
9767
|
79 |
\isacommand{end}\end{isabellebody}%
|
9145
|
80 |
%%% Local Variables:
|
|
81 |
%%% mode: latex
|
|
82 |
%%% TeX-master: "root"
|
|
83 |
%%% End:
|