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