doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy
author haftmann
Sun, 05 Nov 2006 09:36:25 +0100
changeset 21178 c3618fc6a6f7
parent 21147 737a94f047e3
child 21189 5435ccdb4ea1
permissions -rw-r--r--
added gfx
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     1
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     2
(* $Id$ *)
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     3
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
     4
(*<*)
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     5
theory Codegen
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     6
imports Main
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
     7
uses "../../../IsarImplementation/Thy/setup.ML"
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     8
begin
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     9
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    10
ML {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    11
CodegenSerializer.sml_code_width := 74;
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    12
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    13
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    14
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    15
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    16
chapter {* Code generation from Isabelle theories *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    17
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    18
section {* Introduction *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    19
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    20
subsection {* Motivation *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    21
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    22
text {*
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    23
  Executing formal specifications as programs is a well-established
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    24
  topic in the theorem proving community.  With increasing
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    25
  application of theorem proving systems in the area of
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    26
  software development and verification, its relevance manifests
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    27
  for running test cases and rapid prototyping.  In logical
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    28
  calculi like constructive type theory,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    29
  a notion of executability is implicit due to the nature
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    30
  of the calculus.  In contrast, specifications in Isabelle/HOL
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    31
  can be highly non-executable.  In order to bridge
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    32
  the gap between logic and executable specifications,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    33
  an explicit non-trivial transformation has to be applied:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    34
  code generation.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    35
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    36
  This tutorial introduces a generic code generator for the
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    37
  Isabelle system \cite{isa-tutorial}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    38
  Generic in the sense that the
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    39
  \qn{target language} for which code shall ultimately be
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    40
  generated is not fixed but may be an arbitrary state-of-the-art
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    41
  functional programming language (currently, the implementation
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    42
  supports SML \cite{web:sml} and Haskell \cite{web:haskell}).
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    43
  We aim to provide a
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    44
  versatile environment
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    45
  suitable for software development and verification,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    46
  structuring the process
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    47
  of code generation into a small set of orthogonal principles
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    48
  while achieving a big coverage of application areas
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    49
  with maximum flexibility.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    50
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    51
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    52
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    53
subsection {* Overview *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    54
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    55
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    56
  The code generator aims to be usable with no further ado
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    57
  in most cases while allowing for detailed customization.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    58
  This manifests in the structure of this tutorial: this introduction
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    59
  continues with a short introduction of concepts.  Section
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
    60
  \secref{sec:basics} explains how to use the framework naively,
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    61
  presuming a reasonable default setup.  Then, section
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    62
  \secref{sec:advanced} deals with advanced topics,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    63
  introducing further aspects of the code generator framework
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    64
  in a motivation-driven manner.  Last, section \secref{sec:ml}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    65
  introduces the framework's internal programming interfaces.
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    66
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    67
  \begin{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    68
    Ultimately, the code generator which this tutorial deals with
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    69
    is supposed to replace the already established code generator
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    70
    by Stefan Berghofer \cite{Berghofer-Nipkow:2002}.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    71
    So, for the moment, there are two distinct code generators
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    72
    in Isabelle.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    73
    Also note that while the framework itself is largely
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    74
    object-logic independent, only HOL provides a reasonable
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    75
    framework setup.    
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    76
  \end{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    77
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    78
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    79
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    80
subsection {* Code generation process *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    81
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    82
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    83
  The code generator employs a notion of executability
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    84
  for three foundational executable ingredients known
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    85
  from functional programming:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    86
  \emph{function equations}, \emph{datatypes}, and
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    87
  \emph{type classes}. A function equation as a first approximation
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    88
  is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    89
  (an equation headed by a constant @{text f} with arguments
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    90
  @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right hand side @{text t}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    91
  Code generation aims to turn function equations
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    92
  into a functional program by running through
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    93
  a process (see figure \ref{fig:process}):
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    94
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    95
  \begin{itemize}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    96
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    97
    \item Out of the vast collection of theorems proven in a
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    98
      \qn{theory}, a reasonable subset modeling
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    99
      function equations is \qn{selected}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   100
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   101
    \item On those selected theorems, certain
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   102
      transformations are carried out
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   103
      (\qn{preprocessing}).  Their purpose is to turn theorems
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   104
      representing non- or badly executable
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   105
      specifications into equivalent but executable counterparts.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   106
      The result is a structured collection of \qn{code theorems}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   107
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   108
    \item These \qn{code theorems} then are extracted
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   109
      into an Haskell-like intermediate
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   110
      language.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   111
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   112
    \item Finally, out of the intermediate language the final
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   113
      code in the desired \qn{target language} is \qn{serialized}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   114
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   115
  \end{itemize}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   116
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   117
  \begin{figure}[h]
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   118
  \centering
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   119
  \includegraphics[width=0.3\textwidth]{codegen_process}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   120
  \caption{code generator -- processing overview}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   121
  \label{fig:process}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   122
  \end{figure}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   123
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   124
  From these steps, only the two last are carried out
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   125
  outside the logic; by keeping this layer as
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   126
  thin as possible, the amount of code to trust is
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   127
  kept to a minimum.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   128
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   129
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   130
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   131
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   132
section {* Basics \label{sec:basics} *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   133
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   134
subsection {* Invoking the code generator *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   135
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   136
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   137
  Thanks to a reasonable setup of the HOL theories, in
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   138
  most cases code generation proceeds without further ado:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   139
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   140
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   141
consts
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   142
  fac :: "nat \<Rightarrow> nat"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   143
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   144
primrec
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   145
  "fac 0 = 1"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   146
  "fac (Suc n) = Suc n * fac n"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   147
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   148
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   149
  This executable specification is now turned to SML code:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   150
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   151
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   152
code_gen fac (SML "examples/fac.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   153
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   154
text {*
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   155
  The \isasymCODEGEN command takes a space-separated list of
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   156
  constants together with \qn{serialization directives}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   157
  in parentheses. These start with a \qn{target language}
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   158
  identifier, followed by arguments, their semantics
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   159
  depending on the target. In the SML case, a filename
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   160
  is given where to write the generated code to.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   161
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   162
  Internally, the function equations for all selected
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   163
  constants are taken, including any transitively required
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   164
  constants, datatypes and classes, resulting in the following
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   165
  code:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   166
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   167
  \lstsml{Thy/examples/fac.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   168
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   169
  The code generator will complain when a required
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   170
  ingredient does not provide a executable counterpart.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   171
  This is the case if an involved type is not a datatype:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   172
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   173
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   174
(*<*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   175
setup {* Sign.add_path "foo" *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   176
(*>*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   177
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   178
typedecl 'a foo
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   179
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   180
definition
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   181
  bar :: "'a foo \<Rightarrow> 'a \<Rightarrow> 'a"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   182
  "bar x y = y"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   183
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   184
(*<*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   185
hide type foo
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   186
hide const bar
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   187
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   188
setup {* Sign.parent_path *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   189
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   190
datatype 'a foo = Foo
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   191
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   192
definition
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   193
  bar :: "'a foo \<Rightarrow> 'a \<Rightarrow> 'a"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   194
  "bar x y = y"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   195
(*>*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   196
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   197
code_gen bar (SML "examples/fail_type.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   198
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   199
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   200
  \noindent will result in an error. Likewise, generating code
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   201
  for constants not yielding
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   202
  a function equation will fail, e.g.~the Hilbert choice
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   203
  operation @{text "SOME"}:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   204
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   205
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   206
(*<*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   207
setup {* Sign.add_path "foo" *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   208
(*>*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   209
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   210
definition
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   211
  pick_some :: "'a list \<Rightarrow> 'a"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   212
  "pick_some xs = (SOME x. x \<in> set xs)"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   213
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   214
(*<*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   215
hide const pick_some
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   216
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   217
setup {* Sign.parent_path *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   218
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   219
definition
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   220
  pick_some :: "'a list \<Rightarrow> 'a"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   221
  "pick_some = hd"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   222
(*>*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   223
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   224
code_gen pick_some (SML "examples/fail_const.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   225
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   226
subsection {* Theorem selection *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   227
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   228
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   229
  The list of all function equations in a theory may be inspected
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   230
  using the \isasymPRINTCODETHMS command:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   231
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   232
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   233
print_codethms
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   234
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   235
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   236
  \noindent which displays a table of constant with corresponding
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   237
  function equations (the additional stuff displayed
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   238
  shall not bother us for the moment). If this table does
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   239
  not provide at least one function
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   240
  equation, the table of primitive definitions is searched
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   241
  whether it provides one.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   242
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   243
  The typical HOL tools are already set up in a way that
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   244
  function definitions introduced by \isasymFUN, \isasymFUNCTION,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   245
  \isasymPRIMREC, \isasymRECDEF are implicitly propagated
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   246
  to this function equation table. Specific theorems may be
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   247
  selected using an attribute: \emph{code func}. As example,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   248
  a weight selector function:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   249
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   250
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   251
consts
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   252
  pick :: "(nat \<times> 'a) list \<Rightarrow> nat \<Rightarrow> 'a"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   253
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   254
primrec
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   255
  "pick (x#xs) n = (let (k, v) = x in
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   256
    if n < k then v else pick xs (n - k))"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   257
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   258
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   259
  We want to eliminate the explicit destruction
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   260
  of @{term x} to @{term "(k, v)"}:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   261
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   262
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   263
lemma [code func]:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   264
  "pick ((k, v)#xs) n = (if n < k then v else pick xs (n - k))"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   265
  by simp
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   266
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   267
code_gen pick (SML "examples/pick1.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   268
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   269
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   270
  This theorem is now added to the function equation table:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   271
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   272
  \lstsml{Thy/examples/pick1.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   273
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   274
  It might be convenient to remove the pointless original
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   275
  equation, using the \emph{nofunc} attribute:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   276
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   277
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   278
lemmas [code nofunc] = pick.simps 
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   279
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   280
code_gen pick (SML "examples/pick2.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   281
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   282
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   283
  \lstsml{Thy/examples/pick2.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   284
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   285
  Syntactic redundancies are implicitly dropped. For example,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   286
  using a modified version of the @{const fac} function
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   287
  as function equation, the then redundant (since
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   288
  syntactically subsumed) original function equations
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   289
  are dropped, resulting in a warning:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   290
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   291
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   292
lemma [code func]:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   293
  "fac n = (case n of 0 \<Rightarrow> 1 | Suc m \<Rightarrow> n * fac m)"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   294
  by (cases n) simp_all
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   295
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   296
code_gen fac (SML "examples/fac_case.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   297
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   298
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   299
  \lstsml{Thy/examples/fac_case.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   300
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   301
  \begin{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   302
    Some statements in this section have to be treated with some
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   303
    caution. First, since the HOL function package is still
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   304
    under development, its setup with respect to code generation
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   305
    may differ from what is presumed here.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   306
    Further, the attributes \emph{code} and \emph{code del}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   307
    associated with the existing code generator also apply to
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   308
    the new one: \emph{code} implies \emph{code func},
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   309
    and \emph{code del} implies \emph{code nofunc}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   310
  \end{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   311
*}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   312
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   313
subsection {* Type classes *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   314
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   315
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   316
  Type classes enter the game via the Isar class package.
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   317
  For a short introduction how to use it, see \cite{isabelle-classes};
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   318
  here we just illustrate its impact on code generation.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   319
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   320
  In a target language, type classes may be represented
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   321
  natively (as in the case of Haskell). For languages
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   322
  like SML, they are implemented using \emph{dictionaries}.
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   323
  Our following example specifies a class \qt{null},
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   324
  assigning to each of its inhabitants a \qt{null} value:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   325
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   326
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   327
class null =
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   328
  fixes null :: 'a
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   329
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   330
consts
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   331
  head :: "'a\<Colon>null list \<Rightarrow> 'a"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   332
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   333
primrec
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   334
  "head [] = null"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   335
  "head (x#xs) = x"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   336
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   337
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   338
  We provide some instances for our @{text null}:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   339
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   340
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   341
instance option :: (type) null
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   342
  "null \<equiv> None" ..
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   343
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   344
instance list :: (type) null
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   345
  "null \<equiv> []" ..
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   346
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   347
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   348
  Constructing a dummy example:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   349
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   350
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   351
definition
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   352
  "dummy = head [Some (Suc 0), None]"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   353
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   354
text {*
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   355
  Type classes offer a suitable occasion to introduce
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   356
  the Haskell serializer.  Its usage is almost the same
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   357
  as SML, but, in accordance with conventions
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   358
  some Haskell systems enforce, each module ends
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   359
  up in a single file. The module hierarchy is reflected in
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   360
  the file system, with root given by the user.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   361
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   362
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   363
code_gen dummy (Haskell "examples/")
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   364
  (* NOTE: you may use Haskell only once in this document, otherwise
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   365
  you have to work in distinct subdirectories *)
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   366
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   367
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   368
  \lsthaskell{Thy/examples/Codegen.hs}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   369
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   370
  (we have left out all other modules).
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   371
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   372
  The whole code in SML with explicit dictionary passing:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   373
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   374
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   375
code_gen dummy (SML "examples/class.ML")
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   376
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   377
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   378
  \lstsml{Thy/examples/class.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   379
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   380
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   381
subsection {* Incremental code generation *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   382
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   383
text {*
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   384
  Code generation is \emph{incremental}: theorems
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   385
  and abstract intermediate code are cached and extended on demand.
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   386
  The cache may be partially or fully dropped if the underlying
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   387
  executable content of the theory changes.
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   388
  Implementation of caching is supposed to transparently
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   389
  hid away the details from the user.  Anyway, caching
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   390
  reaches the surface by using a slightly more general form
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   391
  of the \isasymCODEGEN: either the list of constants or the
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   392
  list of serialization expressions may be dropped.  If no
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   393
  serialization expressions are given, only abstract code
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   394
  is generated and cached; if no constants are given, the
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   395
  current cache is serialized.
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   396
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   397
  For explorative reasons, an extended version of the
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   398
  \isasymCODEGEN command may prove useful:
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   399
*}
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   400
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   401
print_codethms ()
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   402
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   403
text {*
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   404
  \noindent print all cached function equations (i.e.~\emph{after}
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   405
  any applied transformation. Inside the brackets a
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   406
  list of constants may be given; their function
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   407
  equations are added to the cache if not already present.
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   408
*}
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   409
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   410
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   411
section {* Recipes and advanced topics \label{sec:advanced} *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   412
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   413
text {*
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   414
  In this tutorial, we do not attempt to give an exhaustive
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   415
  description of the code generator framework; instead,
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   416
  we cast a light on advanced topics by introducing
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   417
  them together with practically motivated examples.  Concerning
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   418
  further reading, see
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   419
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   420
  \begin{itemize}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   421
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   422
  \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   423
    for exhaustive syntax diagrams.
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   424
  \item or \fixme{ref} which deals with foundational issues
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   425
    of the code generator framework.
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   426
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   427
  \end{itemize}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   428
*}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   429
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   430
subsection {* Library theories *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   431
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   432
text {*
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   433
  The HOL \emph{Main} theory already provides a code generator setup
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   434
  which should be suitable for most applications. Common extensions
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   435
  and modifications are available by certain theories of the HOL
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   436
  library; beside being useful in applications, they may serve
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   437
  as a tutorial for customizing the code generator setup.
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   438
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   439
  \begin{description}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   440
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   441
    \item[ExecutableSet] allows to generate code
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   442
       for finite sets using lists.
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   443
    \item[ExecutableRat] implements rational
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   444
       numbers as triples @{text "(sign, enumerator, denominator)"}.
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   445
    \item[EfficientNat] implements natural numbers by integers,
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   446
       which in general will result in higher efficency; pattern
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   447
       matching with @{const "0\<Colon>nat"} / @{const "Suc"}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   448
       is eliminated. \label{eff_nat}
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   449
    \item[MLString] provides an additional datatype @{text "mlstring"};
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   450
       in the HOL default setup, strings in HOL are mapped to list
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   451
       of chars in SML; values of type @{text "mlstring"} are
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   452
       mapped to strings in SML.
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   453
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   454
  \end{description}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   455
*}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   456
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   457
subsection {* Preprocessing *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   458
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   459
text {*
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   460
  Before selected function theorems are turned into abstract
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   461
  code, a chain of definitional transformation steps is carried
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   462
  out: \emph{preprocessing}. There are three possibilities
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   463
  to customize preprocessing: \emph{inline theorems},
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   464
  \emph{inline procedures} and \emph{generic preprocessors}.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   465
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   466
  \emph{Inline theorems} are rewriting rules applied to each
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   467
  function equation.  Due to the interpretation of theorems
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   468
  of function equations, rewrites are applied to the right
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   469
  hand side and the arguments of the left hand side of an
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   470
  equation, but never to the constant heading the left hand side.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   471
  Inline theorems may be declared an undeclared using the
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   472
  \emph{code inline} or \emph{code noinline} attribute respectively.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   473
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   474
  Some common applications:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   475
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   476
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   477
text_raw {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   478
  \begin{itemize}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   479
     \item replacing non-executable constructs by executable ones: \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   480
*}     
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   481
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   482
lemma [code inline]:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   483
  "x \<in> set xs \<longleftrightarrow> x mem xs" by (induct xs) simp_all
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   484
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   485
text_raw {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   486
     \item eliminating superfluous constants: \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   487
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   488
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   489
lemma [code inline]:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   490
  "1 = Suc 0" by simp
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   491
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   492
text_raw {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   493
     \item replacing executable but inconvenient constructs: \\
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   494
*}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   495
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   496
lemma [code inline]:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   497
  "xs = [] \<longleftrightarrow> List.null xs" by (induct xs) simp_all
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   498
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   499
text_raw {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   500
  \end{itemize}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   501
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   502
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   503
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   504
  The current set of inline theorems may be inspected using
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   505
  the \isasymPRINTCODETHMS command.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   506
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   507
  \emph{Inline procedures} are a generalized version of inline
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   508
  theorems written in ML -- rewrite rules are generated dependent
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   509
  on the function theorems for a certain function.  One
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   510
  application is the implicit expanding of @{typ nat} numerals
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   511
  to @{const "0\<Colon>nat"} / @{const Suc} representation.  See further
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   512
  \secref{sec:ml}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   513
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   514
  \emph{Generic preprocessors} provide a most general interface,
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   515
  transforming a list of function theorems to another
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   516
  list of function theorems, provided that neither the heading
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   517
  constant nor its type change.  The @{const "0\<Colon>nat"} / @{const Suc}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   518
  pattern elimination implemented in \secref{eff_nat} uses this
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   519
  interface.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   520
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   521
  \begin{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   522
    The order in which single preprocessing steps are carried
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   523
    out currently is not specified; in particular, preprocessing
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   524
    is \emph{no} fix point process.  Keep this in mind when
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   525
    setting up the preprocessor.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   526
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   527
    Further, the attribute \emph{code unfold}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   528
    associated with the existing code generator also applies to
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   529
    the new one: \emph{code unfold} implies \emph{code inline}.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   530
  \end{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   531
*}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   532
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   533
subsection {* Customizing serialization  *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   534
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   535
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   536
  Consider the following function and its corresponding
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   537
  SML code:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   538
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   539
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   540
fun
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   541
  in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   542
  "in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   543
termination by (auto_term "{}")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   544
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   545
declare in_interval.simps [code func]
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   546
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   547
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   548
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   549
code_type bool
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   550
  (SML)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   551
code_const True and False and "op \<and>" and Not
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   552
  (SML and and and)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   553
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   554
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   555
code_gen in_interval (SML "examples/bool1.ML")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   556
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   557
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   558
  \lstsml{Thy/examples/bool1.ML}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   559
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   560
  Though this is correct code, it is a little bit unsatisfactory:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   561
  boolean values and operators are materialized as distinguished
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   562
  entities with have nothing to do with the SML-builtin notion
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   563
  of \qt{bool}.  This results in less readable code;
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   564
  additionally, eager evaluation may cause programs to
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   565
  loop or break which would perfectly terminate when
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   566
  the existing SML \qt{bool} would be used.  To map
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   567
  the HOL \qt{bool} on SML \qt{bool}, we may use
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   568
  \qn{custom serializations}:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   569
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   570
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   571
code_type bool
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   572
  (SML "bool")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   573
code_const True and False and "op \<and>"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   574
  (SML "true" and "false" and "_ andalso _")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   575
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   576
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   577
  The \isasymCODETYPE commad takes a type constructor
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   578
  as arguments together with a list of custom serializations.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   579
  Each custom serialization starts with a target language
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   580
  identifier followed by an expression, which during
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   581
  code serialization is inserted whenever the type constructor
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   582
  would occur.  For constants, \isasymCODECONST implements
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   583
  the corresponding mechanism.  Each \qt{\_} in
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   584
  a serialization expression is treated as a placeholder
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   585
  for the type constructor's (the constant's) arguments.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   586
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   587
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   588
code_reserved SML
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   589
  bool true false
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   590
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   591
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   592
  To assert that the existing \qt{bool}, \qt{true} and \qt{false}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   593
  is not used for generated code, we use \isasymCODERESERVED.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   594
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   595
  After this setup, code looks quite more readable:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   596
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   597
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   598
code_gen in_interval (SML "examples/bool2.ML")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   599
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   600
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   601
  \lstsml{Thy/examples/bool2.ML}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   602
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   603
  This still is not perfect: the parentheses
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   604
  around \qt{andalso} are superfluous.  Though the serializer
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   605
  by no means attempts to imitate the rich Isabelle syntax
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   606
  framework, it provides some common idioms, notably
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   607
  associative infixes with precedences which may be used here:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   608
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   609
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   610
code_const "op \<and>"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   611
  (SML infixl 1 "andalso")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   612
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   613
code_gen in_interval (SML "examples/bool3.ML")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   614
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   615
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   616
  \lstsml{Thy/examples/bool3.ML}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   617
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   618
  Next, we try to map HOL pairs to SML pairs, using the
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   619
  infix \qt{ * } type constructor and parentheses:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   620
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   621
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   622
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   623
code_type *
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   624
  (SML)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   625
code_const Pair
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   626
  (SML)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   627
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   628
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   629
code_type *
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   630
  (SML infix 2 "*")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   631
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   632
code_const Pair
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   633
  (SML "!((_),/ (_))")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   634
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   635
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   636
  The initial bang \qt{!} tells the serializer to never put
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   637
  parentheses around the whole expression (they are already present),
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   638
  while the parentheses around argument place holders
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   639
  tell not to put parentheses around the arguments.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   640
  The slash \qt{/} (followed by arbitrary white space)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   641
  inserts a space which may be used as a break if necessary
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   642
  during pretty printing.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   643
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   644
  So far, we did only provide more idiomatic serializations for
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   645
  constructs which would be executable on their own.  Target-specific
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   646
  serializations may also be used to \emph{implement} constructs
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   647
  which have no implicit notion of executability.  For example,
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   648
  take the HOL integers:
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   649
*}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   650
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   651
definition
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   652
  double_inc :: "int \<Rightarrow> int"
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   653
  "double_inc k = 2 * k + 1"
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   654
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   655
code_gen double_inc (SML "examples/integers.ML")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   656
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   657
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   658
  will fail: @{typ int} in HOL is implemented using a quotient
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   659
  type, which does not provide any notion of executability.
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   660
  \footnote{Eventually, we also want to provide executability
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   661
  for quotients.}.  However, we could use the SML builtin
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   662
  integers:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   663
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   664
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   665
code_type int
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   666
  (SML "IntInf.int")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   667
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   668
code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   669
    and "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   670
  (SML "IntInf.+ (_, _)" and "IntInf.* (_, _)")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   671
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   672
code_gen double_inc (SML "examples/integers.ML")
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   673
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   674
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   675
  resulting in:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   676
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   677
  \lstsml{Thy/examples/integers.ML}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   678
*}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   679
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   680
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   681
  These examples give a glimpse what powerful mechanisms
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   682
  custom serializations provide; however their usage
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   683
  requires careful thinking in order not to introduce
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   684
  inconsistencies -- or, in other words:
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   685
  custom serializations are completely axiomatic.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   686
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   687
  A further noteworthy details is that any special
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   688
  character in a custom serialization may be quoted
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   689
  using \qt{'}; thus, in \qt{fn '\_ => \_} the first
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   690
  \qt{'\_} is a proper underscore while the
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   691
  second \qt{\_} is a placeholder.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   692
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   693
  The HOL theories provide further
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   694
  examples for custom serializations and form
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   695
  a recommended tutorial on how to use them properly.
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   696
*}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   697
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   698
subsection {* Concerning operational equality *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   699
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   700
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   701
  Surely you have already noticed how equality is treated
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   702
  by the code generator:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   703
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   704
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   705
fun
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   706
  collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   707
  "collect_duplicates xs ys [] = xs"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   708
  "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   709
    then if z \<in> set ys
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   710
      then collect_duplicates xs ys zs
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   711
      else collect_duplicates xs (z#ys) zs
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   712
    else collect_duplicates (z#xs) (z#ys) zs)"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   713
termination by (auto_term "measure (length o snd o snd)")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   714
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   715
lemmas [code func] = collect_duplicates.simps
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   716
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   717
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   718
text {*
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   719
  The membership test during preprocessing is rewriting,
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   720
  resulting in @{const List.memberl}, which itself
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   721
  performs an explicit equality check.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   722
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   723
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   724
code_gen collect_duplicates (SML "examples/collect_duplicates.ML")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   725
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   726
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   727
  \lstsml{Thy/examples/collect_duplicates.ML}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   728
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   729
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   730
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   731
  Obviously, polymorphic equality is implemented the Haskell
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   732
  way using a type class.  How is this achieved?  By an
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   733
  almost trivial definition in the HOL setup:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   734
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   735
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   736
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   737
setup {* Sign.add_path "foo" *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   738
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   739
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   740
class eq =
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   741
  fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   742
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   743
defs
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   744
  eq (*[symmetric, code inline, code func]*): "eq \<equiv> (op =)"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   745
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   746
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   747
  This merely introduces a class @{text eq} with corresponding
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   748
  operation @{const eq}, which by definition is isomorphic
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   749
  to @{const "op ="}; the preprocessing framework does the rest.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   750
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   751
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   752
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   753
lemmas [code noinline] = eq
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   754
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   755
hide (open) "class" eq
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   756
hide (open) const eq
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   757
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   758
lemmas [symmetric, code func] = eq_def
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   759
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   760
setup {* Sign.parent_path *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   761
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   762
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   763
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   764
  For datatypes, instances of @{text eq} are implicitly derived
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   765
  when possible.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   766
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   767
  Though this class is designed to get rarely in the way, there
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   768
  are some cases when it suddenly comes to surface:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   769
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   770
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   771
subsubsection {* code lemmas and customary serializations for equality *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   772
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   773
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   774
  Examine the following:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   775
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   776
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   777
code_const "op = \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   778
  (SML "!(_ : IntInf.int = _)")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   779
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   780
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   781
  What is wrong here? Since @{const "op ="} is nothing else then
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   782
  a plain constant, this customary serialization will refer
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   783
  to polymorphic equality @{const "op = \<Colon> 'a \<Rightarrow> 'a \<Rightarrow> bool"}.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   784
  Instead, we want the specific equality on @{typ int},
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   785
  by using the overloaded constant @{const "Code_Generator.eq"}:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   786
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   787
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   788
code_const "Code_Generator.eq \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   789
  (SML "!(_ : IntInf.int = _)")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   790
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   791
subsubsection {* typedecls interpretated by customary serializations *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   792
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   793
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   794
  A common idiom is to use unspecified types for formalizations
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   795
  and interpret them for a specific target language:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   796
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   797
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   798
typedecl key
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   799
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   800
fun
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   801
  lookup :: "(key \<times> 'a) list \<Rightarrow> key \<Rightarrow> 'a option" where
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   802
  "lookup [] l = None"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   803
  "lookup ((k, v) # xs) l = (if k = l then Some v else lookup xs l)"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   804
termination by (auto_term "measure (length o fst)")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   805
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   806
lemmas [code func] = lookup.simps
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   807
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   808
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   809
code_type key
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   810
  (SML "string")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   811
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   812
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   813
  This, though, is not sufficient: @{typ key} is no instance
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   814
  of @{text eq} since @{typ key} is no datatype; the instance
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   815
  has to be declared manually, including a serialization
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   816
  for the particular instance of @{const "Code_Generator.eq"}:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   817
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   818
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   819
instance key :: eq ..
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   820
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   821
code_const "Code_Generator.eq \<Colon> key \<Rightarrow> key \<Rightarrow> bool"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   822
  (SML "!(_ : string = _)")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   823
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   824
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   825
  Then everything goes fine:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   826
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   827
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   828
code_gen lookup (SML "examples/lookup.ML")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   829
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   830
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   831
  \lstsml{Thy/examples/lookup.ML}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   832
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   833
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   834
subsubsection {* lexicographic orderings and coregularity *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   835
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   836
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   837
  Another subtlety
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   838
  enters the stage when definitions of overloaded constants
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   839
  are dependent on operational equality.  For example, let
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   840
  us define a lexicographic ordering on tuples: \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   841
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   842
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   843
(*<*)
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   844
setup {* Sign.add_path "foobar" *}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   845
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   846
class eq = fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   847
class ord =
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   848
  fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" ("(_/ \<^loc>\<le> _)" [50, 51] 50)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   849
  fixes less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" ("(_/ \<^loc>< _)" [50, 51] 50)
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   850
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   851
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   852
instance * :: (ord, ord) ord
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   853
  "p1 < p2 \<equiv> let (x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) = p1; (x2, y2) = p2 in
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   854
    x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)"
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   855
  "p1 \<le> p2 \<equiv> p1 < p2 \<or> (p1 \<Colon> 'a\<Colon>ord \<times> 'b\<Colon>ord)  = p2" ..
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   856
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   857
(*<*)
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   858
hide "class" eq ord
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   859
hide const eq less_eq less
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   860
setup {* Sign.parent_path *}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   861
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   862
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   863
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   864
  Then code generation will fail.  Why?  The definition
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   865
  of @{const "op \<le>"} depends on equality on both arguments,
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   866
  which are polymorphic and impose an additional @{text eq}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   867
  class constraint, thus violating the type discipline
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   868
  for class operations.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   869
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   870
  The solution is to add @{text eq} to both sort arguments:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   871
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   872
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   873
instance * :: ("{eq, ord}", "{eq, ord}") ord
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   874
  "p1 < p2 \<equiv> let (x1 \<Colon> 'a\<Colon>{eq, ord}, y1 \<Colon> 'b\<Colon>{eq, ord}) = p1; (x2, y2) = p2 in
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   875
    x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)"
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   876
  "p1 \<le> p2 \<equiv> p1 < p2 \<or> (p1 \<Colon> 'a\<Colon>{eq, ord} \<times> 'b\<Colon>{eq, ord})  = p2" ..
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   877
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   878
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   879
  Then code generation succeeds:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   880
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   881
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   882
code_gen "op \<le> \<Colon> 'a\<Colon>{eq, ord} \<times> 'b\<Colon>{eq, ord} \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   883
  (SML "examples/lexicographic.ML")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   884
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   885
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   886
  \lstsml{Thy/examples/lexicographic.ML}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   887
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   888
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   889
subsubsection {* Haskell serialization *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   890
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   891
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   892
  For convenience, the default
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   893
  HOL setup for Haskell maps the @{text eq} class to
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   894
  its counterpart in Haskell, giving custom serializations
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   895
  for the class (\isasymCODECLASS) and its operation:
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   896
*}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   897
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   898
(*<*)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   899
setup {* Sign.add_path "bar" *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   900
class eq = fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   901
(*>*)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   902
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   903
code_class eq
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   904
  (Haskell "Eq" where eq \<equiv> "(==)")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   905
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   906
code_const eq
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   907
  (Haskell infixl 4 "==")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   908
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   909
(*<*)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   910
hide "class" eq
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   911
hide const eq
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   912
setup {* Sign.parent_path *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   913
(*>*)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   914
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   915
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   916
  A problem now occurs whenever a type which
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   917
  is an instance of @{text eq} in HOL is mapped
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   918
  on a Haskell-builtin type which is also an instance
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   919
  of Haskell @{text Eq}:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   920
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   921
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   922
typedecl bar
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   923
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   924
instance bar :: eq ..
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   925
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   926
code_type bar
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   927
  (Haskell "Integer")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   928
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   929
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   930
    The code generator would produce
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   931
    an additional instance, which of course is rejected.
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   932
    To suppress this additional instance, use
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   933
    \isasymCODEINSTANCE:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   934
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   935
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   936
code_instance bar :: eq
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   937
  (Haskell -)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   938
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   939
subsection {* Types matter *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   940
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   941
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   942
  Imagine the following quick-and-dirty setup for implementing
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   943
  some sets as lists:
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   944
*}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   945
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   946
code_type set
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   947
  (SML "_ list")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   948
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   949
code_const "{}" and insert
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   950
  (SML "![]" and infixl 7 "::")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   951
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   952
definition
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   953
  dummy_set :: "nat set"
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   954
  "dummy_set = {1, 2, 3}"
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   955
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   956
(*print_codethms (dummy_set)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   957
code_gen dummy_set (SML -)*)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   958
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   959
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   960
(* shadowing by user-defined serializations, understanding the type game,
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   961
reflexive equations, dangerous equations *)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   962
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   963
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   964
subsection {* Axiomatic extensions *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   965
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   966
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   967
  \begin{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   968
    The extensions introduced in this section, though working
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   969
    in practice, are not the cream of the crop.  They will
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   970
    eventually be replaced by more mature approaches.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   971
  \end{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   972
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   973
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   974
(* code_modulename *)
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   975
(* existing libraries, code inline code_constsubst, code_abstype*)
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   976
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   977
section {* ML interfaces \label{sec:ml} *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   978
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   979
subsection {* Constants with type discipline: codegen\_consts.ML *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   980
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   981
text %mlref {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   982
  \begin{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   983
  @{index_ML_type CodegenConsts.const} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   984
  @{index_ML CodegenConsts.inst_of_typ: "theory -> string * typ -> CodegenConsts.const"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   985
  @{index_ML CodegenConsts.typ_of_inst: "theory -> CodegenConsts.const -> string * typ"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   986
  @{index_ML CodegenConsts.norm: "theory -> CodegenConsts.const -> CodegenConsts.const"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   987
  @{index_ML CodegenConsts.norm_of_typ: "theory -> string * typ -> CodegenConsts.const"}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   988
  \end{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   989
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   990
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   991
subsection {* Executable theory content: codegen\_data.ML *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   992
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   993
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   994
  This Pure module implements the core notions of
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   995
  executable content of a theory.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   996
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   997
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   998
subsubsection {* Suspended theorems *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   999
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1000
text %mlref {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1001
  \begin{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1002
  @{index_ML_type CodegenData.lthms} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1003
  @{index_ML CodegenData.lazy: "(unit -> thm list) -> CodegenData.lthms"}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1004
  \end{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1005
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1006
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1007
subsubsection {* Executable content *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1008
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1009
text %mlref {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1010
  \begin{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1011
  @{index_ML CodegenData.add_func: "thm -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1012
  @{index_ML CodegenData.del_func: "thm -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1013
  @{index_ML CodegenData.add_funcl: "CodegenConsts.const * CodegenData.lthms -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1014
  @{index_ML CodegenData.add_datatype: "string * (((string * sort) list * (string * typ list) list) * CodegenData.lthms) -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1015
  @{index_ML CodegenData.del_datatype: "string -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1016
  @{index_ML CodegenData.add_inline: "thm -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1017
  @{index_ML CodegenData.del_inline: "thm -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1018
  @{index_ML CodegenData.add_inline_proc: "(theory -> cterm list -> thm list) -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1019
  @{index_ML CodegenData.add_preproc: "(theory -> thm list -> thm list) -> theory -> theory"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1020
  @{index_ML CodegenData.these_funcs: "theory -> CodegenConsts.const -> thm list"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1021
  @{index_ML CodegenData.get_datatype: "theory -> string -> ((string * sort) list * (string * typ list) list) option"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1022
  @{index_ML CodegenData.get_datatype_of_constr: "theory -> CodegenConsts.const -> string option"}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1023
  \end{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1024
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1025
  \begin{description}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1026
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1027
  \item @{ML CodegenData.add_func}~@{text "thm"}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1028
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1029
  \end{description}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1030
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1031
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1032
subsection {* Further auxiliary *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1033
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1034
text %mlref {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1035
  \begin{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1036
  @{index_ML CodegenConsts.const_ord: "CodegenConsts.const * CodegenConsts.const -> order"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1037
  @{index_ML CodegenConsts.eq_const: "CodegenConsts.const * CodegenConsts.const -> bool"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1038
  @{index_ML CodegenConsts.consts_of: "theory -> term -> CodegenConsts.const list"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1039
  @{index_ML CodegenConsts.read_const: "theory -> string -> CodegenConsts.const"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1040
  @{index_ML_structure CodegenConsts.Consttab} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1041
  @{index_ML CodegenData.typ_func: "theory -> thm -> typ"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1042
  @{index_ML CodegenData.rewrite_func: "thm list -> thm -> thm"} \\
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1043
  \end{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1044
*}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1045
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1046
subsection {* Implementing code generator applications *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1047
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1048
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1049
  \begin{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1050
    Some interfaces discussed here have not reached
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1051
    a final state yet.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1052
    Changes likely to occur in future.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1053
  \end{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1054
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1055
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1056
subsubsection {* Data depending on the theory's executable content *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1057
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1058
subsubsection {* Datatype hooks *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1059
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1060
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1061
  \emph{Happy proving, happy hacking!}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1062
*}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
  1063
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1064
end