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