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