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