author | wenzelm |
Tue, 24 Jan 2023 11:30:56 +0100 | |
changeset 77066 | 72d87e32b062 |
parent 76987 | 4c275405faae |
child 78664 | d052d61da398 |
permissions | -rw-r--r-- |
20946 | 1 |
theory Classes |
28565 | 2 |
imports Main Setup |
20946 | 3 |
begin |
4 |
||
61419 | 5 |
section \<open>Introduction\<close> |
20946 | 6 |
|
61419 | 7 |
text \<open> |
76987 | 8 |
Type classes were introduced by Wadler and Blott \<^cite>\<open>wadler89how\<close> |
31691 | 9 |
into the Haskell language to allow for a reasonable implementation |
22317 | 10 |
of overloading\footnote{throughout this tutorial, we are referring |
38812 | 11 |
to classical Haskell 1.0 type classes, not considering later |
12 |
additions in expressiveness}. As a canonical example, a polymorphic |
|
69505 | 13 |
equality function \<open>eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool\<close> which is overloaded on |
14 |
different types for \<open>\<alpha>\<close>, which is achieved by splitting |
|
15 |
introduction of the \<open>eq\<close> function from its overloaded |
|
16 |
definitions by means of \<open>class\<close> and \<open>instance\<close> |
|
38812 | 17 |
declarations: \footnote{syntax here is a kind of isabellized |
18 |
Haskell} |
|
22317 | 19 |
|
28565 | 20 |
\begin{quote} |
21 |
||
69505 | 22 |
\<^noindent> \<open>class eq where\<close> \\ |
23 |
\hspace*{2ex}\<open>eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool\<close> |
|
20946 | 24 |
|
69505 | 25 |
\<^medskip>\<^noindent> \<open>instance nat :: eq where\<close> \\ |
26 |
\hspace*{2ex}\<open>eq 0 0 = True\<close> \\ |
|
27 |
\hspace*{2ex}\<open>eq 0 _ = False\<close> \\ |
|
28 |
\hspace*{2ex}\<open>eq _ 0 = False\<close> \\ |
|
29 |
\hspace*{2ex}\<open>eq (Suc n) (Suc m) = eq n m\<close> |
|
22317 | 30 |
|
69505 | 31 |
\<^medskip>\<^noindent> \<open>instance (\<alpha>::eq, \<beta>::eq) pair :: eq where\<close> \\ |
32 |
\hspace*{2ex}\<open>eq (x1, y1) (x2, y2) = eq x1 x2 \<and> eq y1 y2\<close> |
|
20946 | 33 |
|
69505 | 34 |
\<^medskip>\<^noindent> \<open>class ord extends eq where\<close> \\ |
35 |
\hspace*{2ex}\<open>less_eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool\<close> \\ |
|
36 |
\hspace*{2ex}\<open>less :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool\<close> |
|
22317 | 37 |
|
28565 | 38 |
\end{quote} |
39 |
||
61438 | 40 |
\<^noindent> Type variables are annotated with (finitely many) classes; |
22317 | 41 |
these annotations are assertions that a particular polymorphic type |
42 |
provides definitions for overloaded functions. |
|
43 |
||
38812 | 44 |
Indeed, type classes not only allow for simple overloading but form |
45 |
a generic calculus, an instance of order-sorted algebra |
|
76987 | 46 |
\<^cite>\<open>"nipkow-sorts93" and "Nipkow-Prehofer:1993" and "Wenzel:1997:TPHOL"\<close>. |
20946 | 47 |
|
38812 | 48 |
From a software engineering point of view, type classes roughly |
49 |
correspond to interfaces in object-oriented languages like Java; so, |
|
50 |
it is naturally desirable that type classes do not only provide |
|
51 |
functions (class parameters) but also state specifications |
|
69505 | 52 |
implementations must obey. For example, the \<open>class eq\<close> |
22550 | 53 |
above could be given the following specification, demanding that |
69505 | 54 |
\<open>class eq\<close> is an equivalence relation obeying reflexivity, |
22550 | 55 |
symmetry and transitivity: |
22317 | 56 |
|
28565 | 57 |
\begin{quote} |
22317 | 58 |
|
69505 | 59 |
\<^noindent> \<open>class eq where\<close> \\ |
60 |
\hspace*{2ex}\<open>eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool\<close> \\ |
|
61 |
\<open>satisfying\<close> \\ |
|
62 |
\hspace*{2ex}\<open>refl: eq x x\<close> \\ |
|
63 |
\hspace*{2ex}\<open>sym: eq x y \<longleftrightarrow> eq x y\<close> \\ |
|
64 |
\hspace*{2ex}\<open>trans: eq x y \<and> eq y z \<longrightarrow> eq x z\<close> |
|
28565 | 65 |
|
66 |
\end{quote} |
|
67 |
||
61438 | 68 |
\<^noindent> From a theoretical point of view, type classes are |
38812 | 69 |
lightweight modules; Haskell type classes may be emulated by SML |
76987 | 70 |
functors \<^cite>\<open>classes_modules\<close>. Isabelle/Isar offers a discipline |
38812 | 71 |
of type classes which brings all those aspects together: |
22317 | 72 |
|
73 |
\begin{enumerate} |
|
24991 | 74 |
\item specifying abstract parameters together with |
22317 | 75 |
corresponding specifications, |
28540 | 76 |
\item instantiating those abstract parameters by a particular |
22317 | 77 |
type |
78 |
\item in connection with a ``less ad-hoc'' approach to overloading, |
|
31691 | 79 |
\item with a direct link to the Isabelle module system: |
76987 | 80 |
locales \<^cite>\<open>"kammueller-locales"\<close>. |
22317 | 81 |
\end{enumerate} |
82 |
||
61438 | 83 |
\<^noindent> Isar type classes also directly support code generation in |
38812 | 84 |
a Haskell like fashion. Internally, they are mapped to more |
76987 | 85 |
primitive Isabelle concepts \<^cite>\<open>"Haftmann-Wenzel:2006:classes"\<close>. |
22317 | 86 |
|
38812 | 87 |
This tutorial demonstrates common elements of structured |
88 |
specifications and abstract reasoning with type classes by the |
|
89 |
algebraic hierarchy of semigroups, monoids and groups. Our |
|
76987 | 90 |
background theory is that of Isabelle/HOL \<^cite>\<open>"isa-tutorial"\<close>, for |
38812 | 91 |
which some familiarity is assumed. |
61419 | 92 |
\<close> |
20946 | 93 |
|
61419 | 94 |
section \<open>A simple algebra example \label{sec:example}\<close> |
20946 | 95 |
|
61419 | 96 |
subsection \<open>Class definition\<close> |
20946 | 97 |
|
61419 | 98 |
text \<open> |
69505 | 99 |
Depending on an arbitrary type \<open>\<alpha>\<close>, class \<open>semigroup\<close> introduces a binary operator \<open>(\<otimes>)\<close> that is |
20946 | 100 |
assumed to be associative: |
61419 | 101 |
\<close> |
20946 | 102 |
|
29705 | 103 |
class %quote semigroup = |
28565 | 104 |
fixes mult :: "\<alpha> \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" (infixl "\<otimes>" 70) |
105 |
assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" |
|
20946 | 106 |
|
61419 | 107 |
text \<open> |
61438 | 108 |
\<^noindent> This @{command class} specification consists of two parts: |
38812 | 109 |
the \qn{operational} part names the class parameter (@{element |
110 |
"fixes"}), the \qn{logical} part specifies properties on them |
|
111 |
(@{element "assumes"}). The local @{element "fixes"} and @{element |
|
112 |
"assumes"} are lifted to the theory toplevel, yielding the global |
|
61076 | 113 |
parameter @{term [source] "mult :: \<alpha>::semigroup \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"} and the |
114 |
global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y z :: |
|
115 |
\<alpha>::semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}. |
|
61419 | 116 |
\<close> |
20946 | 117 |
|
118 |
||
61419 | 119 |
subsection \<open>Class instantiation \label{sec:class_inst}\<close> |
20946 | 120 |
|
61419 | 121 |
text \<open> |
69597 | 122 |
The concrete type \<^typ>\<open>int\<close> is made a \<^class>\<open>semigroup\<close> instance |
69505 | 123 |
by providing a suitable definition for the class parameter \<open>(\<otimes>)\<close> and a proof for the specification of @{fact assoc}. This is |
38812 | 124 |
accomplished by the @{command instantiation} target: |
61419 | 125 |
\<close> |
20946 | 126 |
|
28565 | 127 |
instantiation %quote int :: semigroup |
128 |
begin |
|
25533 | 129 |
|
28565 | 130 |
definition %quote |
61076 | 131 |
mult_int_def: "i \<otimes> j = i + (j::int)" |
25533 | 132 |
|
28565 | 133 |
instance %quote proof |
134 |
fix i j k :: int have "(i + j) + k = i + (j + k)" by simp |
|
135 |
then show "(i \<otimes> j) \<otimes> k = i \<otimes> (j \<otimes> k)" |
|
28566 | 136 |
unfolding mult_int_def . |
28565 | 137 |
qed |
20946 | 138 |
|
28565 | 139 |
end %quote |
25533 | 140 |
|
61419 | 141 |
text \<open> |
61438 | 142 |
\<^noindent> @{command instantiation} defines class parameters at a |
38812 | 143 |
particular instance using common specification tools (here, |
144 |
@{command definition}). The concluding @{command instance} opens a |
|
145 |
proof that the given parameters actually conform to the class |
|
146 |
specification. Note that the first proof step is the @{method |
|
62939 | 147 |
standard} method, which for such instance proofs maps to the @{method |
38812 | 148 |
intro_classes} method. This reduces an instance judgement to the |
149 |
relevant primitive proof goals; typically it is the first method |
|
150 |
applied in an instantiation proof. |
|
22317 | 151 |
|
69597 | 152 |
From now on, the type-checker will consider \<^typ>\<open>int\<close> as a \<^class>\<open>semigroup\<close> automatically, i.e.\ any general results are immediately |
38812 | 153 |
available on concrete instances. |
28565 | 154 |
|
69597 | 155 |
\<^medskip> Another instance of \<^class>\<open>semigroup\<close> yields the natural |
38812 | 156 |
numbers: |
61419 | 157 |
\<close> |
20946 | 158 |
|
28565 | 159 |
instantiation %quote nat :: semigroup |
160 |
begin |
|
25533 | 161 |
|
28565 | 162 |
primrec %quote mult_nat where |
61076 | 163 |
"(0::nat) \<otimes> n = n" |
28565 | 164 |
| "Suc m \<otimes> n = Suc (m \<otimes> n)" |
25533 | 165 |
|
28565 | 166 |
instance %quote proof |
167 |
fix m n q :: nat |
|
168 |
show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)" |
|
169 |
by (induct m) auto |
|
170 |
qed |
|
20946 | 171 |
|
28565 | 172 |
end %quote |
20946 | 173 |
|
61419 | 174 |
text \<open> |
69505 | 175 |
\<^noindent> Note the occurrence of the name \<open>mult_nat\<close> in the |
38812 | 176 |
primrec declaration; by default, the local name of a class operation |
69505 | 177 |
\<open>f\<close> to be instantiated on type constructor \<open>\<kappa>\<close> is |
178 |
mangled as \<open>f_\<kappa>\<close>. In case of uncertainty, these names may be |
|
58842 | 179 |
inspected using the @{command "print_context"} command. |
61419 | 180 |
\<close> |
25871 | 181 |
|
61419 | 182 |
subsection \<open>Lifting and parametric types\<close> |
25247 | 183 |
|
61419 | 184 |
text \<open> |
38812 | 185 |
Overloaded definitions given at a class instantiation may include |
186 |
recursion over the syntactic structure of types. As a canonical |
|
187 |
example, we model product semigroups using our simple algebra: |
|
61419 | 188 |
\<close> |
25247 | 189 |
|
37706 | 190 |
instantiation %quote prod :: (semigroup, semigroup) semigroup |
28565 | 191 |
begin |
25533 | 192 |
|
28565 | 193 |
definition %quote |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51143
diff
changeset
|
194 |
mult_prod_def: "p\<^sub>1 \<otimes> p\<^sub>2 = (fst p\<^sub>1 \<otimes> fst p\<^sub>2, snd p\<^sub>1 \<otimes> snd p\<^sub>2)" |
25533 | 195 |
|
28565 | 196 |
instance %quote proof |
61076 | 197 |
fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\<alpha>::semigroup \<times> \<beta>::semigroup" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51143
diff
changeset
|
198 |
show "p\<^sub>1 \<otimes> p\<^sub>2 \<otimes> p\<^sub>3 = p\<^sub>1 \<otimes> (p\<^sub>2 \<otimes> p\<^sub>3)" |
28566 | 199 |
unfolding mult_prod_def by (simp add: assoc) |
28565 | 200 |
qed |
25247 | 201 |
|
28565 | 202 |
end %quote |
25533 | 203 |
|
61419 | 204 |
text \<open> |
61438 | 205 |
\<^noindent> Associativity of product semigroups is established using |
69505 | 206 |
the definition of \<open>(\<otimes>)\<close> on products and the hypothetical |
38812 | 207 |
associativity of the type components; these hypotheses are |
69597 | 208 |
legitimate due to the \<^class>\<open>semigroup\<close> constraints imposed on the |
38812 | 209 |
type components by the @{command instance} proposition. Indeed, |
210 |
this pattern often occurs with parametric types and type classes. |
|
61419 | 211 |
\<close> |
25247 | 212 |
|
213 |
||
61419 | 214 |
subsection \<open>Subclassing\<close> |
20946 | 215 |
|
61419 | 216 |
text \<open> |
69505 | 217 |
We define a subclass \<open>monoidl\<close> (a semigroup with a left-hand |
69597 | 218 |
neutral) by extending \<^class>\<open>semigroup\<close> with one additional |
69505 | 219 |
parameter \<open>neutral\<close> together with its characteristic property: |
61419 | 220 |
\<close> |
20946 | 221 |
|
28565 | 222 |
class %quote monoidl = semigroup + |
223 |
fixes neutral :: "\<alpha>" ("\<one>") |
|
224 |
assumes neutl: "\<one> \<otimes> x = x" |
|
20946 | 225 |
|
61419 | 226 |
text \<open> |
61438 | 227 |
\<^noindent> Again, we prove some instances, by providing suitable |
38812 | 228 |
parameter definitions and proofs for the additional specifications. |
229 |
Observe that instantiations for types with the same arity may be |
|
230 |
simultaneous: |
|
61419 | 231 |
\<close> |
20946 | 232 |
|
28566 | 233 |
instantiation %quote nat and int :: monoidl |
234 |
begin |
|
25533 | 235 |
|
28566 | 236 |
definition %quote |
61076 | 237 |
neutral_nat_def: "\<one> = (0::nat)" |
25533 | 238 |
|
28566 | 239 |
definition %quote |
61076 | 240 |
neutral_int_def: "\<one> = (0::int)" |
25533 | 241 |
|
28566 | 242 |
instance %quote proof |
243 |
fix n :: nat |
|
244 |
show "\<one> \<otimes> n = n" |
|
245 |
unfolding neutral_nat_def by simp |
|
246 |
next |
|
247 |
fix k :: int |
|
248 |
show "\<one> \<otimes> k = k" |
|
249 |
unfolding neutral_int_def mult_int_def by simp |
|
250 |
qed |
|
20946 | 251 |
|
28566 | 252 |
end %quote |
25533 | 253 |
|
37706 | 254 |
instantiation %quote prod :: (monoidl, monoidl) monoidl |
28566 | 255 |
begin |
25533 | 256 |
|
28566 | 257 |
definition %quote |
258 |
neutral_prod_def: "\<one> = (\<one>, \<one>)" |
|
25533 | 259 |
|
28566 | 260 |
instance %quote proof |
61076 | 261 |
fix p :: "\<alpha>::monoidl \<times> \<beta>::monoidl" |
28566 | 262 |
show "\<one> \<otimes> p = p" |
263 |
unfolding neutral_prod_def mult_prod_def by (simp add: neutl) |
|
264 |
qed |
|
20946 | 265 |
|
28566 | 266 |
end %quote |
25533 | 267 |
|
61419 | 268 |
text \<open> |
61438 | 269 |
\<^noindent> Fully-fledged monoids are modelled by another subclass, |
24991 | 270 |
which does not add new parameters but tightens the specification: |
61419 | 271 |
\<close> |
20946 | 272 |
|
28566 | 273 |
class %quote monoid = monoidl + |
274 |
assumes neutr: "x \<otimes> \<one> = x" |
|
20946 | 275 |
|
28566 | 276 |
instantiation %quote nat and int :: monoid |
277 |
begin |
|
25533 | 278 |
|
28566 | 279 |
instance %quote proof |
280 |
fix n :: nat |
|
281 |
show "n \<otimes> \<one> = n" |
|
282 |
unfolding neutral_nat_def by (induct n) simp_all |
|
283 |
next |
|
284 |
fix k :: int |
|
285 |
show "k \<otimes> \<one> = k" |
|
286 |
unfolding neutral_int_def mult_int_def by simp |
|
287 |
qed |
|
25247 | 288 |
|
28566 | 289 |
end %quote |
25533 | 290 |
|
37706 | 291 |
instantiation %quote prod :: (monoid, monoid) monoid |
28566 | 292 |
begin |
25533 | 293 |
|
28566 | 294 |
instance %quote proof |
61076 | 295 |
fix p :: "\<alpha>::monoid \<times> \<beta>::monoid" |
28566 | 296 |
show "p \<otimes> \<one> = p" |
297 |
unfolding neutral_prod_def mult_prod_def by (simp add: neutr) |
|
298 |
qed |
|
22317 | 299 |
|
28566 | 300 |
end %quote |
25533 | 301 |
|
61419 | 302 |
text \<open> |
69505 | 303 |
\<^noindent> To finish our small algebra example, we add a \<open>group\<close> class with a corresponding instance: |
61419 | 304 |
\<close> |
20946 | 305 |
|
28566 | 306 |
class %quote group = monoidl + |
307 |
fixes inverse :: "\<alpha> \<Rightarrow> \<alpha>" ("(_\<div>)" [1000] 999) |
|
308 |
assumes invl: "x\<div> \<otimes> x = \<one>" |
|
20946 | 309 |
|
28566 | 310 |
instantiation %quote int :: group |
311 |
begin |
|
312 |
||
313 |
definition %quote |
|
61076 | 314 |
inverse_int_def: "i\<div> = - (i::int)" |
25533 | 315 |
|
28566 | 316 |
instance %quote proof |
317 |
fix i :: int |
|
318 |
have "-i + i = 0" by simp |
|
319 |
then show "i\<div> \<otimes> i = \<one>" |
|
320 |
unfolding mult_int_def neutral_int_def inverse_int_def . |
|
321 |
qed |
|
25533 | 322 |
|
28566 | 323 |
end %quote |
20946 | 324 |
|
25533 | 325 |
|
61419 | 326 |
section \<open>Type classes as locales\<close> |
22317 | 327 |
|
61419 | 328 |
subsection \<open>A look behind the scenes\<close> |
22317 | 329 |
|
61419 | 330 |
text \<open> |
38812 | 331 |
The example above gives an impression how Isar type classes work in |
332 |
practice. As stated in the introduction, classes also provide a |
|
333 |
link to Isar's locale system. Indeed, the logical core of a class |
|
31691 | 334 |
is nothing other than a locale: |
61419 | 335 |
\<close> |
22347 | 336 |
|
29705 | 337 |
class %quote idem = |
22347 | 338 |
fixes f :: "\<alpha> \<Rightarrow> \<alpha>" |
339 |
assumes idem: "f (f x) = f x" |
|
22317 | 340 |
|
61419 | 341 |
text \<open> |
61438 | 342 |
\<^noindent> essentially introduces the locale |
61419 | 343 |
\<close> (*<*)setup %invisible \<open>Sign.add_path "foo"\<close> |
30227 | 344 |
(*>*) |
28566 | 345 |
locale %quote idem = |
22347 | 346 |
fixes f :: "\<alpha> \<Rightarrow> \<alpha>" |
347 |
assumes idem: "f (f x) = f x" |
|
348 |
||
61438 | 349 |
text \<open>\<^noindent> together with corresponding constant(s):\<close> |
22347 | 350 |
|
28566 | 351 |
consts %quote f :: "\<alpha> \<Rightarrow> \<alpha>" |
22347 | 352 |
|
61419 | 353 |
text \<open> |
61438 | 354 |
\<^noindent> The connection to the type system is done by means of a |
69505 | 355 |
primitive type class \<open>idem\<close>, together with a corresponding |
55385
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
wenzelm
parents:
53015
diff
changeset
|
356 |
interpretation: |
61419 | 357 |
\<close> |
22347 | 358 |
|
29513 | 359 |
interpretation %quote idem_class: |
61076 | 360 |
idem "f :: (\<alpha>::idem) \<Rightarrow> \<alpha>" |
55385
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
wenzelm
parents:
53015
diff
changeset
|
361 |
(*<*)sorry(*>*) |
28565 | 362 |
|
61419 | 363 |
text \<open> |
61438 | 364 |
\<^noindent> This gives you the full power of the Isabelle module system; |
69505 | 365 |
conclusions in locale \<open>idem\<close> are implicitly propagated |
366 |
to class \<open>idem\<close>. |
|
61419 | 367 |
\<close> (*<*)setup %invisible \<open>Sign.parent_path\<close> |
30227 | 368 |
(*>*) |
61419 | 369 |
subsection \<open>Abstract reasoning\<close> |
20946 | 370 |
|
61419 | 371 |
text \<open> |
22347 | 372 |
Isabelle locales enable reasoning at a general level, while results |
20946 | 373 |
are implicitly transferred to all instances. For example, we can |
69505 | 374 |
now establish the \<open>left_cancel\<close> lemma for groups, which |
375 |
states that the function \<open>(x \<otimes>)\<close> is injective: |
|
61419 | 376 |
\<close> |
20946 | 377 |
|
28566 | 378 |
lemma %quote (in group) left_cancel: "x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z" |
379 |
proof |
|
380 |
assume "x \<otimes> y = x \<otimes> z" |
|
381 |
then have "x\<div> \<otimes> (x \<otimes> y) = x\<div> \<otimes> (x \<otimes> z)" by simp |
|
382 |
then have "(x\<div> \<otimes> x) \<otimes> y = (x\<div> \<otimes> x) \<otimes> z" using assoc by simp |
|
383 |
then show "y = z" using neutl and invl by simp |
|
384 |
next |
|
385 |
assume "y = z" |
|
386 |
then show "x \<otimes> y = x \<otimes> z" by simp |
|
387 |
qed |
|
20946 | 388 |
|
61419 | 389 |
text \<open> |
69597 | 390 |
\<^noindent> Here the \qt{@{keyword "in"} \<^class>\<open>group\<close>} target |
38812 | 391 |
specification indicates that the result is recorded within that |
392 |
context for later use. This local theorem is also lifted to the |
|
61076 | 393 |
global one @{fact "group.left_cancel:"} @{prop [source] "\<And>x y z :: |
69505 | 394 |
\<alpha>::group. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}. Since type \<open>int\<close> has been |
395 |
made an instance of \<open>group\<close> before, we may refer to that |
|
61076 | 396 |
fact as well: @{prop [source] "\<And>x y z :: int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = |
38812 | 397 |
z"}. |
61419 | 398 |
\<close> |
20946 | 399 |
|
400 |
||
61419 | 401 |
subsection \<open>Derived definitions\<close> |
23956 | 402 |
|
61419 | 403 |
text \<open> |
35282
8fd9d555d04d
dropped references to old axclass from documentation
haftmann
parents:
31931
diff
changeset
|
404 |
Isabelle locales are targets which support local definitions: |
61419 | 405 |
\<close> |
23956 | 406 |
|
28566 | 407 |
primrec %quote (in monoid) pow_nat :: "nat \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where |
408 |
"pow_nat 0 x = \<one>" |
|
409 |
| "pow_nat (Suc n) x = x \<otimes> pow_nat n x" |
|
20946 | 410 |
|
61419 | 411 |
text \<open> |
69505 | 412 |
\<^noindent> If the locale \<open>group\<close> is also a class, this local |
38812 | 413 |
definition is propagated onto a global definition of @{term [source] |
61076 | 414 |
"pow_nat :: nat \<Rightarrow> \<alpha>::monoid \<Rightarrow> \<alpha>::monoid"} with corresponding theorems |
23956 | 415 |
|
416 |
@{thm pow_nat.simps [no_vars]}. |
|
20946 | 417 |
|
61438 | 418 |
\<^noindent> As you can see from this example, for local definitions |
38812 | 419 |
you may use any specification tool which works together with |
420 |
locales, such as Krauss's recursive function package |
|
76987 | 421 |
\<^cite>\<open>krauss2006\<close>. |
61419 | 422 |
\<close> |
23956 | 423 |
|
424 |
||
61419 | 425 |
subsection \<open>A functor analogy\<close> |
25247 | 426 |
|
61419 | 427 |
text \<open> |
38812 | 428 |
We introduced Isar classes by analogy to type classes in functional |
429 |
programming; if we reconsider this in the context of what has been |
|
430 |
said about type classes and locales, we can drive this analogy |
|
431 |
further by stating that type classes essentially correspond to |
|
432 |
functors that have a canonical interpretation as type classes. |
|
433 |
There is also the possibility of other interpretations. For |
|
69505 | 434 |
example, \<open>list\<close>s also form a monoid with \<open>append\<close> and |
69597 | 435 |
\<^term>\<open>[]\<close> as operations, but it seems inappropriate to apply to |
38812 | 436 |
lists the same operations as for genuinely algebraic types. In such |
437 |
a case, we can simply make a particular interpretation of monoids |
|
438 |
for lists: |
|
61419 | 439 |
\<close> |
25247 | 440 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
30227
diff
changeset
|
441 |
interpretation %quote list_monoid: monoid append "[]" |
28947
ac1a14b5a085
unfold_locales is default method - no need for explicit references
haftmann
parents:
28566
diff
changeset
|
442 |
proof qed auto |
25247 | 443 |
|
61419 | 444 |
text \<open> |
61438 | 445 |
\<^noindent> This enables us to apply facts on monoids |
25247 | 446 |
to lists, e.g. @{thm list_monoid.neutl [no_vars]}. |
447 |
||
448 |
When using this interpretation pattern, it may also |
|
449 |
be appropriate to map derived definitions accordingly: |
|
61419 | 450 |
\<close> |
25247 | 451 |
|
28566 | 452 |
primrec %quote replicate :: "nat \<Rightarrow> \<alpha> list \<Rightarrow> \<alpha> list" where |
453 |
"replicate 0 _ = []" |
|
454 |
| "replicate (Suc n) xs = xs @ replicate n xs" |
|
25247 | 455 |
|
61566
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents:
61438
diff
changeset
|
456 |
interpretation %quote list_monoid: monoid append "[]" rewrites |
28566 | 457 |
"monoid.pow_nat append [] = replicate" |
458 |
proof - |
|
29513 | 459 |
interpret monoid append "[]" .. |
28566 | 460 |
show "monoid.pow_nat append [] = replicate" |
461 |
proof |
|
462 |
fix n |
|
463 |
show "monoid.pow_nat append [] n = replicate n" |
|
464 |
by (induct n) auto |
|
465 |
qed |
|
466 |
qed intro_locales |
|
25247 | 467 |
|
61419 | 468 |
text \<open> |
61438 | 469 |
\<^noindent> This pattern is also helpful to reuse abstract |
31255 | 470 |
specifications on the \emph{same} type. For example, think of a |
69597 | 471 |
class \<open>preorder\<close>; for type \<^typ>\<open>nat\<close>, there are at least two |
31255 | 472 |
possible instances: the natural order or the order induced by the |
473 |
divides relation. But only one of these instances can be used for |
|
69505 | 474 |
@{command instantiation}; using the locale behind the class \<open>preorder\<close>, it is still possible to utilise the same abstract |
31255 | 475 |
specification again using @{command interpretation}. |
61419 | 476 |
\<close> |
25247 | 477 |
|
61419 | 478 |
subsection \<open>Additional subclass relations\<close> |
22347 | 479 |
|
61419 | 480 |
text \<open> |
69505 | 481 |
Any \<open>group\<close> is also a \<open>monoid\<close>; this can be made |
31255 | 482 |
explicit by claiming an additional subclass relation, together with |
483 |
a proof of the logical difference: |
|
61419 | 484 |
\<close> |
22347 | 485 |
|
28566 | 486 |
subclass %quote (in group) monoid |
28947
ac1a14b5a085
unfold_locales is default method - no need for explicit references
haftmann
parents:
28566
diff
changeset
|
487 |
proof |
28566 | 488 |
fix x |
489 |
from invl have "x\<div> \<otimes> x = \<one>" by simp |
|
490 |
with assoc [symmetric] neutl invl have "x\<div> \<otimes> (x \<otimes> \<one>) = x\<div> \<otimes> x" by simp |
|
491 |
with left_cancel show "x \<otimes> \<one> = x" by simp |
|
492 |
qed |
|
23956 | 493 |
|
61419 | 494 |
text \<open> |
38812 | 495 |
The logical proof is carried out on the locale level. Afterwards it |
69505 | 496 |
is propagated to the type system, making \<open>group\<close> an instance |
497 |
of \<open>monoid\<close> by adding an additional edge to the graph of |
|
38812 | 498 |
subclass relations (\figref{fig:subclass}). |
25247 | 499 |
|
500 |
\begin{figure}[htbp] |
|
501 |
\begin{center} |
|
502 |
\small |
|
503 |
\unitlength 0.6mm |
|
504 |
\begin{picture}(40,60)(0,0) |
|
69505 | 505 |
\put(20,60){\makebox(0,0){\<open>semigroup\<close>}} |
506 |
\put(20,40){\makebox(0,0){\<open>monoidl\<close>}} |
|
507 |
\put(00,20){\makebox(0,0){\<open>monoid\<close>}} |
|
508 |
\put(40,00){\makebox(0,0){\<open>group\<close>}} |
|
25247 | 509 |
\put(20,55){\vector(0,-1){10}} |
510 |
\put(15,35){\vector(-1,-1){10}} |
|
511 |
\put(25,35){\vector(1,-3){10}} |
|
512 |
\end{picture} |
|
513 |
\hspace{8em} |
|
514 |
\begin{picture}(40,60)(0,0) |
|
69505 | 515 |
\put(20,60){\makebox(0,0){\<open>semigroup\<close>}} |
516 |
\put(20,40){\makebox(0,0){\<open>monoidl\<close>}} |
|
517 |
\put(00,20){\makebox(0,0){\<open>monoid\<close>}} |
|
518 |
\put(40,00){\makebox(0,0){\<open>group\<close>}} |
|
25247 | 519 |
\put(20,55){\vector(0,-1){10}} |
520 |
\put(15,35){\vector(-1,-1){10}} |
|
521 |
\put(05,15){\vector(3,-1){30}} |
|
522 |
\end{picture} |
|
523 |
\caption{Subclass relationship of monoids and groups: |
|
524 |
before and after establishing the relationship |
|
69505 | 525 |
\<open>group \<subseteq> monoid\<close>; transitive edges are left out.} |
25247 | 526 |
\label{fig:subclass} |
527 |
\end{center} |
|
528 |
\end{figure} |
|
30227 | 529 |
|
69505 | 530 |
For illustration, a derived definition in \<open>group\<close> using \<open>pow_nat\<close> |
61419 | 531 |
\<close> |
23956 | 532 |
|
28565 | 533 |
definition %quote (in group) pow_int :: "int \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where |
534 |
"pow_int k x = (if k >= 0 |
|
535 |
then pow_nat (nat k) x |
|
536 |
else (pow_nat (nat (- k)) x)\<div>)" |
|
23956 | 537 |
|
61419 | 538 |
text \<open> |
61438 | 539 |
\<^noindent> yields the global definition of @{term [source] "pow_int :: |
61076 | 540 |
int \<Rightarrow> \<alpha>::group \<Rightarrow> \<alpha>::group"} with the corresponding theorem @{thm |
38812 | 541 |
pow_int_def [no_vars]}. |
61419 | 542 |
\<close> |
23956 | 543 |
|
61419 | 544 |
subsection \<open>A note on syntax\<close> |
25868 | 545 |
|
61419 | 546 |
text \<open> |
38812 | 547 |
As a convenience, class context syntax allows references to local |
548 |
class operations and their global counterparts uniformly; type |
|
549 |
inference resolves ambiguities. For example: |
|
61419 | 550 |
\<close> |
25868 | 551 |
|
28565 | 552 |
context %quote semigroup |
25868 | 553 |
begin |
554 |
||
67406 | 555 |
term %quote "x \<otimes> y" \<comment> \<open>example 1\<close> |
556 |
term %quote "(x::nat) \<otimes> y" \<comment> \<open>example 2\<close> |
|
25868 | 557 |
|
28566 | 558 |
end %quote |
25868 | 559 |
|
67406 | 560 |
term %quote "x \<otimes> y" \<comment> \<open>example 3\<close> |
25868 | 561 |
|
61419 | 562 |
text \<open> |
61438 | 563 |
\<^noindent> Here in example 1, the term refers to the local class |
69505 | 564 |
operation \<open>mult [\<alpha>]\<close>, whereas in example 2 the type |
565 |
constraint enforces the global class operation \<open>mult [nat]\<close>. |
|
38812 | 566 |
In the global context in example 3, the reference is to the |
69505 | 567 |
polymorphic global class operation \<open>mult [?\<alpha> :: semigroup]\<close>. |
61419 | 568 |
\<close> |
22347 | 569 |
|
61419 | 570 |
section \<open>Further issues\<close> |
29705 | 571 |
|
61419 | 572 |
subsection \<open>Type classes and code generation\<close> |
22317 | 573 |
|
61419 | 574 |
text \<open> |
38812 | 575 |
Turning back to the first motivation for type classes, namely |
576 |
overloading, it is obvious that overloading stemming from @{command |
|
577 |
class} statements and @{command instantiation} targets naturally |
|
578 |
maps to Haskell type classes. The code generator framework |
|
76987 | 579 |
\<^cite>\<open>"isabelle-codegen"\<close> takes this into account. If the target |
38812 | 580 |
language (e.g.~SML) lacks type classes, then they are implemented by |
581 |
an explicit dictionary construction. As example, let's go back to |
|
582 |
the power function: |
|
61419 | 583 |
\<close> |
22317 | 584 |
|
28565 | 585 |
definition %quote example :: int where |
586 |
"example = pow_int 10 (-2)" |
|
22317 | 587 |
|
61419 | 588 |
text \<open> |
61438 | 589 |
\<^noindent> This maps to Haskell as follows: |
61419 | 590 |
\<close> |
69660
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
haftmann
parents:
69597
diff
changeset
|
591 |
text %quote \<open> |
39680 | 592 |
@{code_stmts example (Haskell)} |
61419 | 593 |
\<close> |
22317 | 594 |
|
61419 | 595 |
text \<open> |
61438 | 596 |
\<^noindent> The code in SML has explicit dictionary passing: |
61419 | 597 |
\<close> |
69660
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
haftmann
parents:
69597
diff
changeset
|
598 |
text %quote \<open> |
39680 | 599 |
@{code_stmts example (SML)} |
61419 | 600 |
\<close> |
39664
0afaf89ab591
more canonical type setting of type writer code examples
haftmann
parents:
38812
diff
changeset
|
601 |
|
20946 | 602 |
|
61419 | 603 |
text \<open> |
61438 | 604 |
\<^noindent> In Scala, implicits are used as dictionaries: |
61419 | 605 |
\<close> |
69660
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
haftmann
parents:
69597
diff
changeset
|
606 |
text %quote \<open> |
39680 | 607 |
@{code_stmts example (Scala)} |
61419 | 608 |
\<close> |
38812 | 609 |
|
610 |
||
61419 | 611 |
subsection \<open>Inspecting the type class universe\<close> |
29705 | 612 |
|
61419 | 613 |
text \<open> |
38812 | 614 |
To facilitate orientation in complex subclass structures, two |
615 |
diagnostics commands are provided: |
|
29705 | 616 |
|
617 |
\begin{description} |
|
618 |
||
619 |
\item[@{command "print_classes"}] print a list of all classes |
|
620 |
together with associated operations etc. |
|
621 |
||
622 |
\item[@{command "class_deps"}] visualizes the subclass relation |
|
58202 | 623 |
between all classes as a Hasse diagram. An optional first sort argument |
624 |
constrains the set of classes to all subclasses of this sort, |
|
625 |
an optional second sort argument to all superclasses of this sort. |
|
29705 | 626 |
|
627 |
\end{description} |
|
61419 | 628 |
\<close> |
29705 | 629 |
|
20946 | 630 |
end |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48985
diff
changeset
|
631 |