| 20967 |      1 | %
 | 
|  |      2 | \begin{isabellebody}%
 | 
|  |      3 | \def\isabellecontext{Classes}%
 | 
|  |      4 | %
 | 
|  |      5 | \isadelimtheory
 | 
|  |      6 | \isanewline
 | 
|  |      7 | \isanewline
 | 
|  |      8 | %
 | 
|  |      9 | \endisadelimtheory
 | 
|  |     10 | %
 | 
|  |     11 | \isatagtheory
 | 
| 22479 |     12 | %
 | 
| 20967 |     13 | \endisatagtheory
 | 
|  |     14 | {\isafoldtheory}%
 | 
|  |     15 | %
 | 
|  |     16 | \isadelimtheory
 | 
|  |     17 | %
 | 
|  |     18 | \endisadelimtheory
 | 
|  |     19 | %
 | 
|  |     20 | \isadelimML
 | 
| 22317 |     21 | %
 | 
|  |     22 | \endisadelimML
 | 
|  |     23 | %
 | 
|  |     24 | \isatagML
 | 
|  |     25 | %
 | 
|  |     26 | \endisatagML
 | 
|  |     27 | {\isafoldML}%
 | 
|  |     28 | %
 | 
|  |     29 | \isadelimML
 | 
|  |     30 | %
 | 
|  |     31 | \endisadelimML
 | 
|  |     32 | %
 | 
|  |     33 | \isadelimML
 | 
| 20967 |     34 | %
 | 
|  |     35 | \endisadelimML
 | 
|  |     36 | %
 | 
|  |     37 | \isatagML
 | 
|  |     38 | %
 | 
|  |     39 | \endisatagML
 | 
|  |     40 | {\isafoldML}%
 | 
|  |     41 | %
 | 
|  |     42 | \isadelimML
 | 
|  |     43 | %
 | 
|  |     44 | \endisadelimML
 | 
|  |     45 | %
 | 
|  |     46 | \isamarkupchapter{Haskell-style classes with Isabelle/Isar%
 | 
|  |     47 | }
 | 
|  |     48 | \isamarkuptrue%
 | 
|  |     49 | %
 | 
|  |     50 | \isamarkupsection{Introduction%
 | 
|  |     51 | }
 | 
|  |     52 | \isamarkuptrue%
 | 
|  |     53 | %
 | 
|  |     54 | \begin{isamarkuptext}%
 | 
| 22317 |     55 | Type classes were introduces by Wadler and Blott \cite{wadler89how}
 | 
|  |     56 |   into the Haskell language, to allow for a reasonable implementation
 | 
|  |     57 |   of overloading\footnote{throughout this tutorial, we are referring
 | 
|  |     58 |   to classical Haskell 1.0 type classes, not considering
 | 
|  |     59 |   later additions in expressiveness}.
 | 
|  |     60 |   As a canonical example, a polymorphic equality function
 | 
|  |     61 |   \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different
 | 
| 22550 |     62 |   types for \isa{{\isasymalpha}}, which is achieved by splitting introduction
 | 
| 22317 |     63 |   of the \isa{eq} function from its overloaded definitions by means
 | 
|  |     64 |   of \isa{class} and \isa{instance} declarations:
 | 
|  |     65 | 
 | 
|  |     66 |   \medskip\noindent\hspace*{2ex}\isa{class\ eq\ where}\footnote{syntax here is a kind of isabellized Haskell} \\
 | 
|  |     67 |   \hspace*{4ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool}
 | 
|  |     68 | 
 | 
|  |     69 |   \medskip\noindent\hspace*{2ex}\isa{instance\ nat\ {\isasymColon}\ eq\ where} \\
 | 
|  |     70 |   \hspace*{4ex}\isa{eq\ {\isadigit{0}}\ {\isadigit{0}}\ {\isacharequal}\ True} \\
 | 
|  |     71 |   \hspace*{4ex}\isa{eq\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ False} \\
 | 
|  |     72 |   \hspace*{4ex}\isa{eq\ {\isacharunderscore}\ {\isadigit{0}}\ {\isacharequal}\ False} \\
 | 
|  |     73 |   \hspace*{4ex}\isa{eq\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ eq\ n\ m}
 | 
|  |     74 | 
 | 
|  |     75 |   \medskip\noindent\hspace*{2ex}\isa{instance\ {\isacharparenleft}{\isasymalpha}{\isasymColon}eq{\isacharcomma}\ {\isasymbeta}{\isasymColon}eq{\isacharparenright}\ pair\ {\isasymColon}\ eq\ where} \\
 | 
| 22550 |     76 |   \hspace*{4ex}\isa{eq\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ eq\ x{\isadigit{1}}\ x{\isadigit{2}}\ {\isasymand}\ eq\ y{\isadigit{1}}\ y{\isadigit{2}}}
 | 
| 22317 |     77 | 
 | 
|  |     78 |   \medskip\noindent\hspace*{2ex}\isa{class\ ord\ extends\ eq\ where} \\
 | 
|  |     79 |   \hspace*{4ex}\isa{less{\isacharunderscore}eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\
 | 
|  |     80 |   \hspace*{4ex}\isa{less\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool}
 | 
|  |     81 | 
 | 
| 22550 |     82 |   \medskip\noindent Type variables are annotated with (finitly many) classes;
 | 
| 22317 |     83 |   these annotations are assertions that a particular polymorphic type
 | 
|  |     84 |   provides definitions for overloaded functions.
 | 
|  |     85 | 
 | 
|  |     86 |   Indeed, type classes not only allow for simple overloading
 | 
|  |     87 |   but form a generic calculus, an instance of order-sorted
 | 
|  |     88 |   algebra \cite{Nipkow-Prehofer:1993,Nipkow:1993,Wenzel:1997}.
 | 
| 20967 |     89 | 
 | 
| 22317 |     90 |   From a software enigineering point of view, type classes
 | 
|  |     91 |   correspond to interfaces in object-oriented languages like Java;
 | 
|  |     92 |   so, it is naturally desirable that type classes do not only
 | 
|  |     93 |   provide functions (class operations) but also state specifications
 | 
|  |     94 |   implementations must obey.  For example, the \isa{class\ eq}
 | 
| 22550 |     95 |   above could be given the following specification, demanding that
 | 
|  |     96 |   \isa{class\ eq} is an equivalence relation obeying reflexivity,
 | 
|  |     97 |   symmetry and transitivity:
 | 
| 22317 |     98 | 
 | 
|  |     99 |   \medskip\noindent\hspace*{2ex}\isa{class\ eq\ where} \\
 | 
|  |    100 |   \hspace*{4ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\
 | 
|  |    101 |   \hspace*{2ex}\isa{satisfying} \\
 | 
|  |    102 |   \hspace*{4ex}\isa{refl{\isacharcolon}\ eq\ x\ x} \\
 | 
|  |    103 |   \hspace*{4ex}\isa{sym{\isacharcolon}\ eq\ x\ y\ {\isasymlongleftrightarrow}\ eq\ x\ y} \\
 | 
|  |    104 |   \hspace*{4ex}\isa{trans{\isacharcolon}\ eq\ x\ y\ {\isasymand}\ eq\ y\ z\ {\isasymlongrightarrow}\ eq\ x\ z}
 | 
|  |    105 | 
 | 
| 22550 |    106 |   \medskip\noindent From a theoretic point of view, type classes are leightweight
 | 
| 22479 |    107 |   modules; Haskell type classes may be emulated by
 | 
| 22317 |    108 |   SML functors \cite{classes_modules}. 
 | 
|  |    109 |   Isabelle/Isar offers a discipline of type classes which brings
 | 
|  |    110 |   all those aspects together:
 | 
|  |    111 | 
 | 
|  |    112 |   \begin{enumerate}
 | 
|  |    113 |     \item specifying abstract operations togehter with
 | 
|  |    114 |        corresponding specifications,
 | 
|  |    115 |     \item instantating those abstract operations by a particular
 | 
|  |    116 |        type
 | 
|  |    117 |     \item in connection with a ``less ad-hoc'' approach to overloading,
 | 
|  |    118 |     \item with a direct link to the Isabelle module system (aka locales).
 | 
|  |    119 |   \end{enumerate}
 | 
|  |    120 | 
 | 
| 22550 |    121 |   \noindent Isar type classes also directly support code generation
 | 
|  |    122 |   in a Haskell like fashion.
 | 
| 22317 |    123 | 
 | 
|  |    124 |   This tutorial demonstrates common elements of structured specifications
 | 
|  |    125 |   and abstract reasoning with type classes by the algebraic hierarchy of
 | 
|  |    126 |   semigroups, monoids and groups.  Our background theory is that of
 | 
|  |    127 |   Isabelle/HOL \cite{Nipkow-et-al:2002:tutorial}, for which some
 | 
|  |    128 |   familiarity is assumed.
 | 
|  |    129 | 
 | 
|  |    130 |   Here we merely present the look-and-feel for end users.
 | 
|  |    131 |   Internally, those are mapped to more primitive Isabelle concepts.
 | 
|  |    132 |   See \cite{haftmann_wenzel2006classes} for more detail.%
 | 
| 20967 |    133 | \end{isamarkuptext}%
 | 
|  |    134 | \isamarkuptrue%
 | 
|  |    135 | %
 | 
|  |    136 | \isamarkupsection{A simple algebra example \label{sec:example}%
 | 
|  |    137 | }
 | 
|  |    138 | \isamarkuptrue%
 | 
|  |    139 | %
 | 
|  |    140 | \isamarkupsubsection{Class definition%
 | 
|  |    141 | }
 | 
|  |    142 | \isamarkuptrue%
 | 
|  |    143 | %
 | 
|  |    144 | \begin{isamarkuptext}%
 | 
|  |    145 | Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operation \isa{{\isasymcirc}} that is
 | 
|  |    146 |   assumed to be associative:%
 | 
|  |    147 | \end{isamarkuptext}%
 | 
|  |    148 | \isamarkuptrue%
 | 
|  |    149 | \ \ \ \ \isacommand{class}\isamarkupfalse%
 | 
| 22479 |    150 | \ semigroup\ {\isacharequal}\ type\ {\isacharplus}\isanewline
 | 
| 20967 |    151 | \ \ \ \ \ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}\isactrlloc {\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline
 | 
|  |    152 | \ \ \ \ \ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ y{\isacharparenright}\ \isactrlloc {\isasymotimes}\ z\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}y\ \isactrlloc {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}%
 | 
|  |    153 | \begin{isamarkuptext}%
 | 
|  |    154 | \noindent This \isa{{\isasymCLASS}} specification consists of two
 | 
|  |    155 |   parts: the \qn{operational} part names the class operation (\isa{{\isasymFIXES}}), the \qn{logical} part specifies properties on them
 | 
|  |    156 |   (\isa{{\isasymASSUMES}}).  The local \isa{{\isasymFIXES}} and \isa{{\isasymASSUMES}} are lifted to the theory toplevel, yielding the global
 | 
| 22479 |    157 |   operation \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the
 | 
|  |    158 |   global theorem \isa{semigroup{\isachardot}assoc{\isacharcolon}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.%
 | 
| 20967 |    159 | \end{isamarkuptext}%
 | 
|  |    160 | \isamarkuptrue%
 | 
|  |    161 | %
 | 
|  |    162 | \isamarkupsubsection{Class instantiation \label{sec:class_inst}%
 | 
|  |    163 | }
 | 
|  |    164 | \isamarkuptrue%
 | 
|  |    165 | %
 | 
|  |    166 | \begin{isamarkuptext}%
 | 
|  |    167 | The concrete type \isa{int} is made a \isa{semigroup}
 | 
|  |    168 |   instance by providing a suitable definition for the class operation
 | 
|  |    169 |   \isa{mult} and a proof for the specification of \isa{assoc}.%
 | 
|  |    170 | \end{isamarkuptext}%
 | 
|  |    171 | \isamarkuptrue%
 | 
|  |    172 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    173 | \ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline
 | 
| 22479 |    174 | \ \ \ \ \ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}i\ j\ {\isasymColon}\ int{\isachardot}\ i\ {\isasymotimes}\ j\ {\isasymequiv}\ i\ {\isacharplus}\ j{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    175 | %
 | 
|  |    176 | \isadelimproof
 | 
|  |    177 | \ \ \ \ %
 | 
|  |    178 | \endisadelimproof
 | 
|  |    179 | %
 | 
|  |    180 | \isatagproof
 | 
|  |    181 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    182 | \isanewline
 | 
| 22317 |    183 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
| 20967 |    184 | \ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse%
 | 
|  |    185 | \ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
|  |    186 | \ simp\isanewline
 | 
| 22317 |    187 | \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
 | 
| 20967 |    188 | \ \isacommand{show}\isamarkupfalse%
 | 
|  |    189 | \ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
 | 
|  |    190 | \ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
|  |    191 | \isanewline
 | 
|  |    192 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    193 | %
 | 
|  |    194 | \endisatagproof
 | 
|  |    195 | {\isafoldproof}%
 | 
|  |    196 | %
 | 
|  |    197 | \isadelimproof
 | 
|  |    198 | %
 | 
|  |    199 | \endisadelimproof
 | 
|  |    200 | %
 | 
|  |    201 | \begin{isamarkuptext}%
 | 
|  |    202 | \noindent From now on, the type-checker will consider \isa{int}
 | 
|  |    203 |   as a \isa{semigroup} automatically, i.e.\ any general results
 | 
|  |    204 |   are immediately available on concrete instances.
 | 
|  |    205 | 
 | 
| 22317 |    206 |   Note that the first proof step is the \isa{default} method,
 | 
|  |    207 |   which for instantiation proofs maps to the \isa{intro{\isacharunderscore}classes} method.
 | 
|  |    208 |   This boils down an instantiation judgement to the relevant primitive
 | 
|  |    209 |   proof goals and should conveniently always be the first method applied
 | 
|  |    210 |   in an instantiation proof.
 | 
|  |    211 | 
 | 
| 22550 |    212 |   \medskip Another instance of \isa{semigroup} are the natural numbers:%
 | 
| 20967 |    213 | \end{isamarkuptext}%
 | 
|  |    214 | \isamarkuptrue%
 | 
|  |    215 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    216 | \ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline
 | 
| 22317 |    217 | \ \ \ \ \ \ mult{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymequiv}\ m\ {\isacharplus}\ n{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    218 | %
 | 
|  |    219 | \isadelimproof
 | 
|  |    220 | \ \ \ \ %
 | 
|  |    221 | \endisadelimproof
 | 
|  |    222 | %
 | 
|  |    223 | \isatagproof
 | 
|  |    224 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    225 | \isanewline
 | 
|  |    226 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    227 | \ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline
 | 
|  |    228 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    229 | \ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
 | 
| 22317 |    230 | \ mult{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | 
| 20967 |    231 | \ simp\isanewline
 | 
|  |    232 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    233 | %
 | 
|  |    234 | \endisatagproof
 | 
|  |    235 | {\isafoldproof}%
 | 
|  |    236 | %
 | 
|  |    237 | \isadelimproof
 | 
|  |    238 | %
 | 
|  |    239 | \endisadelimproof
 | 
|  |    240 | %
 | 
|  |    241 | \begin{isamarkuptext}%
 | 
| 22550 |    242 | \noindent Also \isa{list}s form a semigroup with \isa{op\ {\isacharat}} as
 | 
| 20967 |    243 |   operation:%
 | 
|  |    244 | \end{isamarkuptext}%
 | 
|  |    245 | \isamarkuptrue%
 | 
|  |    246 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    247 | \ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ semigroup\isanewline
 | 
| 22317 |    248 | \ \ \ \ \ \ mult{\isacharunderscore}list{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    249 | %
 | 
|  |    250 | \isadelimproof
 | 
|  |    251 | \ \ \ \ %
 | 
|  |    252 | \endisadelimproof
 | 
|  |    253 | %
 | 
|  |    254 | \isatagproof
 | 
|  |    255 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    256 | \isanewline
 | 
|  |    257 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    258 | \ xs\ ys\ zs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ list{\isachardoublequoteclose}\isanewline
 | 
|  |    259 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    260 | \ {\isachardoublequoteopen}xs\ {\isasymotimes}\ ys\ {\isasymotimes}\ zs\ {\isacharequal}\ xs\ {\isasymotimes}\ {\isacharparenleft}ys\ {\isasymotimes}\ zs{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
|  |    261 | \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
 | 
|  |    262 | \ {\isacharminus}\isanewline
 | 
|  |    263 | \ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
 | 
| 22317 |    264 | \ mult{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse%
 | 
| 20967 |    265 | \ {\isachardoublequoteopen}{\isasymAnd}xs\ ys{\isasymColon}{\isasymalpha}\ list{\isachardot}\ xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
|  |    266 | \isanewline
 | 
|  |    267 | \ \ \ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
 | 
|  |    268 | \ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
 | 
|  |    269 | \ simp\isanewline
 | 
|  |    270 | \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    271 | \isanewline
 | 
|  |    272 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    273 | %
 | 
|  |    274 | \endisatagproof
 | 
|  |    275 | {\isafoldproof}%
 | 
|  |    276 | %
 | 
|  |    277 | \isadelimproof
 | 
|  |    278 | %
 | 
|  |    279 | \endisadelimproof
 | 
|  |    280 | %
 | 
|  |    281 | \isamarkupsubsection{Subclasses%
 | 
|  |    282 | }
 | 
|  |    283 | \isamarkuptrue%
 | 
|  |    284 | %
 | 
|  |    285 | \begin{isamarkuptext}%
 | 
| 22317 |    286 | We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral)
 | 
| 20967 |    287 |   by extending \isa{semigroup}
 | 
|  |    288 |   with one additional operation \isa{neutral} together
 | 
|  |    289 |   with its property:%
 | 
|  |    290 | \end{isamarkuptext}%
 | 
|  |    291 | \isamarkuptrue%
 | 
|  |    292 | \ \ \ \ \isacommand{class}\isamarkupfalse%
 | 
|  |    293 | \ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline
 | 
|  |    294 | \ \ \ \ \ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}\isactrlloc {\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline
 | 
|  |    295 | \ \ \ \ \ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}\isactrlloc {\isasymone}\ \isactrlloc {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}%
 | 
|  |    296 | \begin{isamarkuptext}%
 | 
|  |    297 | \noindent Again, we make some instances, by
 | 
|  |    298 |   providing suitable operation definitions and proofs for the
 | 
|  |    299 |   additional specifications.%
 | 
|  |    300 | \end{isamarkuptext}%
 | 
|  |    301 | \isamarkuptrue%
 | 
|  |    302 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    303 | \ nat\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline
 | 
| 22317 |    304 | \ \ \ \ \ \ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    305 | %
 | 
|  |    306 | \isadelimproof
 | 
|  |    307 | \ \ \ \ %
 | 
|  |    308 | \endisadelimproof
 | 
|  |    309 | %
 | 
|  |    310 | \isatagproof
 | 
|  |    311 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    312 | \isanewline
 | 
|  |    313 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    314 | \ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline
 | 
|  |    315 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    316 | \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
 | 
|  |    317 | \ neutral{\isacharunderscore}nat{\isacharunderscore}def\ mult{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | 
|  |    318 | \ simp\isanewline
 | 
|  |    319 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    320 | %
 | 
|  |    321 | \endisatagproof
 | 
|  |    322 | {\isafoldproof}%
 | 
|  |    323 | %
 | 
|  |    324 | \isadelimproof
 | 
|  |    325 | \isanewline
 | 
|  |    326 | %
 | 
|  |    327 | \endisadelimproof
 | 
|  |    328 | \isanewline
 | 
|  |    329 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    330 | \ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline
 | 
| 22317 |    331 | \ \ \ \ \ \ neutral{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    332 | %
 | 
|  |    333 | \isadelimproof
 | 
|  |    334 | \ \ \ \ %
 | 
|  |    335 | \endisadelimproof
 | 
|  |    336 | %
 | 
|  |    337 | \isatagproof
 | 
|  |    338 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    339 | \isanewline
 | 
|  |    340 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    341 | \ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline
 | 
|  |    342 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    343 | \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
 | 
|  |    344 | \ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | 
|  |    345 | \ simp\isanewline
 | 
|  |    346 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    347 | %
 | 
|  |    348 | \endisatagproof
 | 
|  |    349 | {\isafoldproof}%
 | 
|  |    350 | %
 | 
|  |    351 | \isadelimproof
 | 
|  |    352 | \isanewline
 | 
|  |    353 | %
 | 
|  |    354 | \endisadelimproof
 | 
|  |    355 | \isanewline
 | 
|  |    356 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    357 | \ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ monoidl\isanewline
 | 
| 22317 |    358 | \ \ \ \ \ \ neutral{\isacharunderscore}list{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    359 | %
 | 
|  |    360 | \isadelimproof
 | 
|  |    361 | \ \ \ \ %
 | 
|  |    362 | \endisadelimproof
 | 
|  |    363 | %
 | 
|  |    364 | \isatagproof
 | 
|  |    365 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    366 | \isanewline
 | 
|  |    367 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    368 | \ xs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ list{\isachardoublequoteclose}\isanewline
 | 
|  |    369 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    370 | \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ xs\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline
 | 
|  |    371 | \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
 | 
|  |    372 | \ {\isacharminus}\isanewline
 | 
|  |    373 | \ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
 | 
|  |    374 | \ mult{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse%
 | 
| 22479 |    375 | \ {\isachardoublequoteopen}{\isasymAnd}xs\ ys{\isasymColon}{\isasymalpha}\ list{\isachardot}\ xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
| 20967 |    376 | \isanewline
 | 
|  |    377 | \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
 | 
|  |    378 | \ \isacommand{from}\isamarkupfalse%
 | 
|  |    379 | \ mult{\isacharunderscore}list{\isacharunderscore}def\ neutral{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse%
 | 
|  |    380 | \ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isacharbrackleft}{\isacharbrackright}{\isasymColon}{\isasymalpha}\ list{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
|  |    381 | \ simp\isanewline
 | 
|  |    382 | \ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
 | 
|  |    383 | \ \isacommand{show}\isamarkupfalse%
 | 
|  |    384 | \ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
 | 
|  |    385 | \ simp\isanewline
 | 
|  |    386 | \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    387 | \isanewline
 | 
|  |    388 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    389 | %
 | 
|  |    390 | \endisatagproof
 | 
|  |    391 | {\isafoldproof}%
 | 
|  |    392 | %
 | 
|  |    393 | \isadelimproof
 | 
|  |    394 | %
 | 
|  |    395 | \endisadelimproof
 | 
|  |    396 | %
 | 
|  |    397 | \begin{isamarkuptext}%
 | 
| 22317 |    398 | \noindent Fully-fledged monoids are modelled by another subclass
 | 
|  |    399 |   which does not add new operations but tightens the specification:%
 | 
| 20967 |    400 | \end{isamarkuptext}%
 | 
|  |    401 | \isamarkuptrue%
 | 
|  |    402 | \ \ \ \ \isacommand{class}\isamarkupfalse%
 | 
|  |    403 | \ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline
 | 
| 22317 |    404 | \ \ \ \ \ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ \isactrlloc {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}%
 | 
|  |    405 | \begin{isamarkuptext}%
 | 
|  |    406 | \noindent Instantiations may also be given simultaneously for different
 | 
|  |    407 |   type constructors:%
 | 
|  |    408 | \end{isamarkuptext}%
 | 
|  |    409 | \isamarkuptrue%
 | 
| 20967 |    410 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
| 22317 |    411 | \ nat\ {\isacharcolon}{\isacharcolon}\ monoid\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoid\ \isakeyword{and}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ monoid\isanewline
 | 
| 20967 |    412 | %
 | 
|  |    413 | \isadelimproof
 | 
|  |    414 | \ \ \ \ %
 | 
|  |    415 | \endisadelimproof
 | 
|  |    416 | %
 | 
|  |    417 | \isatagproof
 | 
|  |    418 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    419 | \isanewline
 | 
|  |    420 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    421 | \ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline
 | 
|  |    422 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    423 | \ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
 | 
|  |    424 | \ neutral{\isacharunderscore}nat{\isacharunderscore}def\ mult{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | 
|  |    425 | \ simp\isanewline
 | 
| 22317 |    426 | \ \ \ \ \isacommand{next}\isamarkupfalse%
 | 
| 20967 |    427 | \isanewline
 | 
|  |    428 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    429 | \ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline
 | 
|  |    430 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    431 | \ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
 | 
|  |    432 | \ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | 
|  |    433 | \ simp\isanewline
 | 
| 22317 |    434 | \ \ \ \ \isacommand{next}\isamarkupfalse%
 | 
| 20967 |    435 | \isanewline
 | 
|  |    436 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    437 | \ xs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ list{\isachardoublequoteclose}\isanewline
 | 
|  |    438 | \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
 | 
|  |    439 | \ {\isachardoublequoteopen}xs\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline
 | 
|  |    440 | \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
 | 
|  |    441 | \ {\isacharminus}\isanewline
 | 
|  |    442 | \ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
 | 
|  |    443 | \ mult{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse%
 | 
|  |    444 | \ {\isachardoublequoteopen}{\isasymAnd}xs\ ys{\isasymColon}{\isasymalpha}\ list{\isachardot}\ xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
|  |    445 | \isanewline
 | 
|  |    446 | \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
 | 
|  |    447 | \ \isacommand{from}\isamarkupfalse%
 | 
|  |    448 | \ mult{\isacharunderscore}list{\isacharunderscore}def\ neutral{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse%
 | 
| 22479 |    449 | \ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isacharbrackleft}{\isacharbrackright}{\isasymColon}{\isasymalpha}\ list{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
| 20967 |    450 | \ simp\isanewline
 | 
|  |    451 | \ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
 | 
|  |    452 | \ \isacommand{show}\isamarkupfalse%
 | 
|  |    453 | \ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
 | 
|  |    454 | \ simp\isanewline
 | 
|  |    455 | \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    456 | \isanewline
 | 
|  |    457 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    458 | %
 | 
|  |    459 | \endisatagproof
 | 
|  |    460 | {\isafoldproof}%
 | 
|  |    461 | %
 | 
|  |    462 | \isadelimproof
 | 
|  |    463 | %
 | 
|  |    464 | \endisadelimproof
 | 
| 22317 |    465 | %
 | 
|  |    466 | \begin{isamarkuptext}%
 | 
|  |    467 | \noindent To finish our small algebra example, we add a \isa{group} class
 | 
|  |    468 |   with a corresponding instance:%
 | 
|  |    469 | \end{isamarkuptext}%
 | 
|  |    470 | \isamarkuptrue%
 | 
| 20967 |    471 | \ \ \ \ \isacommand{class}\isamarkupfalse%
 | 
|  |    472 | \ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline
 | 
|  |    473 | \ \ \ \ \ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}\isactrlloc {\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline
 | 
|  |    474 | \ \ \ \ \ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x\ {\isacharequal}\ \isactrlloc {\isasymone}{\isachardoublequoteclose}\isanewline
 | 
|  |    475 | \isanewline
 | 
|  |    476 | \ \ \ \ \isacommand{instance}\isamarkupfalse%
 | 
|  |    477 | \ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline
 | 
| 22317 |    478 | \ \ \ \ \ \ inverse{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymequiv}\ {\isacharminus}\ i{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    479 | %
 | 
|  |    480 | \isadelimproof
 | 
|  |    481 | \ \ \ \ %
 | 
|  |    482 | \endisadelimproof
 | 
|  |    483 | %
 | 
|  |    484 | \isatagproof
 | 
|  |    485 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    486 | \isanewline
 | 
|  |    487 | \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
 | 
|  |    488 | \ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline
 | 
|  |    489 | \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
 | 
|  |    490 | \ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
|  |    491 | \ simp\isanewline
 | 
|  |    492 | \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
 | 
|  |    493 | \ \isacommand{show}\isamarkupfalse%
 | 
| 22550 |    494 | \ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline
 | 
|  |    495 | \ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
 | 
| 20967 |    496 | \ mult{\isacharunderscore}int{\isacharunderscore}def\ \isakeyword{and}\ neutral{\isacharunderscore}int{\isacharunderscore}def\ \isakeyword{and}\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
|  |    497 | \isanewline
 | 
|  |    498 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    499 | %
 | 
|  |    500 | \endisatagproof
 | 
|  |    501 | {\isafoldproof}%
 | 
|  |    502 | %
 | 
|  |    503 | \isadelimproof
 | 
|  |    504 | %
 | 
|  |    505 | \endisadelimproof
 | 
|  |    506 | %
 | 
| 22317 |    507 | \isamarkupsection{Type classes as locales%
 | 
|  |    508 | }
 | 
|  |    509 | \isamarkuptrue%
 | 
|  |    510 | %
 | 
|  |    511 | \isamarkupsubsection{A look behind the scene%
 | 
|  |    512 | }
 | 
|  |    513 | \isamarkuptrue%
 | 
|  |    514 | %
 | 
|  |    515 | \begin{isamarkuptext}%
 | 
| 22479 |    516 | The example above gives an impression how Isar type classes work
 | 
|  |    517 |   in practice.  As stated in the introduction, classes also provide
 | 
|  |    518 |   a link to Isar's locale system.  Indeed, the logical core of a class
 | 
|  |    519 |   is nothing else than a locale:%
 | 
|  |    520 | \end{isamarkuptext}%
 | 
|  |    521 | \isamarkuptrue%
 | 
|  |    522 | \isacommand{class}\isamarkupfalse%
 | 
|  |    523 | \ idem\ {\isacharequal}\ type\ {\isacharplus}\isanewline
 | 
|  |    524 | \ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline
 | 
|  |    525 | \ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}%
 | 
|  |    526 | \begin{isamarkuptext}%
 | 
|  |    527 | \noindent essentially introduces the locale%
 | 
|  |    528 | \end{isamarkuptext}%
 | 
|  |    529 | \isamarkuptrue%
 | 
| 22317 |    530 | %
 | 
| 22479 |    531 | \isadelimML
 | 
|  |    532 | %
 | 
|  |    533 | \endisadelimML
 | 
|  |    534 | %
 | 
|  |    535 | \isatagML
 | 
|  |    536 | %
 | 
|  |    537 | \endisatagML
 | 
|  |    538 | {\isafoldML}%
 | 
|  |    539 | %
 | 
|  |    540 | \isadelimML
 | 
|  |    541 | %
 | 
|  |    542 | \endisadelimML
 | 
|  |    543 | \isacommand{locale}\isamarkupfalse%
 | 
|  |    544 | \ idem\ {\isacharequal}\isanewline
 | 
|  |    545 | \ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline
 | 
|  |    546 | \ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}%
 | 
|  |    547 | \begin{isamarkuptext}%
 | 
| 22550 |    548 | \noindent together with corresponding constant(s):%
 | 
| 22479 |    549 | \end{isamarkuptext}%
 | 
|  |    550 | \isamarkuptrue%
 | 
|  |    551 | \isacommand{consts}\isamarkupfalse%
 | 
| 22550 |    552 | \ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}%
 | 
|  |    553 | \begin{isamarkuptext}%
 | 
|  |    554 | \noindent The connection to the type system is done by means
 | 
|  |    555 |   of a primitive axclass%
 | 
|  |    556 | \end{isamarkuptext}%
 | 
|  |    557 | \isamarkuptrue%
 | 
| 22479 |    558 | \isacommand{axclass}\isamarkupfalse%
 | 
|  |    559 | \ idem\ {\isacharless}\ type\isanewline
 | 
|  |    560 | \ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}%
 | 
|  |    561 | \begin{isamarkuptext}%
 | 
| 22649 |    562 | \noindent together with a corresponding interpretation:%
 | 
| 22479 |    563 | \end{isamarkuptext}%
 | 
|  |    564 | \isamarkuptrue%
 | 
|  |    565 | \isacommand{interpretation}\isamarkupfalse%
 | 
|  |    566 | \ idem{\isacharunderscore}class{\isacharcolon}\isanewline
 | 
|  |    567 | \ \ idem\ {\isacharbrackleft}{\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isacharprime}a{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}{\isacharbrackright}\isanewline
 | 
|  |    568 | %
 | 
|  |    569 | \isadelimproof
 | 
|  |    570 | %
 | 
|  |    571 | \endisadelimproof
 | 
|  |    572 | %
 | 
|  |    573 | \isatagproof
 | 
|  |    574 | \isacommand{by}\isamarkupfalse%
 | 
| 22550 |    575 | \ unfold{\isacharunderscore}locales\ {\isacharparenleft}rule\ idem{\isacharparenright}%
 | 
| 22479 |    576 | \endisatagproof
 | 
|  |    577 | {\isafoldproof}%
 | 
|  |    578 | %
 | 
|  |    579 | \isadelimproof
 | 
|  |    580 | %
 | 
|  |    581 | \endisadelimproof
 | 
|  |    582 | %
 | 
|  |    583 | \isadelimML
 | 
|  |    584 | %
 | 
|  |    585 | \endisadelimML
 | 
|  |    586 | %
 | 
|  |    587 | \isatagML
 | 
|  |    588 | %
 | 
|  |    589 | \endisatagML
 | 
|  |    590 | {\isafoldML}%
 | 
|  |    591 | %
 | 
|  |    592 | \isadelimML
 | 
|  |    593 | %
 | 
|  |    594 | \endisadelimML
 | 
|  |    595 | %
 | 
|  |    596 | \begin{isamarkuptext}%
 | 
|  |    597 | This give you at hand the full power of the Isabelle module system;
 | 
|  |    598 |   conclusions in locale \isa{idem} are implicitly propagated
 | 
|  |    599 |   to class \isa{idem}.%
 | 
| 22317 |    600 | \end{isamarkuptext}%
 | 
|  |    601 | \isamarkuptrue%
 | 
|  |    602 | %
 | 
| 20967 |    603 | \isamarkupsubsection{Abstract reasoning%
 | 
|  |    604 | }
 | 
|  |    605 | \isamarkuptrue%
 | 
|  |    606 | %
 | 
|  |    607 | \begin{isamarkuptext}%
 | 
| 22479 |    608 | Isabelle locales enable reasoning at a general level, while results
 | 
| 20967 |    609 |   are implicitly transferred to all instances.  For example, we can
 | 
|  |    610 |   now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which
 | 
|  |    611 |   states that the function \isa{{\isacharparenleft}x\ {\isasymcirc}{\isacharparenright}} is injective:%
 | 
|  |    612 | \end{isamarkuptext}%
 | 
|  |    613 | \isamarkuptrue%
 | 
|  |    614 | \ \ \ \ \isacommand{lemma}\isamarkupfalse%
 | 
|  |    615 | \ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline
 | 
|  |    616 | %
 | 
|  |    617 | \isadelimproof
 | 
|  |    618 | \ \ \ \ %
 | 
|  |    619 | \endisadelimproof
 | 
|  |    620 | %
 | 
|  |    621 | \isatagproof
 | 
|  |    622 | \isacommand{proof}\isamarkupfalse%
 | 
|  |    623 | \isanewline
 | 
|  |    624 | \ \ \ \ \isacommand{assume}\isamarkupfalse%
 | 
|  |    625 | \ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline
 | 
| 22479 |    626 | \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
 | 
| 20967 |    627 | \ \isacommand{have}\isamarkupfalse%
 | 
|  |    628 | \ {\isachardoublequoteopen}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
|  |    629 | \ simp\isanewline
 | 
| 22479 |    630 | \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
 | 
| 20967 |    631 | \ \isacommand{have}\isamarkupfalse%
 | 
|  |    632 | \ {\isachardoublequoteopen}{\isacharparenleft}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x{\isacharparenright}\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x{\isacharparenright}\ \isactrlloc {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
 | 
|  |    633 | \ assoc\ \isacommand{by}\isamarkupfalse%
 | 
|  |    634 | \ simp\isanewline
 | 
| 22479 |    635 | \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
 | 
| 20967 |    636 | \ \isacommand{show}\isamarkupfalse%
 | 
|  |    637 | \ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
 | 
|  |    638 | \ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse%
 | 
|  |    639 | \ simp\isanewline
 | 
|  |    640 | \ \ \ \ \isacommand{next}\isamarkupfalse%
 | 
|  |    641 | \isanewline
 | 
|  |    642 | \ \ \ \ \isacommand{assume}\isamarkupfalse%
 | 
|  |    643 | \ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline
 | 
| 22479 |    644 | \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
 | 
| 20967 |    645 | \ \isacommand{show}\isamarkupfalse%
 | 
|  |    646 | \ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
|  |    647 | \ simp\isanewline
 | 
|  |    648 | \ \ \ \ \isacommand{qed}\isamarkupfalse%
 | 
|  |    649 | %
 | 
|  |    650 | \endisatagproof
 | 
|  |    651 | {\isafoldproof}%
 | 
|  |    652 | %
 | 
|  |    653 | \isadelimproof
 | 
|  |    654 | %
 | 
|  |    655 | \endisadelimproof
 | 
|  |    656 | %
 | 
|  |    657 | \begin{isamarkuptext}%
 | 
|  |    658 | \noindent Here the \qt{\isa{{\isasymIN}\ group}} target specification
 | 
|  |    659 |   indicates that the result is recorded within that context for later
 | 
| 22479 |    660 |   use.  This local theorem is also lifted to the global one \isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.  Since type \isa{int} has been made an instance of
 | 
|  |    661 |   \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.%
 | 
| 20967 |    662 | \end{isamarkuptext}%
 | 
|  |    663 | \isamarkuptrue%
 | 
|  |    664 | %
 | 
| 22317 |    665 | \isamarkupsection{Further issues%
 | 
| 20967 |    666 | }
 | 
|  |    667 | \isamarkuptrue%
 | 
|  |    668 | %
 | 
| 22317 |    669 | \isamarkupsubsection{Code generation%
 | 
| 20967 |    670 | }
 | 
|  |    671 | \isamarkuptrue%
 | 
|  |    672 | %
 | 
|  |    673 | \begin{isamarkuptext}%
 | 
| 22317 |    674 | Turning back to the first motivation for type classes,
 | 
|  |    675 |   namely overloading, it is obvious that overloading
 | 
|  |    676 |   stemming from \isa{{\isasymCLASS}} and \isa{{\isasymINSTANCE}}
 | 
|  |    677 |   statements naturally maps to Haskell type classes.
 | 
|  |    678 |   The code generator framework \cite{isabelle-codegen} 
 | 
|  |    679 |   takes this into account.  Concerning target languages
 | 
|  |    680 |   lacking type classes (e.g.~SML), type classes
 | 
|  |    681 |   are implemented by explicit dictionary construction.
 | 
|  |    682 |   As example, the natural power function on groups:%
 | 
| 20967 |    683 | \end{isamarkuptext}%
 | 
|  |    684 | \isamarkuptrue%
 | 
| 22317 |    685 | \ \ \ \ \isacommand{fun}\isamarkupfalse%
 | 
| 20967 |    686 | \isanewline
 | 
| 22479 |    687 | \ \ \ \ \ \ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoidl\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoidl{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | 
| 20967 |    688 | \ \ \ \ \ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline
 | 
| 22479 |    689 | \ \ \ \ \ \ {\isacharbar}\ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}\isanewline
 | 
| 20967 |    690 | \isanewline
 | 
|  |    691 | \ \ \ \ \isacommand{definition}\isamarkupfalse%
 | 
|  |    692 | \isanewline
 | 
| 22479 |    693 | \ \ \ \ \ \ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | 
| 20967 |    694 | \ \ \ \ \ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline
 | 
|  |    695 | \ \ \ \ \ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline
 | 
|  |    696 | \ \ \ \ \ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
|  |    697 | \isanewline
 | 
|  |    698 | \ \ \ \ \isacommand{definition}\isamarkupfalse%
 | 
|  |    699 | \isanewline
 | 
| 22317 |    700 | \ \ \ \ \ \ example\ {\isacharcolon}{\isacharcolon}\ int\ \isakeyword{where}\isanewline
 | 
| 20967 |    701 | \ \ \ \ \ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}%
 | 
|  |    702 | \begin{isamarkuptext}%
 | 
| 22317 |    703 | \noindent This maps to Haskell as:%
 | 
| 20967 |    704 | \end{isamarkuptext}%
 | 
|  |    705 | \isamarkuptrue%
 | 
| 22317 |    706 | \isacommand{code{\isacharunderscore}gen}\isamarkupfalse%
 | 
| 23015 |    707 | \ example\ \isakeyword{in}\ Haskell\ \isakeyword{file}\ {\isachardoublequoteopen}code{\isacharunderscore}examples{\isacharslash}{\isachardoublequoteclose}%
 | 
| 20967 |    708 | \begin{isamarkuptext}%
 | 
| 22317 |    709 | \lsthaskell{Thy/code_examples/Classes.hs}
 | 
|  |    710 | 
 | 
|  |    711 |   \noindent (we have left out all other modules).
 | 
|  |    712 | 
 | 
|  |    713 |   \noindent The whole code in SML with explicit dictionary passing:%
 | 
|  |    714 | \end{isamarkuptext}%
 | 
|  |    715 | \isamarkuptrue%
 | 
|  |    716 | \isacommand{code{\isacharunderscore}gen}\isamarkupfalse%
 | 
| 23015 |    717 | \ example\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}code{\isacharunderscore}examples{\isacharslash}classes{\isachardot}ML{\isachardoublequoteclose}%
 | 
| 22317 |    718 | \begin{isamarkuptext}%
 | 
|  |    719 | \lstsml{Thy/code_examples/classes.ML}%
 | 
| 20967 |    720 | \end{isamarkuptext}%
 | 
|  |    721 | \isamarkuptrue%
 | 
|  |    722 | %
 | 
|  |    723 | \isadelimtheory
 | 
|  |    724 | %
 | 
|  |    725 | \endisadelimtheory
 | 
|  |    726 | %
 | 
|  |    727 | \isatagtheory
 | 
|  |    728 | \isacommand{end}\isamarkupfalse%
 | 
|  |    729 | %
 | 
|  |    730 | \endisatagtheory
 | 
|  |    731 | {\isafoldtheory}%
 | 
|  |    732 | %
 | 
|  |    733 | \isadelimtheory
 | 
|  |    734 | %
 | 
|  |    735 | \endisadelimtheory
 | 
|  |    736 | \isanewline
 | 
|  |    737 | \end{isabellebody}%
 | 
|  |    738 | %%% Local Variables:
 | 
|  |    739 | %%% mode: latex
 | 
|  |    740 | %%% TeX-master: "root"
 | 
|  |    741 | %%% End:
 |