doc-src/IsarAdvanced/Classes/Thy/Classes.thy
author haftmann
Tue, 30 Oct 2007 08:45:55 +0100
changeset 25231 1aa9c8f022d0
parent 25200 f1d2e106f2fe
child 25247 7bacd1798fc4
permissions -rw-r--r--
simplified proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
     1
(* $Id$ *)
75b56e51fade initial draft
haftmann
parents:
diff changeset
     2
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
     3
(*<*)
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
     4
theory Classes
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
     5
imports Main Code_Integer
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
     6
begin
75b56e51fade initial draft
haftmann
parents:
diff changeset
     7
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
     8
ML {*
24628
33137422d7fd updated
haftmann
parents: 24348
diff changeset
     9
CodeTarget.code_width := 74;
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    10
*}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    11
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
    12
syntax
75b56e51fade initial draft
haftmann
parents:
diff changeset
    13
  "_alpha" :: "type"  ("\<alpha>")
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
    14
  "_alpha_ofsort" :: "sort \<Rightarrow> type"  ("\<alpha>()\<Colon>_" [0] 1000)
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
    15
75b56e51fade initial draft
haftmann
parents:
diff changeset
    16
parse_ast_translation {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
    17
  let
75b56e51fade initial draft
haftmann
parents:
diff changeset
    18
    fun alpha_ast_tr [] = Syntax.Variable "'a"
75b56e51fade initial draft
haftmann
parents:
diff changeset
    19
      | alpha_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts);
75b56e51fade initial draft
haftmann
parents:
diff changeset
    20
    fun alpha_ofsort_ast_tr [ast] =
75b56e51fade initial draft
haftmann
parents:
diff changeset
    21
      Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'a", ast]
75b56e51fade initial draft
haftmann
parents:
diff changeset
    22
      | alpha_ofsort_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts);
75b56e51fade initial draft
haftmann
parents:
diff changeset
    23
  in [
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
    24
    ("_alpha", alpha_ast_tr), ("_alpha_ofsort", alpha_ofsort_ast_tr)
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
    25
  ] end
75b56e51fade initial draft
haftmann
parents:
diff changeset
    26
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
    27
(*>*)
75b56e51fade initial draft
haftmann
parents:
diff changeset
    28
75b56e51fade initial draft
haftmann
parents:
diff changeset
    29
75b56e51fade initial draft
haftmann
parents:
diff changeset
    30
chapter {* Haskell-style classes with Isabelle/Isar *}
75b56e51fade initial draft
haftmann
parents:
diff changeset
    31
75b56e51fade initial draft
haftmann
parents:
diff changeset
    32
section {* Introduction *}
75b56e51fade initial draft
haftmann
parents:
diff changeset
    33
75b56e51fade initial draft
haftmann
parents:
diff changeset
    34
text {*
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    35
  Type classes were introduces by Wadler and Blott \cite{wadler89how}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    36
  into the Haskell language, to allow for a reasonable implementation
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    37
  of overloading\footnote{throughout this tutorial, we are referring
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    38
  to classical Haskell 1.0 type classes, not considering
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    39
  later additions in expressiveness}.
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    40
  As a canonical example, a polymorphic equality function
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    41
  @{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on different
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    42
  types for @{text "\<alpha>"}, which is achieved by splitting introduction
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    43
  of the @{text eq} function from its overloaded definitions by means
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    44
  of @{text class} and @{text instance} declarations:
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    45
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    46
  \medskip\noindent\hspace*{2ex}@{text "class eq where"}\footnote{syntax here is a kind of isabellized Haskell} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    47
  \hspace*{4ex}@{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
    48
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    49
  \medskip\noindent\hspace*{2ex}@{text "instance nat \<Colon> eq where"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    50
  \hspace*{4ex}@{text "eq 0 0 = True"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    51
  \hspace*{4ex}@{text "eq 0 _ = False"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    52
  \hspace*{4ex}@{text "eq _ 0 = False"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    53
  \hspace*{4ex}@{text "eq (Suc n) (Suc m) = eq n m"}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    54
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    55
  \medskip\noindent\hspace*{2ex}@{text "instance (\<alpha>\<Colon>eq, \<beta>\<Colon>eq) pair \<Colon> eq where"} \\
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    56
  \hspace*{4ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \<and> eq y1 y2"}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
    57
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    58
  \medskip\noindent\hspace*{2ex}@{text "class ord extends eq where"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    59
  \hspace*{4ex}@{text "less_eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    60
  \hspace*{4ex}@{text "less \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    61
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    62
  \medskip\noindent Type variables are annotated with (finitly many) classes;
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    63
  these annotations are assertions that a particular polymorphic type
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    64
  provides definitions for overloaded functions.
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    65
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    66
  Indeed, type classes not only allow for simple overloading
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    67
  but form a generic calculus, an instance of order-sorted
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
    68
  algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
    69
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    70
  From a software enigineering point of view, type classes
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    71
  correspond to interfaces in object-oriented languages like Java;
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    72
  so, it is naturally desirable that type classes do not only
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
    73
  provide functions (class parameters) but also state specifications
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    74
  implementations must obey.  For example, the @{text "class eq"}
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    75
  above could be given the following specification, demanding that
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    76
  @{text "class eq"} is an equivalence relation obeying reflexivity,
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    77
  symmetry and transitivity:
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    78
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    79
  \medskip\noindent\hspace*{2ex}@{text "class eq where"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    80
  \hspace*{4ex}@{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    81
  \hspace*{2ex}@{text "satisfying"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    82
  \hspace*{4ex}@{text "refl: eq x x"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    83
  \hspace*{4ex}@{text "sym: eq x y \<longleftrightarrow> eq x y"} \\
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    84
  \hspace*{4ex}@{text "trans: eq x y \<and> eq y z \<longrightarrow> eq x z"}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    85
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    86
  \medskip\noindent From a theoretic point of view, type classes are leightweight
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
    87
  modules; Haskell type classes may be emulated by
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    88
  SML functors \cite{classes_modules}. 
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    89
  Isabelle/Isar offers a discipline of type classes which brings
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    90
  all those aspects together:
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    91
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    92
  \begin{enumerate}
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
    93
    \item specifying abstract parameters together with
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    94
       corresponding specifications,
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
    95
    \item instantating those abstract parameters by a particular
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    96
       type
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
    97
    \item in connection with a ``less ad-hoc'' approach to overloading,
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
    98
    \item with a direct link to the Isabelle module system
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
    99
      (aka locales \cite{kammueller-locales}).
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   100
  \end{enumerate}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   101
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   102
  \noindent Isar type classes also directly support code generation
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   103
  in a Haskell like fashion.
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   104
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   105
  This tutorial demonstrates common elements of structured specifications
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   106
  and abstract reasoning with type classes by the algebraic hierarchy of
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   107
  semigroups, monoids and groups.  Our background theory is that of
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   108
  Isabelle/HOL \cite{isa-tutorial}, for which some
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   109
  familiarity is assumed.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   110
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   111
  Here we merely present the look-and-feel for end users.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   112
  Internally, those are mapped to more primitive Isabelle concepts.
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   113
  See \cite{Haftmann-Wenzel:2006:classes} for more detail.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   114
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   115
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   116
section {* A simple algebra example \label{sec:example} *}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   117
75b56e51fade initial draft
haftmann
parents:
diff changeset
   118
subsection {* Class definition *}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   119
75b56e51fade initial draft
haftmann
parents:
diff changeset
   120
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   121
  Depending on an arbitrary type @{text "\<alpha>"}, class @{text
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   122
  "semigroup"} introduces a binary operator @{text "\<otimes>"} that is
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   123
  assumed to be associative:
75b56e51fade initial draft
haftmann
parents:
diff changeset
   124
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   125
22473
753123c89d72 explizit "type" superclass
haftmann
parents: 22347
diff changeset
   126
    class semigroup = type +
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   127
      fixes mult :: "\<alpha> \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"    (infixl "\<otimes>" 70)
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   128
      assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   129
75b56e51fade initial draft
haftmann
parents:
diff changeset
   130
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   131
  \noindent This @{text "\<CLASS>"} specification consists of two
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   132
  parts: the \qn{operational} part names the class parameter (@{text
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   133
  "\<FIXES>"}), the \qn{logical} part specifies properties on them
75b56e51fade initial draft
haftmann
parents:
diff changeset
   134
  (@{text "\<ASSUMES>"}).  The local @{text "\<FIXES>"} and @{text
75b56e51fade initial draft
haftmann
parents:
diff changeset
   135
  "\<ASSUMES>"} are lifted to the theory toplevel, yielding the global
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   136
  parameter @{term [source] "mult \<Colon> \<alpha>\<Colon>semigroup \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"} and the
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   137
  global theorem @{text "semigroup.assoc:"}~@{prop [source] "\<And>x y
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   138
  z \<Colon> \<alpha>\<Colon>semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   139
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   140
75b56e51fade initial draft
haftmann
parents:
diff changeset
   141
75b56e51fade initial draft
haftmann
parents:
diff changeset
   142
subsection {* Class instantiation \label{sec:class_inst} *}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   143
75b56e51fade initial draft
haftmann
parents:
diff changeset
   144
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   145
  The concrete type @{text "int"} is made a @{text "semigroup"}
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   146
  instance by providing a suitable definition for the class parameter
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   147
  @{text "mult"} and a proof for the specification of @{text "assoc"}.
75b56e51fade initial draft
haftmann
parents:
diff changeset
   148
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   149
75b56e51fade initial draft
haftmann
parents:
diff changeset
   150
    instance int :: semigroup
24628
33137422d7fd updated
haftmann
parents: 24348
diff changeset
   151
      mult_int_def: "i \<otimes> j \<equiv> i + j"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   152
    proof
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   153
      fix i j k :: int have "(i + j) + k = i + (j + k)" by simp
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   154
      then show "(i \<otimes> j) \<otimes> k = i \<otimes> (j \<otimes> k)" unfolding mult_int_def .
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   155
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   156
75b56e51fade initial draft
haftmann
parents:
diff changeset
   157
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   158
  \noindent From now on, the type-checker will consider @{text "int"}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   159
  as a @{text "semigroup"} automatically, i.e.\ any general results
75b56e51fade initial draft
haftmann
parents:
diff changeset
   160
  are immediately available on concrete instances.
75b56e51fade initial draft
haftmann
parents:
diff changeset
   161
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   162
  Note that the first proof step is the @{text default} method,
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   163
  which for instantiation proofs maps to the @{text intro_classes} method.
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   164
  This boils down an instantiation judgement to the relevant primitive
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   165
  proof goals and should conveniently always be the first method applied
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   166
  in an instantiation proof.
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   167
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   168
  \medskip Another instance of @{text "semigroup"} are the natural numbers:
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   169
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   170
75b56e51fade initial draft
haftmann
parents:
diff changeset
   171
    instance nat :: semigroup
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   172
      mult_nat_def: "m \<otimes> n \<equiv> m + n"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   173
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   174
      fix m n q :: nat 
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   175
      show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)" unfolding mult_nat_def by simp
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   176
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   177
75b56e51fade initial draft
haftmann
parents:
diff changeset
   178
text {*
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   179
  \noindent Also @{text "list"}s form a semigroup with @{const "op @"} as
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   180
  parameter:
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   181
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   182
75b56e51fade initial draft
haftmann
parents:
diff changeset
   183
    instance list :: (type) semigroup
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   184
      mult_list_def: "xs \<otimes> ys \<equiv> xs @ ys"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   185
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   186
      fix xs ys zs :: "\<alpha> list"
75b56e51fade initial draft
haftmann
parents:
diff changeset
   187
      show "xs \<otimes> ys \<otimes> zs = xs \<otimes> (ys \<otimes> zs)"
75b56e51fade initial draft
haftmann
parents:
diff changeset
   188
      proof -
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   189
        from mult_list_def have "\<And>xs ys\<Colon>\<alpha> list. xs \<otimes> ys \<equiv> xs @ ys" .
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   190
        thus ?thesis by simp
75b56e51fade initial draft
haftmann
parents:
diff changeset
   191
      qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   192
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   193
75b56e51fade initial draft
haftmann
parents:
diff changeset
   194
75b56e51fade initial draft
haftmann
parents:
diff changeset
   195
subsection {* Subclasses *}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   196
75b56e51fade initial draft
haftmann
parents:
diff changeset
   197
text {*
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   198
  We define a subclass @{text "monoidl"} (a semigroup with a left-hand neutral)
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   199
  by extending @{text "semigroup"}
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   200
  with one additional parameter @{text "neutral"} together
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   201
  with its property:
75b56e51fade initial draft
haftmann
parents:
diff changeset
   202
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   203
75b56e51fade initial draft
haftmann
parents:
diff changeset
   204
    class monoidl = semigroup +
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   205
      fixes neutral :: "\<alpha>" ("\<one>")
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   206
      assumes neutl: "\<one> \<otimes> x = x"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   207
75b56e51fade initial draft
haftmann
parents:
diff changeset
   208
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   209
  \noindent Again, we make some instances, by
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   210
  providing suitable parameter definitions and proofs for the
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   211
  additional specifications.
75b56e51fade initial draft
haftmann
parents:
diff changeset
   212
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   213
75b56e51fade initial draft
haftmann
parents:
diff changeset
   214
    instance nat :: monoidl
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   215
      neutral_nat_def: "\<one> \<equiv> 0"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   216
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   217
      fix n :: nat
75b56e51fade initial draft
haftmann
parents:
diff changeset
   218
      show "\<one> \<otimes> n = n" unfolding neutral_nat_def mult_nat_def by simp
75b56e51fade initial draft
haftmann
parents:
diff changeset
   219
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   220
75b56e51fade initial draft
haftmann
parents:
diff changeset
   221
    instance int :: monoidl
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   222
      neutral_int_def: "\<one> \<equiv> 0"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   223
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   224
      fix k :: int
75b56e51fade initial draft
haftmann
parents:
diff changeset
   225
      show "\<one> \<otimes> k = k" unfolding neutral_int_def mult_int_def by simp
75b56e51fade initial draft
haftmann
parents:
diff changeset
   226
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   227
75b56e51fade initial draft
haftmann
parents:
diff changeset
   228
    instance list :: (type) monoidl
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   229
      neutral_list_def: "\<one> \<equiv> []"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   230
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   231
      fix xs :: "\<alpha> list"
75b56e51fade initial draft
haftmann
parents:
diff changeset
   232
      show "\<one> \<otimes> xs = xs"
75b56e51fade initial draft
haftmann
parents:
diff changeset
   233
      proof -
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   234
	from mult_list_def have "\<And>xs ys\<Colon>\<alpha> list. xs \<otimes> ys \<equiv> xs @ ys" .
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   235
	moreover from mult_list_def neutral_list_def have "\<one> \<equiv> []\<Colon>\<alpha> list" by simp
75b56e51fade initial draft
haftmann
parents:
diff changeset
   236
	ultimately show ?thesis by simp
75b56e51fade initial draft
haftmann
parents:
diff changeset
   237
      qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   238
    qed  
75b56e51fade initial draft
haftmann
parents:
diff changeset
   239
75b56e51fade initial draft
haftmann
parents:
diff changeset
   240
text {*
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   241
  \noindent Fully-fledged monoids are modelled by another subclass
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   242
  which does not add new parameters but tightens the specification:
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   243
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   244
75b56e51fade initial draft
haftmann
parents:
diff changeset
   245
    class monoid = monoidl +
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   246
      assumes neutr: "x \<otimes> \<one> = x"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   247
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   248
text {*
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   249
  \noindent Instantiations may also be given simultaneously for different
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   250
  type constructors:
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   251
*}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   252
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   253
    instance nat :: monoid and int :: monoid and list :: (type) monoid
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   254
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   255
      fix n :: nat
75b56e51fade initial draft
haftmann
parents:
diff changeset
   256
      show "n \<otimes> \<one> = n" unfolding neutral_nat_def mult_nat_def by simp
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   257
    next
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   258
      fix k :: int
75b56e51fade initial draft
haftmann
parents:
diff changeset
   259
      show "k \<otimes> \<one> = k" unfolding neutral_int_def mult_int_def by simp
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   260
    next
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   261
      fix xs :: "\<alpha> list"
75b56e51fade initial draft
haftmann
parents:
diff changeset
   262
      show "xs \<otimes> \<one> = xs"
75b56e51fade initial draft
haftmann
parents:
diff changeset
   263
      proof -
75b56e51fade initial draft
haftmann
parents:
diff changeset
   264
	from mult_list_def have "\<And>xs ys\<Colon>\<alpha> list. xs \<otimes> ys \<equiv> xs @ ys" .
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   265
	moreover from mult_list_def neutral_list_def have "\<one> \<equiv> []\<Colon>\<alpha> list" by simp
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   266
	ultimately show ?thesis by simp
75b56e51fade initial draft
haftmann
parents:
diff changeset
   267
      qed
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   268
    qed
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   269
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   270
text {*
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   271
  \noindent To finish our small algebra example, we add a @{text "group"} class
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   272
  with a corresponding instance:
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   273
*}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   274
75b56e51fade initial draft
haftmann
parents:
diff changeset
   275
    class group = monoidl +
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   276
      fixes inverse :: "\<alpha> \<Rightarrow> \<alpha>"    ("(_\<div>)" [1000] 999)
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   277
      assumes invl: "x\<div> \<otimes> x = \<one>"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   278
75b56e51fade initial draft
haftmann
parents:
diff changeset
   279
    instance int :: group
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   280
      inverse_int_def: "i\<div> \<equiv> - i"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   281
    proof
75b56e51fade initial draft
haftmann
parents:
diff changeset
   282
      fix i :: int
75b56e51fade initial draft
haftmann
parents:
diff changeset
   283
      have "-i + i = 0" by simp
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   284
      then show "i\<div> \<otimes> i = \<one>"
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   285
      unfolding mult_int_def and neutral_int_def and inverse_int_def .
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   286
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   287
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   288
section {* Type classes as locales *}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   289
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   290
subsection {* A look behind the scene *}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   291
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   292
text {*
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   293
  The example above gives an impression how Isar type classes work
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   294
  in practice.  As stated in the introduction, classes also provide
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   295
  a link to Isar's locale system.  Indeed, the logical core of a class
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   296
  is nothing else than a locale:
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   297
*}
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   298
22473
753123c89d72 explizit "type" superclass
haftmann
parents: 22347
diff changeset
   299
class idem = type +
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   300
  fixes f :: "\<alpha> \<Rightarrow> \<alpha>"
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   301
  assumes idem: "f (f x) = f x"
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   302
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   303
text {*
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   304
  \noindent essentially introduces the locale
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   305
*}
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   306
(*<*) setup {* Sign.add_path "foo" *} (*>*)
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   307
locale idem =
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   308
  fixes f :: "\<alpha> \<Rightarrow> \<alpha>"
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   309
  assumes idem: "f (f x) = f x"
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   310
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   311
text {* \noindent together with corresponding constant(s): *}
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   312
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   313
consts f :: "\<alpha> \<Rightarrow> \<alpha>"
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   314
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   315
text {*
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   316
  \noindent The connection to the type system is done by means
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   317
  of a primitive axclass
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   318
*}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   319
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   320
axclass idem < type
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   321
  idem: "f (f x) = f x"
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   322
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   323
text {* \noindent together with a corresponding interpretation: *}
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   324
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   325
interpretation idem_class:
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   326
  idem ["f \<Colon> ('a\<Colon>idem) \<Rightarrow> \<alpha>"]
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   327
by unfold_locales (rule idem)
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   328
(*<*) setup {* Sign.parent_path *} (*>*)
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   329
text {*
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   330
  This give you at hand the full power of the Isabelle module system;
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   331
  conclusions in locale @{text idem} are implicitly propagated
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   332
  to class @{text idem}.
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   333
*}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   334
75b56e51fade initial draft
haftmann
parents:
diff changeset
   335
subsection {* Abstract reasoning *}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   336
75b56e51fade initial draft
haftmann
parents:
diff changeset
   337
text {*
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   338
  Isabelle locales enable reasoning at a general level, while results
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   339
  are implicitly transferred to all instances.  For example, we can
75b56e51fade initial draft
haftmann
parents:
diff changeset
   340
  now establish the @{text "left_cancel"} lemma for groups, which
75b56e51fade initial draft
haftmann
parents:
diff changeset
   341
  states that the function @{text "(x \<circ>)"} is injective:
75b56e51fade initial draft
haftmann
parents:
diff changeset
   342
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   343
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   344
    lemma (in group) left_cancel: "x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   345
    proof
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   346
    assume "x \<otimes> y = x \<otimes> z"
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   347
      then have "x\<div> \<otimes> (x \<otimes> y) = x\<div> \<otimes> (x \<otimes> z)" by simp
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   348
      then have "(x\<div> \<otimes> x) \<otimes> y = (x\<div> \<otimes> x) \<otimes> z" using assoc by simp
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   349
      then show "y = z" using neutl and invl by simp
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   350
    next
75b56e51fade initial draft
haftmann
parents:
diff changeset
   351
    assume "y = z"
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   352
      then show "x \<otimes> y = x \<otimes> z" by simp
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   353
    qed
75b56e51fade initial draft
haftmann
parents:
diff changeset
   354
75b56e51fade initial draft
haftmann
parents:
diff changeset
   355
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   356
  \noindent Here the \qt{@{text "\<IN> group"}} target specification
75b56e51fade initial draft
haftmann
parents:
diff changeset
   357
  indicates that the result is recorded within that context for later
75b56e51fade initial draft
haftmann
parents:
diff changeset
   358
  use.  This local theorem is also lifted to the global one @{text
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   359
  "group.left_cancel:"} @{prop [source] "\<And>x y z \<Colon> \<alpha>\<Colon>group. x \<otimes> y = x \<otimes>
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   360
  z \<longleftrightarrow> y = z"}.  Since type @{text "int"} has been made an instance of
75b56e51fade initial draft
haftmann
parents:
diff changeset
   361
  @{text "group"} before, we may refer to that fact as well: @{prop
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   362
  [source] "\<And>x y z \<Colon> int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   363
*}
75b56e51fade initial draft
haftmann
parents:
diff changeset
   364
75b56e51fade initial draft
haftmann
parents:
diff changeset
   365
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   366
subsection {* Derived definitions *}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   367
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   368
text {*
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   369
  Isabelle locales support a concept of local definitions
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   370
  in locales:
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   371
*}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   372
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   373
    fun (in monoid)
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   374
      pow_nat :: "nat \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   375
      "pow_nat 0 x = \<one>"
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   376
      | "pow_nat (Suc n) x = x \<otimes> pow_nat n x"
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   377
75b56e51fade initial draft
haftmann
parents:
diff changeset
   378
text {*
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   379
  \noindent If the locale @{text group} is also a class, this local
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   380
  definition is propagated onto a global definition of
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   381
  @{term [source] "pow_nat \<Colon> nat \<Rightarrow> \<alpha>\<Colon>monoid \<Rightarrow> \<alpha>\<Colon>monoid"}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   382
  with corresponding theorems
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   383
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   384
  @{thm pow_nat.simps [no_vars]}.
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   385
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   386
  \noindent As you can see from this example, for local
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   387
  definitions you may use any specification tool
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   388
  which works together with locales (e.g. \cite{krauss2006}).
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   389
*}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   390
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   391
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   392
subsection {* Additional subclass relations *}
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   393
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   394
text {*
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   395
  Any @{text "group"} is also a @{text "monoid"};  this
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   396
  can be made explicit by claiming an additional subclass relation,
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   397
  together with a proof of the logical difference:
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   398
*}
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   399
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   400
    subclass (in group) monoid
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   401
    proof unfold_locales
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   402
      fix x
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   403
      from invl have "x\<div> \<otimes> x = \<one>" by simp
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   404
      with assoc [symmetric] neutl invl have "x\<div> \<otimes> (x \<otimes> \<one>) = x\<div> \<otimes> x" by simp
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   405
      with left_cancel show "x \<otimes> \<one> = x" by simp
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   406
    qed
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   407
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   408
text {*
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   409
  \noindent The logical proof is carried out on the locale level
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   410
  and thus conveniently is opened using the @{text unfold_locales}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   411
  method which only leaves the logical differences still
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   412
  open to proof to the user.  Afterwards it is propagated
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   413
  to the type system, making @{text group} an instance of
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   414
  @{text monoid}.  For illustration, a derived definition
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   415
  in @{text group} which uses @{text pow_nat}:
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   416
*}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   417
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   418
    definition (in group)
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   419
      pow_int :: "int \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   420
      "pow_int k x = (if k >= 0
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   421
        then pow_nat (nat k) x
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   422
        else (pow_nat (nat (- k)) x)\<div>)"
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   423
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   424
text {*
25200
f1d2e106f2fe adjusted
haftmann
parents: 24991
diff changeset
   425
  \noindent   yields the global definition of
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   426
  @{term [source] "pow_int \<Colon> int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group"}
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   427
  with the corresponding theorem @{thm pow_int_def [no_vars]}.
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   428
*}
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   429
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   430
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   431
section {* Further issues *}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   432
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   433
subsection {* Code generation *}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   434
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   435
text {*
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   436
  Turning back to the first motivation for type classes,
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   437
  namely overloading, it is obvious that overloading
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   438
  stemming from @{text "\<CLASS>"} and @{text "\<INSTANCE>"}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   439
  statements naturally maps to Haskell type classes.
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   440
  The code generator framework \cite{isabelle-codegen} 
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   441
  takes this into account.  Concerning target languages
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   442
  lacking type classes (e.g.~SML), type classes
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   443
  are implemented by explicit dictionary construction.
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   444
  For example, lets go back to the power function:
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   445
*}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   446
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   447
(*    fun
23956
48494ccfabaf updated
haftmann
parents: 22845
diff changeset
   448
      pow_nat :: "nat \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group" where
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   449
      "pow_nat 0 x = \<one>"
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   450
      | "pow_nat (Suc n) x = x \<otimes> pow_nat n x"
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   451
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   452
    definition
22347
ddbf185a3be0 continued
haftmann
parents: 22317
diff changeset
   453
      pow_int :: "int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group" where
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   454
      "pow_int k x = (if k >= 0
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   455
        then pow_nat (nat k) x
24991
c6f5cc939c29 added subclass command
haftmann
parents: 24628
diff changeset
   456
        else (pow_nat (nat (- k)) x)\<div>)"*)
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   457
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   458
    definition
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   459
      example :: int where
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   460
      "example = pow_int 10 (-2)"
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   461
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   462
text {*
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   463
  \noindent This maps to Haskell as:
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   464
*}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   465
24628
33137422d7fd updated
haftmann
parents: 24348
diff changeset
   466
export_code example in Haskell module_name Classes file "code_examples/"
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   467
  (* NOTE: you may use Haskell only once in this document, otherwise
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   468
  you have to work in distinct subdirectories *)
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   469
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   470
text {*
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   471
  \lsthaskell{Thy/code_examples/Classes.hs}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   472
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   473
  \noindent The whole code in SML with explicit dictionary passing:
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   474
*}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   475
24628
33137422d7fd updated
haftmann
parents: 24348
diff changeset
   476
export_code example (*<*)in SML module_name Classes(*>*)in SML module_name Classes file "code_examples/classes.ML"
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   477
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   478
text {*
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   479
  \lstsml{Thy/code_examples/classes.ML}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   480
*}
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   481
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   482
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   483
(* subsection {* Different syntax for same specifications *}
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   484
75b56e51fade initial draft
haftmann
parents:
diff changeset
   485
text {*
75b56e51fade initial draft
haftmann
parents:
diff changeset
   486
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   487
subsection {* Syntactic classes *}
22317
b550d2c6ca90 continued class tutorial
haftmann
parents: 20946
diff changeset
   488
20946
75b56e51fade initial draft
haftmann
parents:
diff changeset
   489
*} *)
75b56e51fade initial draft
haftmann
parents:
diff changeset
   490
75b56e51fade initial draft
haftmann
parents:
diff changeset
   491
end