doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy
author wenzelm
Mon, 14 Jul 2008 17:51:48 +0200
changeset 27579 97ce525f664c
parent 27557 151731493264
child 27609 b23c9ad0fe7d
permissions -rw-r--r--
end_theory: no result;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
     1
20948
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
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
     7
uses "../../../antiquote_setup.ML"
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     8
begin
9b9910b82645 initial draft
haftmann
parents:
diff changeset
     9
27103
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 27025
diff changeset
    10
ML {* CodeTarget.code_width := 74 *}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    11
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    12
(*>*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    13
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    14
chapter {* Code generation from Isabelle theories *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    15
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    16
section {* Introduction *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    17
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    18
subsection {* Motivation *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    19
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    20
text {*
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    21
  Executing formal specifications as programs is a well-established
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    22
  topic in the theorem proving community.  With increasing
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    23
  application of theorem proving systems in the area of
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    24
  software development and verification, its relevance manifests
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    25
  for running test cases and rapid prototyping.  In logical
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    26
  calculi like constructive type theory,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    27
  a notion of executability is implicit due to the nature
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    28
  of the calculus.  In contrast, specifications in Isabelle
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    29
  can be highly non-executable.  In order to bridge
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    30
  the gap between logic and executable specifications,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    31
  an explicit non-trivial transformation has to be applied:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    32
  code generation.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    33
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    34
  This tutorial introduces a generic code generator for the
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    35
  Isabelle system \cite{isa-tutorial}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    36
  Generic in the sense that the
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    37
  \qn{target language} for which code shall ultimately be
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    38
  generated is not fixed but may be an arbitrary state-of-the-art
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    39
  functional programming language (currently, the implementation
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    40
  supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    41
  \cite{haskell-revised-report}).
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    42
  We aim to provide a
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    43
  versatile environment
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    44
  suitable for software development and verification,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    45
  structuring the process
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    46
  of code generation into a small set of orthogonal principles
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    47
  while achieving a big coverage of application areas
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    48
  with maximum flexibility.
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
    49
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    50
  Conceptually the code generator framework is part
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    51
  of Isabelle's @{text Pure} meta logic; the object logic
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    52
  @{text HOL} which is an extension of @{text Pure}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    53
  already comes with a reasonable framework setup and thus provides
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    54
  a good working horse for raising code-generation-driven
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    55
  applications.  So, we assume some familiarity and experience
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    56
  with the ingredients of the @{text HOL} \emph{Main} theory
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    57
  (see also \cite{isa-tutorial}).
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    58
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    59
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    60
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    61
subsection {* Overview *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    62
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    63
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    64
  The code generator aims to be usable with no further ado
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    65
  in most cases while allowing for detailed customization.
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    66
  This manifests in the structure of this tutorial:
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    67
  we start with a generic example \secref{sec:example}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    68
  and introduce code generation concepts \secref{sec:concept}.
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    69
  Section
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
    70
  \secref{sec:basics} explains how to use the framework naively,
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    71
  presuming a reasonable default setup.  Then, section
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    72
  \secref{sec:advanced} deals with advanced topics,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    73
  introducing further aspects of the code generator framework
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    74
  in a motivation-driven manner.  Last, section \secref{sec:ml}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    75
  introduces the framework's internal programming interfaces.
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
    76
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    77
  \begin{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    78
    Ultimately, the code generator which this tutorial deals with
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    79
    is supposed to replace the already established code generator
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    80
    by Stefan Berghofer \cite{Berghofer-Nipkow:2002}.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
    81
    So, for the moment, there are two distinct code generators
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    82
    in Isabelle.
22916
haftmann
parents: 22885
diff changeset
    83
    Also note that while the framework itself is
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    84
    object-logic independent, only @{text HOL} provides a reasonable
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    85
    framework setup.    
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    86
  \end{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    87
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    88
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
    89
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    90
section {* An example: a simple theory of search trees \label{sec:example} *}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    91
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    92
text {*
22916
haftmann
parents: 22885
diff changeset
    93
  When writing executable specifications using @{text HOL},
haftmann
parents: 22885
diff changeset
    94
  it is convenient to use
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    95
  three existing packages: the datatype package for defining
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    96
  datatypes, the function package for (recursive) functions,
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    97
  and the class package for overloaded definitions.
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    98
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
    99
  We develope a small theory of search trees; trees are represented
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   100
  as a datatype with key type @{typ "'a"} and value type @{typ "'b"}:
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   101
*}
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   102
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   103
datatype ('a, 'b) searchtree = Leaf "'a\<Colon>linorder" 'b
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   104
  | Branch "('a, 'b) searchtree" "'a" "('a, 'b) searchtree"
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   105
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   106
text {*
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   107
  \noindent Note that we have constrained the type of keys
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   108
  to the class of total orders, @{text linorder}.
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   109
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   110
  We define @{text find} and @{text update} functions:
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   111
*}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   112
25870
haftmann
parents: 25533
diff changeset
   113
primrec
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   114
  find :: "('a\<Colon>linorder, 'b) searchtree \<Rightarrow> 'a \<Rightarrow> 'b option" where
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   115
  "find (Leaf key val) it = (if it = key then Some val else None)"
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   116
  | "find (Branch t1 key t2) it = (if it \<le> key then find t1 it else find t2 it)"
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   117
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   118
fun
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   119
  update :: "'a\<Colon>linorder \<times> 'b \<Rightarrow> ('a, 'b) searchtree \<Rightarrow> ('a, 'b) searchtree" where
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   120
  "update (it, entry) (Leaf key val) = (
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   121
    if it = key then Leaf key entry
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   122
      else if it \<le> key
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   123
      then Branch (Leaf it entry) it (Leaf key val)
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   124
      else Branch (Leaf key val) it (Leaf it entry)
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   125
   )"
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   126
  | "update (it, entry) (Branch t1 key t2) = (
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   127
    if it \<le> key
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   128
      then (Branch (update (it, entry) t1) key t2)
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   129
      else (Branch t1 key (update (it, entry) t2))
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   130
   )"
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   131
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   132
text {*
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   133
  \noindent For testing purpose, we define a small example
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   134
  using natural numbers @{typ nat} (which are a @{text linorder})
23130
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   135
  as keys and list of nats as values:
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   136
*}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   137
23130
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   138
definition
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   139
  example :: "(nat, nat list) searchtree"
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   140
where
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   141
  "example = update (Suc (Suc (Suc (Suc 0))), [Suc (Suc 0), Suc (Suc 0)]) (update (Suc (Suc (Suc 0)), [Suc (Suc (Suc 0))])
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   142
    (update (Suc (Suc 0), [Suc (Suc 0)]) (Leaf (Suc 0) [])))"
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   143
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   144
text {*
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   145
  \noindent Then we generate code
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   146
*}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   147
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   148
export_code example (*<*)in SML (*>*)in SML file "examples/tree.ML"
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   149
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   150
text {*
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   151
  \noindent which looks like:
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   152
  \lstsml{Thy/examples/tree.ML}
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   153
*}
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   154
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   155
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   156
section {* Code generation concepts and process \label{sec:concept} *}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   157
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   158
text {*
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   159
  \begin{figure}[h]
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   160
  \centering
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   161
  \includegraphics[width=0.7\textwidth]{codegen_process}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   162
  \caption{code generator -- processing overview}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   163
  \label{fig:process}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   164
  \end{figure}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
   165
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   166
  The code generator employs a notion of executability
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   167
  for three foundational executable ingredients known
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   168
  from functional programming:
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   169
  \emph{defining equations}, \emph{datatypes}, and
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   170
  \emph{type classes}. A defining equation as a first approximation
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   171
  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
   172
  (an equation headed by a constant @{text f} with arguments
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   173
  @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right hand side @{text t}).
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   174
  Code generation aims to turn defining equations
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   175
  into a functional program by running through
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   176
  a process (see figure \ref{fig:process}):
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   177
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   178
  \begin{itemize}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   179
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   180
    \item Out of the vast collection of theorems proven in a
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   181
      \qn{theory}, a reasonable subset modeling
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   182
      defining equations is \qn{selected}.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   183
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   184
    \item On those selected theorems, certain
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   185
      transformations are carried out
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   186
      (\qn{preprocessing}).  Their purpose is to turn theorems
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   187
      representing non- or badly executable
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   188
      specifications into equivalent but executable counterparts.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   189
      The result is a structured collection of \qn{code theorems}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   190
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
   191
    \item These \qn{code theorems} then are \qn{translated}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   192
      into an Haskell-like intermediate
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   193
      language.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   194
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   195
    \item Finally, out of the intermediate language the final
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   196
      code in the desired \qn{target language} is \qn{serialized}.
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   197
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   198
  \end{itemize}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   199
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   200
  From these steps, only the two last are carried out
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   201
  outside the logic; by keeping this layer as
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   202
  thin as possible, the amount of code to trust is
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   203
  kept to a minimum.
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
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   208
section {* Basics \label{sec:basics} *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   209
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   210
subsection {* Invoking the code generator *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   211
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   212
text {*
22916
haftmann
parents: 22885
diff changeset
   213
  Thanks to a reasonable setup of the @{text HOL} theories, in
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   214
  most cases code generation proceeds without further ado:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   215
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   216
25870
haftmann
parents: 25533
diff changeset
   217
primrec
22473
753123c89d72 explizit "type" superclass
haftmann
parents: 22423
diff changeset
   218
  fac :: "nat \<Rightarrow> nat" where
753123c89d72 explizit "type" superclass
haftmann
parents: 22423
diff changeset
   219
    "fac 0 = 1"
753123c89d72 explizit "type" superclass
haftmann
parents: 22423
diff changeset
   220
  | "fac (Suc n) = Suc n * fac n"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   221
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   222
text {*
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   223
  \noindent This executable specification is now turned to SML code:
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   224
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   225
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   226
export_code fac (*<*)in SML(*>*)in SML file "examples/fac.ML"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   227
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   228
text {*
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   229
  \noindent  The @{text "\<EXPORTCODE>"} command takes a space-separated list of
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   230
  constants together with \qn{serialization directives}
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   231
  These start with a \qn{target language}
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   232
  identifier, followed by a file specification
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   233
  where to write the generated code to.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   234
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   235
  Internally, the defining equations for all selected
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   236
  constants are taken, including any transitively required
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   237
  constants, datatypes and classes, resulting in the following
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   238
  code:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   239
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   240
  \lstsml{Thy/examples/fac.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   241
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   242
  The code generator will complain when a required
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   243
  ingredient does not provide a executable counterpart,
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   244
  e.g.~generating code
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   245
  for constants not yielding
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   246
  a defining equation (e.g.~the Hilbert choice
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   247
  operation @{text "SOME"}):
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   248
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   249
(*<*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   250
setup {* Sign.add_path "foo" *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   251
(*>*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   252
definition
21993
4b802a9e0738 updated manual
haftmann
parents: 21545
diff changeset
   253
  pick_some :: "'a list \<Rightarrow> 'a" where
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   254
  "pick_some xs = (SOME x. x \<in> set xs)"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   255
(*<*)
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   256
hide const pick_some
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   257
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   258
setup {* Sign.parent_path *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   259
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   260
definition
21993
4b802a9e0738 updated manual
haftmann
parents: 21545
diff changeset
   261
  pick_some :: "'a list \<Rightarrow> 'a" where
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   262
  "pick_some = hd"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   263
(*>*)
25870
haftmann
parents: 25533
diff changeset
   264
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   265
export_code pick_some in SML file "examples/fail_const.ML"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   266
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   267
text {* \noindent will fail. *}
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   268
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   269
subsection {* Theorem selection *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   270
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   271
text {*
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   272
  The list of all defining equations in a theory may be inspected
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   273
  using the @{text "\<PRINTCODESETUP>"} command:
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   274
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   275
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   276
print_codesetup
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   277
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   278
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   279
  \noindent which displays a table of constant with corresponding
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   280
  defining equations (the additional stuff displayed
22751
1bfd75c1f232 updated
haftmann
parents: 22550
diff changeset
   281
  shall not bother us for the moment).
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   282
22916
haftmann
parents: 22885
diff changeset
   283
  The typical @{text HOL} tools are already set up in a way that
22751
1bfd75c1f232 updated
haftmann
parents: 22550
diff changeset
   284
  function definitions introduced by @{text "\<DEFINITION>"},
25870
haftmann
parents: 25533
diff changeset
   285
  @{text "\<PRIMREC>"}, @{text "\<FUN>"},
haftmann
parents: 25533
diff changeset
   286
  @{text "\<FUNCTION>"}, @{text "\<CONSTDEFS>"},
21323
haftmann
parents: 21223
diff changeset
   287
  @{text "\<RECDEF>"} are implicitly propagated
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   288
  to this defining equation table. Specific theorems may be
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   289
  selected using an attribute: \emph{code func}. As example,
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   290
  a weight selector function:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   291
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   292
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   293
primrec
25870
haftmann
parents: 25533
diff changeset
   294
  pick :: "(nat \<times> 'a) list \<Rightarrow> nat \<Rightarrow> 'a" where
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   295
  "pick (x#xs) n = (let (k, v) = x in
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   296
    if n < k then v else pick xs (n - k))"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   297
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   298
text {*
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   299
  \noindent We want to eliminate the explicit destruction
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   300
  of @{term x} to @{term "(k, v)"}:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   301
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   302
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   303
lemma [code func]:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   304
  "pick ((k, v)#xs) n = (if n < k then v else pick xs (n - k))"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   305
  by simp
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   306
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   307
export_code pick (*<*)in SML(*>*) in SML file "examples/pick1.ML"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   308
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   309
text {*
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   310
  \noindent This theorem now is used for generating code:
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   311
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   312
  \lstsml{Thy/examples/pick1.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   313
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   314
  \noindent It might be convenient to remove the pointless original
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   315
  equation, using the \emph{func del} attribute:
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   316
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   317
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   318
lemmas [code func del] = pick.simps 
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   319
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   320
export_code pick (*<*)in SML(*>*) in SML file "examples/pick2.ML"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   321
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   322
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   323
  \lstsml{Thy/examples/pick2.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   324
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   325
  \noindent Syntactic redundancies are implicitly dropped. For example,
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   326
  using a modified version of the @{const fac} function
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   327
  as defining equation, the then redundant (since
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   328
  syntactically subsumed) original defining equations
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   329
  are dropped, resulting in a warning:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   330
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   331
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   332
lemma [code func]:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   333
  "fac n = (case n of 0 \<Rightarrow> 1 | Suc m \<Rightarrow> n * fac m)"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   334
  by (cases n) simp_all
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   335
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   336
export_code fac (*<*)in SML(*>*)in SML file "examples/fac_case.ML"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   337
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   338
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   339
  \lstsml{Thy/examples/fac_case.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   340
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   341
  \begin{warn}
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   342
    The attributes \emph{code} and \emph{code del}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   343
    associated with the existing code generator also apply to
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   344
    the new one: \emph{code} implies \emph{code func},
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   345
    and \emph{code del} implies \emph{code func del}.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   346
  \end{warn}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   347
*}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   348
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   349
subsection {* Type classes *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   350
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   351
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   352
  Type classes enter the game via the Isar class package.
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   353
  For a short introduction how to use it, see \cite{isabelle-classes};
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   354
  here we just illustrate its impact on code generation.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   355
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   356
  In a target language, type classes may be represented
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   357
  natively (as in the case of Haskell).  For languages
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   358
  like SML, they are implemented using \emph{dictionaries}.
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   359
  Our following example specifies a class \qt{null},
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   360
  assigning to each of its inhabitants a \qt{null} value:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   361
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   362
22473
753123c89d72 explizit "type" superclass
haftmann
parents: 22423
diff changeset
   363
class null = type +
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   364
  fixes null :: 'a
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   365
25870
haftmann
parents: 25533
diff changeset
   366
primrec
haftmann
parents: 25533
diff changeset
   367
  head :: "'a\<Colon>null list \<Rightarrow> 'a" where
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   368
  "head [] = null"
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   369
  | "head (x#xs) = x"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   370
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   371
text {*
25533
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   372
 \noindent  We provide some instances for our @{text null}:
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   373
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   374
25533
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   375
instantiation option and list :: (type) null
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   376
begin
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   377
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   378
definition
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   379
  "null = None"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   380
25533
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   381
definition
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   382
  "null = []"
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   383
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   384
instance ..
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   385
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   386
end
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   387
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   388
text {*
25533
0140cc7b26ad added something about instantiation target
haftmann
parents: 25411
diff changeset
   389
  \noindent Constructing a dummy example:
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   390
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   391
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   392
definition
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   393
  "dummy = head [Some (Suc 0), None]"
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   394
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   395
text {*
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   396
  Type classes offer a suitable occasion to introduce
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   397
  the Haskell serializer.  Its usage is almost the same
21075
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   398
  as SML, but, in accordance with conventions
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   399
  some Haskell systems enforce, each module ends
d6742ff3b522 continued
haftmann
parents: 21058
diff changeset
   400
  up in a single file. The module hierarchy is reflected in
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   401
  the file system, with root directory given as file specification.
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   402
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   403
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   404
export_code dummy in Haskell file "examples/"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   405
  (* NOTE: you may use Haskell only once in this document, otherwise
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   406
  you have to work in distinct subdirectories *)
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   407
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   408
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   409
  \lsthaskell{Thy/examples/Codegen.hs}
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   410
  \noindent (we have left out all other modules).
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   411
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   412
  \medskip
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   413
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   414
  The whole code in SML with explicit dictionary passing:
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   415
*}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   416
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   417
export_code dummy (*<*)in SML(*>*)in SML file "examples/class.ML"
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   418
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   419
text {*
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   420
  \lstsml{Thy/examples/class.ML}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   421
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   422
  \medskip
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   423
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   424
  \noindent or in OCaml:
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   425
*}
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   426
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   427
export_code dummy in OCaml file "examples/class.ocaml"
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   428
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   429
text {*
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   430
  \lstsml{Thy/examples/class.ocaml}
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   431
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   432
  \medskip The explicit association of constants
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   433
  to classes can be inspected using the @{text "\<PRINTCLASSES>"}
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   434
  command.
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   435
*}
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
   436
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
   437
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   438
section {* Recipes and advanced topics \label{sec:advanced} *}
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   439
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   440
text {*
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   441
  In this tutorial, we do not attempt to give an exhaustive
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   442
  description of the code generator framework; instead,
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   443
  we cast a light on advanced topics by introducing
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   444
  them together with practically motivated examples.  Concerning
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   445
  further reading, see
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   446
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   447
  \begin{itemize}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   448
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   449
  \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   450
    for exhaustive syntax diagrams.
24193
926dde4d96de updated
haftmann
parents: 23850
diff changeset
   451
  \item or \cite{Haftmann-Nipkow:2007:codegen} which deals with foundational issues
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   452
    of the code generator framework.
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   453
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   454
  \end{itemize}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   455
*}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   456
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   457
subsection {* Library theories \label{sec:library} *}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   458
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   459
text {*
22916
haftmann
parents: 22885
diff changeset
   460
  The @{text HOL} @{text Main} theory already provides a code
haftmann
parents: 22885
diff changeset
   461
  generator setup
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   462
  which should be suitable for most applications. Common extensions
22916
haftmann
parents: 22885
diff changeset
   463
  and modifications are available by certain theories of the @{text HOL}
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   464
  library; beside being useful in applications, they may serve
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   465
  as a tutorial for customizing the code generator setup.
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   466
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   467
  \begin{description}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   468
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   469
    \item[@{text "Code_Integer"}] represents @{text HOL} integers by big
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   470
       integer literals in target languages.
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   471
    \item[@{text "Code_Char"}] represents @{text HOL} characters by 
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   472
       character literals in target languages.
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   473
    \item[@{text "Code_Char_chr"}] like @{text "Code_Char"},
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   474
       but also offers treatment of character codes; includes
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   475
       @{text "Code_Integer"}.
23850
f1434532a562 updated
haftmann
parents: 23130
diff changeset
   476
    \item[@{text "Efficient_Nat"}] \label{eff_nat} implements natural numbers by integers,
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   477
       which in general will result in higher efficency; pattern
25369
5200374fda5d replaced @{const} (allows name only) by proper @{term};
wenzelm
parents: 24994
diff changeset
   478
       matching with @{term "0\<Colon>nat"} / @{const "Suc"}
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   479
       is eliminated;  includes @{text "Code_Integer"}.
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   480
    \item[@{text "Code_Index"}] provides an additional datatype
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   481
       @{text index} which is mapped to target-language built-in integers.
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   482
       Useful for code setups which involve e.g. indexing of
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   483
       target-language arrays.
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   484
    \item[@{text "Code_Message"}] provides an additional datatype
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   485
       @{text message_string} which is isomorphic to strings;
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   486
       @{text message_string}s are mapped to target-language strings.
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24628
diff changeset
   487
       Useful for code setups which involve e.g. printing (error) messages.
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   488
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   489
  \end{description}
22916
haftmann
parents: 22885
diff changeset
   490
haftmann
parents: 22885
diff changeset
   491
  \begin{warn}
haftmann
parents: 22885
diff changeset
   492
    When importing any of these theories, they should form the last
haftmann
parents: 22885
diff changeset
   493
    items in an import list.  Since these theories adapt the
haftmann
parents: 22885
diff changeset
   494
    code generator setup in a non-conservative fashion,
haftmann
parents: 22885
diff changeset
   495
    strange effects may occur otherwise.
haftmann
parents: 22885
diff changeset
   496
  \end{warn}
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   497
*}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   498
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   499
subsection {* Preprocessing *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   500
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   501
text {*
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   502
  Before selected function theorems are turned into abstract
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   503
  code, a chain of definitional transformation steps is carried
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   504
  out: \emph{preprocessing}.  In essence, the preprocessort
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   505
  consists of two components: a \emph{simpset} and \emph{function transformators}.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   506
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   507
  The \emph{simpset} allows to employ the full generality of the Isabelle
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   508
  simplifier.  Due to the interpretation of theorems
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
   509
  of defining equations, rewrites are applied to the right
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   510
  hand side and the arguments of the left hand side of an
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   511
  equation, but never to the constant heading the left hand side.
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   512
  An important special case are \emph{inline theorems} which may be
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   513
  declared an undeclared using the
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   514
  \emph{code inline} or \emph{code inline del} attribute respectively.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   515
  Some common applications:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   516
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   517
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   518
text_raw {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   519
  \begin{itemize}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   520
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   521
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   522
text {*
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   523
     \item replacing non-executable constructs by executable ones:
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   524
*}     
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   525
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   526
  lemma [code inline]:
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   527
    "x \<in> set xs \<longleftrightarrow> x mem xs" by (induct xs) simp_all
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   528
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   529
text {*
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   530
     \item eliminating superfluous constants:
21089
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   531
*}
cf6defa31209 continued
haftmann
parents: 21075
diff changeset
   532
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   533
  lemma [code inline]:
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   534
    "1 = Suc 0" by simp
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   535
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   536
text {*
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   537
     \item replacing executable but inconvenient constructs:
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   538
*}
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   539
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   540
  lemma [code inline]:
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   541
    "xs = [] \<longleftrightarrow> List.null xs" by (induct xs) simp_all
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   542
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   543
text_raw {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   544
  \end{itemize}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   545
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   546
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   547
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   548
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   549
  \emph{Function transformators} provide a most general interface,
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   550
  transforming a list of function theorems to another
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   551
  list of function theorems, provided that neither the heading
25369
5200374fda5d replaced @{const} (allows name only) by proper @{term};
wenzelm
parents: 24994
diff changeset
   552
  constant nor its type change.  The @{term "0\<Colon>nat"} / @{const Suc}
21323
haftmann
parents: 21223
diff changeset
   553
  pattern elimination implemented in
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
   554
  theory @{text "EfficientNat"} (\secref{eff_nat}) uses this
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   555
  interface.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   556
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   557
  \noindent The current setup of the preprocessor may be inspected using
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   558
  the @{text "\<PRINTCODESETUP>"} command.
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   559
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   560
  \begin{warn}
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
   561
    Preprocessing is \emph{no} fix point process.  Keep this in mind when
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   562
    setting up the preprocessor.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   563
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   564
    Further, the attribute \emph{code unfold}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   565
    associated with the existing code generator also applies to
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   566
    the new one: \emph{code unfold} implies \emph{code inline}.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   567
  \end{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   568
*}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   569
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   570
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   571
subsection {* Concerning operational equality *}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   572
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   573
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   574
  Surely you have already noticed how equality is treated
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   575
  by the code generator:
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   576
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   577
25870
haftmann
parents: 25533
diff changeset
   578
primrec
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   579
  collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   580
    "collect_duplicates xs ys [] = xs"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   581
  | "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   582
      then if z \<in> set ys
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   583
        then collect_duplicates xs ys zs
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   584
        else collect_duplicates xs (z#ys) zs
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   585
      else collect_duplicates (z#xs) (z#ys) zs)"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   586
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   587
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   588
  The membership test during preprocessing is rewritten,
23850
f1434532a562 updated
haftmann
parents: 23130
diff changeset
   589
  resulting in @{const List.member}, which itself
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   590
  performs an explicit equality check.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   591
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   592
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   593
export_code collect_duplicates (*<*)in SML(*>*)in SML file "examples/collect_duplicates.ML"
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   594
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   595
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   596
  \lstsml{Thy/examples/collect_duplicates.ML}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   597
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   598
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   599
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   600
  Obviously, polymorphic equality is implemented the Haskell
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   601
  way using a type class.  How is this achieved?  HOL introduces
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   602
  an explicit class @{text eq} with a corresponding operation
26732
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
   603
  @{const eq_class.eq} such that @{thm eq [no_vars]}.
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   604
  The preprocessing framework does the rest.
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   605
  For datatypes, instances of @{text eq} are implicitly derived
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   606
  when possible.  For other types, you may instantiate @{text eq}
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   607
  manually like any other type class.
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   608
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   609
  Though this @{text eq} class is designed to get rarely in
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   610
  the way, a subtlety
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   611
  enters the stage when definitions of overloaded constants
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   612
  are dependent on operational equality.  For example, let
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   613
  us define a lexicographic ordering on tuples:
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   614
*}
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   615
25870
haftmann
parents: 25533
diff changeset
   616
instantiation * :: (ord, ord) ord
haftmann
parents: 25533
diff changeset
   617
begin
haftmann
parents: 25533
diff changeset
   618
haftmann
parents: 25533
diff changeset
   619
definition
haftmann
parents: 25533
diff changeset
   620
  [code func del]: "p1 < p2 \<longleftrightarrow> (let (x1, y1) = p1; (x2, y2) = p2 in
haftmann
parents: 25533
diff changeset
   621
    x1 < x2 \<or> (x1 = x2 \<and> y1 < y2))"
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   622
25870
haftmann
parents: 25533
diff changeset
   623
definition
haftmann
parents: 25533
diff changeset
   624
  [code func del]: "p1 \<le> p2 \<longleftrightarrow> (let (x1, y1) = p1; (x2, y2) = p2 in
haftmann
parents: 25533
diff changeset
   625
    x1 < x2 \<or> (x1 = x2 \<and> y1 \<le> y2))"
haftmann
parents: 25533
diff changeset
   626
haftmann
parents: 25533
diff changeset
   627
instance ..
haftmann
parents: 25533
diff changeset
   628
haftmann
parents: 25533
diff changeset
   629
end
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   630
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   631
lemma ord_prod [code func(*<*), code func del(*>*)]:
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   632
  "(x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   633
  "(x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> (x1 = x2 \<and> y1 \<le> y2)"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   634
  unfolding less_prod_def less_eq_prod_def by simp_all
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   635
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   636
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   637
  Then code generation will fail.  Why?  The definition
25369
5200374fda5d replaced @{const} (allows name only) by proper @{term};
wenzelm
parents: 24994
diff changeset
   638
  of @{term "op \<le>"} depends on equality on both arguments,
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   639
  which are polymorphic and impose an additional @{text eq}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   640
  class constraint, thus violating the type discipline
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   641
  for class operations.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   642
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   643
  The solution is to add @{text eq} explicitly to the first sort arguments in the
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   644
  code theorems:
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   645
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   646
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   647
lemma ord_prod_code [code func]:
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   648
  "(x1 \<Colon> 'a\<Colon>{ord, eq}, y1 \<Colon> 'b\<Colon>ord) < (x2, y2) \<longleftrightarrow>
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   649
    x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   650
  "(x1 \<Colon> 'a\<Colon>{ord, eq}, y1 \<Colon> 'b\<Colon>ord) \<le> (x2, y2) \<longleftrightarrow>
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   651
    x1 < x2 \<or> (x1 = x2 \<and> y1 \<le> y2)"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   652
  unfolding ord_prod by rule+
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   653
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   654
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   655
  \noindent Then code generation succeeds:
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   656
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   657
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   658
export_code "op \<le> \<Colon> 'a\<Colon>{eq, ord} \<times> 'b\<Colon>ord \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22798
diff changeset
   659
  (*<*)in SML(*>*)in SML file "examples/lexicographic.ML"
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   660
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   661
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   662
  \lstsml{Thy/examples/lexicographic.ML}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   663
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   664
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   665
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   666
  In general, code theorems for overloaded constants may have more
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   667
  restrictive sort constraints than the underlying instance relation
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   668
  between class and type constructor as long as the whole system of
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   669
  constraints is coregular; code theorems violating coregularity
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   670
  are rejected immediately.  Consequently, it might be necessary
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   671
  to delete disturbing theorems in the code theorem table,
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   672
  as we have done here with the original definitions @{text less_prod_def}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   673
  and @{text less_eq_prod_def}.
22885
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   674
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   675
  In some cases, the automatically derived defining equations
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   676
  for equality on a particular type may not be appropriate.
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   677
  As example, watch the following datatype representing
23130
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   678
  monomorphic parametric types (where type constructors
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   679
  are referred to by natural numbers):
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   680
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   681
23130
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   682
datatype monotype = Mono nat "monotype list"
22885
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   683
(*<*)
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   684
lemma monotype_eq:
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   685
  "Mono tyco1 typargs1 = Mono tyco2 typargs2 \<equiv> 
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   686
     tyco1 = tyco2 \<and> typargs1 = typargs2" by simp
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   687
(*>*)
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   688
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   689
text {*
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   690
  Then code generation for SML would fail with a message
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   691
  that the generated code conains illegal mutual dependencies:
23130
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   692
  the theorem @{thm monotype_eq [no_vars]} already requires the
22885
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   693
  instance @{text "monotype \<Colon> eq"}, which itself requires
23130
cf50eb79d107 eliminated strings
haftmann
parents: 23106
diff changeset
   694
  @{thm monotype_eq [no_vars]};  Haskell has no problem with mutually
22885
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   695
  recursive @{text instance} and @{text function} definitions,
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   696
  but the SML serializer does not support this.
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   697
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   698
  In such cases, you have to provide you own equality equations
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   699
  involving auxiliary constants.  In our case,
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   700
  @{const [show_types] list_all2} can do the job:
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   701
*}
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   702
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   703
lemma monotype_eq_list_all2 [code func]:
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   704
  "Mono tyco1 typargs1 = Mono tyco2 typargs2 \<longleftrightarrow>
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   705
     tyco1 = tyco2 \<and> list_all2 (op =) typargs1 typargs2"
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   706
  by (simp add: list_all2_eq [symmetric])
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   707
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   708
text {*
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   709
  does not depend on instance @{text "monotype \<Colon> eq"}:
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   710
*}
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   711
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   712
export_code "op = :: monotype \<Rightarrow> monotype \<Rightarrow> bool"
22885
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   713
  (*<*)in SML(*>*)in SML file "examples/monotype.ML"
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   714
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   715
text {*
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   716
  \lstsml{Thy/examples/monotype.ML}
ebde66a71ab0 continued
haftmann
parents: 22845
diff changeset
   717
*}
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   718
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   719
subsection {* Programs as sets of theorems *}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   720
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   721
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   722
  As told in \secref{sec:concept}, code generation is based
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   723
  on a structured collection of code theorems.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   724
  For explorative purpose, this collection
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   725
  may be inspected using the @{text "\<CODETHMS>"} command:
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   726
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   727
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   728
code_thms "op mod :: nat \<Rightarrow> nat \<Rightarrow> nat"
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   729
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   730
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   731
  \noindent prints a table with \emph{all} defining equations
25369
5200374fda5d replaced @{const} (allows name only) by proper @{term};
wenzelm
parents: 24994
diff changeset
   732
  for @{term "op mod :: nat \<Rightarrow> nat \<Rightarrow> nat"}, including
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   733
  \emph{all} defining equations those equations depend
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   734
  on recursivly.  @{text "\<CODETHMS>"} provides a convenient
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   735
  mechanism to inspect the impact of a preprocessor setup
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   736
  on defining equations.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   737
  
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   738
  Similarly, the @{text "\<CODEDEPS>"} command shows a graph
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   739
  visualizing dependencies between defining equations.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   740
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   741
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   742
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   743
subsection {* Constructor sets for datatypes *}
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   744
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   745
text {*
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   746
  Conceptually, any datatype is spanned by a set of
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   747
  \emph{constructors} of type @{text "\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^isub>1 \<dots> \<alpha>\<^isub>n"}
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   748
  where @{text "{\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n}"} is excactly the set of \emph{all}
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   749
  type variables in @{text "\<tau>"}.  The HOL datatype package
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   750
  by default registers any new datatype in the table
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   751
  of datatypes, which may be inspected using
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   752
  the @{text "\<PRINTCODESETUP>"} command.
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   753
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   754
  In some cases, it may be convenient to alter or
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   755
  extend this table;  as an example, we will develope an alternative
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   756
  representation of natural numbers as binary digits, whose
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   757
  size does increase logarithmically with its value, not linear
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   758
  \footnote{Indeed, the @{text "Efficient_Nat"} theory \ref{eff_nat}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   759
    does something similar}.  First, the digit representation:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   760
*}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   761
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   762
definition Dig0 :: "nat \<Rightarrow> nat" where
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   763
  "Dig0 n = 2 * n"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   764
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   765
definition Dig1 :: "nat \<Rightarrow> nat" where
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   766
  "Dig1 n = Suc (2 * n)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   767
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   768
text {*
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   769
  \noindent We will use these two ">digits"< to represent natural numbers
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   770
  in binary digits, e.g.:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   771
*}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   772
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   773
lemma 42: "42 = Dig0 (Dig1 (Dig0 (Dig1 (Dig0 1))))"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   774
  by (simp add: Dig0_def Dig1_def)
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   775
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   776
text {*
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   777
  \noindent Of course we also have to provide proper code equations for
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   778
  the operations, e.g. @{term "op + \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"}:
26973
6d52187fc2a6 temporary adjustment
haftmann
parents: 26732
diff changeset
   779
*}
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   780
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   781
lemma plus_Dig [code func]:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   782
  "0 + n = n"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   783
  "m + 0 = m"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   784
  "1 + Dig0 n = Dig1 n"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   785
  "Dig0 m + 1 = Dig1 m"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   786
  "1 + Dig1 n = Dig0 (n + 1)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   787
  "Dig1 m + 1 = Dig0 (m + 1)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   788
  "Dig0 m + Dig0 n = Dig0 (m + n)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   789
  "Dig0 m + Dig1 n = Dig1 (m + n)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   790
  "Dig1 m + Dig0 n = Dig1 (m + n)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   791
  "Dig1 m + Dig1 n = Dig0 (m + n + 1)"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   792
  by (simp_all add: Dig0_def Dig1_def)
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   793
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   794
text {*
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   795
  \noindent We then instruct the code generator to view @{term "0\<Colon>nat"},
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   796
  @{term "1\<Colon>nat"}, @{term Dig0} and @{term Dig1} as
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   797
  datatype constructors:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   798
*}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   799
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   800
code_datatype "0\<Colon>nat" "1\<Colon>nat" Dig0 Dig1
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   801
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   802
text {*
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   803
  \noindent For the former constructor @{term Suc}, we provide a code
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   804
  equation and remove some parts of the default code generator setup
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   805
  which are an obstacle here:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   806
*}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   807
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   808
lemma Suc_Dig [code func]:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   809
  "Suc n = n + 1"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   810
  by simp
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   811
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   812
declare One_nat_def [code inline del]
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   813
declare add_Suc_shift [code func del] 
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   814
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   815
text {*
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   816
  \noindent This yields the following code:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   817
*}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   818
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   819
export_code "op + \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat" (*<*)in SML(*>*) in SML file "examples/nat_binary.ML"
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   820
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   821
text {* \lstsml{Thy/examples/nat_binary.ML} *}
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   822
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   823
text {*
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   824
  \medskip
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   825
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   826
  From this example, it can be easily glimpsed that using own constructor sets
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   827
  is a little delicate since it changes the set of valid patterns for values
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   828
  of that type.  Without going into much detail, here some practical hints:
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   829
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   830
  \begin{itemize}
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   831
    \item When changing the constuctor set for datatypes, take care to
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   832
      provide an alternative for the @{text case} combinator (e.g. by replacing
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   833
      it using the preprocessor).
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   834
    \item Values in the target language need not to be normalized -- different
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   835
      values in the target language may represent the same value in the
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   836
      logic (e.g. @{text "Dig1 0 = 1"}).
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   837
    \item Usually, a good methodology to deal with the subleties of pattern
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   838
      matching is to see the type as an abstract type: provide a set
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   839
      of operations which operate on the concrete representation of the type,
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   840
      and derive further operations by combinations of these primitive ones,
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   841
      without relying on a particular representation.
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   842
  \end{itemize}
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   843
*}
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   844
26999
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   845
code_datatype %invisible "0::nat" Suc
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   846
declare %invisible plus_Dig [code func del]
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   847
declare %invisible One_nat_def [code inline]
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   848
declare %invisible add_Suc_shift [code func] 
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   849
lemma %invisible [code func]: "0 + n = (n \<Colon> nat)" by simp
284c871d3acb added new code_datatype example
haftmann
parents: 26973
diff changeset
   850
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   851
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
   852
subsection {* Customizing serialization  *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
   853
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   854
subsubsection {* Basics *}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   855
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   856
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   857
  Consider the following function and its corresponding
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   858
  SML code:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   859
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   860
25870
haftmann
parents: 25533
diff changeset
   861
primrec
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   862
  in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   863
  "in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l"
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   864
(*<*)
21323
haftmann
parents: 21223
diff changeset
   865
code_type %tt bool
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   866
  (SML)
21323
haftmann
parents: 21223
diff changeset
   867
code_const %tt True and False and "op \<and>" and Not
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   868
  (SML and and and)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   869
(*>*)
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   870
export_code in_interval (*<*)in SML(*>*)in SML file "examples/bool_literal.ML"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   871
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   872
text {*
21323
haftmann
parents: 21223
diff changeset
   873
  \lstsml{Thy/examples/bool_literal.ML}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   874
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   875
  \noindent Though this is correct code, it is a little bit unsatisfactory:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   876
  boolean values and operators are materialized as distinguished
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   877
  entities with have nothing to do with the SML-builtin notion
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   878
  of \qt{bool}.  This results in less readable code;
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   879
  additionally, eager evaluation may cause programs to
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   880
  loop or break which would perfectly terminate when
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   881
  the existing SML \qt{bool} would be used.  To map
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   882
  the HOL \qt{bool} on SML \qt{bool}, we may use
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   883
  \qn{custom serializations}:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   884
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   885
21323
haftmann
parents: 21223
diff changeset
   886
code_type %tt bool
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   887
  (SML "bool")
21323
haftmann
parents: 21223
diff changeset
   888
code_const %tt True and False and "op \<and>"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   889
  (SML "true" and "false" and "_ andalso _")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   890
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   891
text {*
21323
haftmann
parents: 21223
diff changeset
   892
  The @{text "\<CODETYPE>"} commad takes a type constructor
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   893
  as arguments together with a list of custom serializations.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   894
  Each custom serialization starts with a target language
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   895
  identifier followed by an expression, which during
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   896
  code serialization is inserted whenever the type constructor
21323
haftmann
parents: 21223
diff changeset
   897
  would occur.  For constants, @{text "\<CODECONST>"} implements
haftmann
parents: 21223
diff changeset
   898
  the corresponding mechanism.  Each ``@{verbatim "_"}'' in
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   899
  a serialization expression is treated as a placeholder
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   900
  for the type constructor's (the constant's) arguments.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   901
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   902
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   903
export_code in_interval (*<*)in SML(*>*) in SML file "examples/bool_mlbool.ML"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   904
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   905
text {*
21323
haftmann
parents: 21223
diff changeset
   906
  \lstsml{Thy/examples/bool_mlbool.ML}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   907
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   908
  \noindent This still is not perfect: the parentheses
21323
haftmann
parents: 21223
diff changeset
   909
  around the \qt{andalso} expression are superfluous.
haftmann
parents: 21223
diff changeset
   910
  Though the serializer
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   911
  by no means attempts to imitate the rich Isabelle syntax
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   912
  framework, it provides some common idioms, notably
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   913
  associative infixes with precedences which may be used here:
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   914
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   915
21323
haftmann
parents: 21223
diff changeset
   916
code_const %tt "op \<and>"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   917
  (SML infixl 1 "andalso")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   918
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
   919
export_code in_interval (*<*)in SML(*>*) in SML file "examples/bool_infix.ML"
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   920
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   921
text {*
21323
haftmann
parents: 21223
diff changeset
   922
  \lstsml{Thy/examples/bool_infix.ML}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   923
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   924
  \medskip
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   925
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   926
  Next, we try to map HOL pairs to SML pairs, using the
21323
haftmann
parents: 21223
diff changeset
   927
  infix ``@{verbatim "*"}'' type constructor and parentheses:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   928
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   929
(*<*)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   930
code_type *
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   931
  (SML)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   932
code_const Pair
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   933
  (SML)
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   934
(*>*)
21323
haftmann
parents: 21223
diff changeset
   935
code_type %tt *
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   936
  (SML infix 2 "*")
21323
haftmann
parents: 21223
diff changeset
   937
code_const %tt Pair
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   938
  (SML "!((_),/ (_))")
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   939
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   940
text {*
21323
haftmann
parents: 21223
diff changeset
   941
  The initial bang ``@{verbatim "!"}'' tells the serializer to never put
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   942
  parentheses around the whole expression (they are already present),
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   943
  while the parentheses around argument place holders
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   944
  tell not to put parentheses around the arguments.
21323
haftmann
parents: 21223
diff changeset
   945
  The slash ``@{verbatim "/"}'' (followed by arbitrary white space)
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   946
  inserts a space which may be used as a break if necessary
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   947
  during pretty printing.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   948
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   949
  These examples give a glimpse what mechanisms
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   950
  custom serializations provide; however their usage
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   951
  requires careful thinking in order not to introduce
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   952
  inconsistencies -- or, in other words:
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   953
  custom serializations are completely axiomatic.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   954
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   955
  A further noteworthy details is that any special
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   956
  character in a custom serialization may be quoted
21323
haftmann
parents: 21223
diff changeset
   957
  using ``@{verbatim "'"}''; thus, in
haftmann
parents: 21223
diff changeset
   958
  ``@{verbatim "fn '_ => _"}'' the first
haftmann
parents: 21223
diff changeset
   959
  ``@{verbatim "_"}'' is a proper underscore while the
haftmann
parents: 21223
diff changeset
   960
  second ``@{verbatim "_"}'' is a placeholder.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   961
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   962
  The HOL theories provide further
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
   963
  examples for custom serializations.
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   964
*}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   965
22188
a63889770d57 adjusted manual to improved treatment of overloaded constants
haftmann
parents: 22175
diff changeset
   966
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   967
subsubsection {* Haskell serialization *}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   968
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   969
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   970
  For convenience, the default
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   971
  HOL setup for Haskell maps the @{text eq} class to
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   972
  its counterpart in Haskell, giving custom serializations
21323
haftmann
parents: 21223
diff changeset
   973
  for the class (@{text "\<CODECLASS>"}) and its operation:
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   974
*}
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   975
21323
haftmann
parents: 21223
diff changeset
   976
code_class %tt eq
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   977
  (Haskell "Eq" where "op =" \<equiv> "(==)")
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   978
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
   979
code_const %tt "op ="
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   980
  (Haskell infixl 4 "==")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   981
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   982
text {*
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   983
  A problem now occurs whenever a type which
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   984
  is an instance of @{text eq} in HOL is mapped
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   985
  on a Haskell-builtin type which is also an instance
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   986
  of Haskell @{text Eq}:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   987
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
   988
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   989
typedecl bar
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   990
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   991
instantiation bar :: eq
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   992
begin
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   993
26732
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
   994
definition "eq_class.eq (x\<Colon>bar) y \<longleftrightarrow> x = y"
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   995
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   996
instance by default (simp add: eq_bar_def)
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   997
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25870
diff changeset
   998
end
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
   999
21323
haftmann
parents: 21223
diff changeset
  1000
code_type %tt bar
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1001
  (Haskell "Integer")
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1002
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1003
text {*
22188
a63889770d57 adjusted manual to improved treatment of overloaded constants
haftmann
parents: 22175
diff changeset
  1004
  The code generator would produce
a63889770d57 adjusted manual to improved treatment of overloaded constants
haftmann
parents: 22175
diff changeset
  1005
  an additional instance, which of course is rejected.
a63889770d57 adjusted manual to improved treatment of overloaded constants
haftmann
parents: 22175
diff changeset
  1006
  To suppress this additional instance, use
a63889770d57 adjusted manual to improved treatment of overloaded constants
haftmann
parents: 22175
diff changeset
  1007
  @{text "\<CODEINSTANCE>"}:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1008
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1009
21323
haftmann
parents: 21223
diff changeset
  1010
code_instance %tt bar :: eq
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1011
  (Haskell -)
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1012
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1013
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1014
subsubsection {* Pretty printing *}
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1015
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1016
text {*
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1017
  The serializer provides ML interfaces to set up
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1018
  pretty serializations for expressions like lists, numerals
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1019
  and characters;  these are
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1020
  monolithic stubs and should only be used with the
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1021
  theories introduces in \secref{sec:library}.
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1022
*}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1023
22550
c5039bee2602 updated
haftmann
parents: 22479
diff changeset
  1024
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1025
subsection {* Cyclic module dependencies *}
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1026
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1027
text {*
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1028
  Sometimes the awkward situation occurs that dependencies
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1029
  between definitions introduce cyclic dependencies
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1030
  between modules, which in the Haskell world leaves
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1031
  you to the mercy of the Haskell implementation you are using,
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1032
  while for SML code generation is not possible.
21178
c3618fc6a6f7 added gfx
haftmann
parents: 21147
diff changeset
  1033
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1034
  A solution is to declare module names explicitly.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1035
  Let use assume the three cyclically dependent
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1036
  modules are named \emph{A}, \emph{B} and \emph{C}.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1037
  Then, by stating
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1038
*}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1039
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1040
code_modulename SML
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1041
  A ABC
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1042
  B ABC
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1043
  C ABC
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1044
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1045
text {*
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1046
  we explicitly map all those modules on \emph{ABC},
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1047
  resulting in an ad-hoc merge of this three modules
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1048
  at serialization time.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1049
*}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1050
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1051
subsection {* Incremental code generation *}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1052
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1053
text {*
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1054
  Code generation is \emph{incremental}: theorems
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1055
  and abstract intermediate code are cached and extended on demand.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1056
  The cache may be partially or fully dropped if the underlying
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1057
  executable content of the theory changes.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1058
  Implementation of caching is supposed to transparently
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1059
  hid away the details from the user.  Anyway, caching
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1060
  reaches the surface by using a slightly more general form
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1061
  of the @{text "\<CODETHMS>"}, @{text "\<CODEDEPS>"}
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 24279
diff changeset
  1062
  and @{text "\<EXPORTCODE>"} commands: the list of constants
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1063
  may be omitted.  Then, all constants with code theorems
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1064
  in the current cache are referred to.
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1065
*}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1066
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
  1067
(*subsection {* Code generation and proof extraction *}
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1068
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1069
text {*
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1070
  \fixme
25411
ac31c92e4bf5 updated
haftmann
parents: 25369
diff changeset
  1071
*}*)
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1072
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1073
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
  1074
section {* ML interfaces \label{sec:ml} *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1075
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1076
text {*
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1077
  Since the code generator framework not only aims to provide
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1078
  a nice Isar interface but also to form a base for
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1079
  code-generation-based applications, here a short
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1080
  description of the most important ML interfaces.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1081
*}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1082
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1083
subsection {* Executable theory content: @{text Code} *}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1084
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1085
text {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1086
  This Pure module implements the core notions of
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1087
  executable content of a theory.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1088
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1089
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
  1090
subsubsection {* Managing executable content *}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1091
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1092
text %mlref {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1093
  \begin{mldecls}
25369
5200374fda5d replaced @{const} (allows name only) by proper @{term};
wenzelm
parents: 24994
diff changeset
  1094
  @{index_ML Code.add_func: "thm -> theory -> theory"} \\
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1095
  @{index_ML Code.del_func: "thm -> theory -> theory"} \\
24421
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1096
  @{index_ML Code.add_funcl: "string * thm list Susp.T -> theory -> theory"} \\
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1097
  @{index_ML Code.map_pre: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1098
  @{index_ML Code.map_post: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1099
  @{index_ML Code.add_functrans: "string * (theory -> thm list -> thm list)
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1100
    -> theory -> theory"} \\
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1101
  @{index_ML Code.del_functrans: "string -> theory -> theory"} \\
24421
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1102
  @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1103
  @{index_ML Code.get_datatype: "theory -> string
22479
de15ea8fb348 updated code generation sections
haftmann
parents: 22473
diff changeset
  1104
    -> (string * sort) list * (string * typ list) list"} \\
24421
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1105
  @{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1106
  \end{mldecls}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1107
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1108
  \begin{description}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1109
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1110
  \item @{ML Code.add_func}~@{text "thm"}~@{text "thy"} adds function
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1111
     theorem @{text "thm"} to executable content.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1112
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1113
  \item @{ML Code.del_func}~@{text "thm"}~@{text "thy"} removes function
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1114
     theorem @{text "thm"} from executable content, if present.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1115
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1116
  \item @{ML Code.add_funcl}~@{text "(const, lthms)"}~@{text "thy"} adds
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
  1117
     suspended defining equations @{text lthms} for constant
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1118
     @{text const} to executable content.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1119
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1120
  \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1121
     the preprocessor simpset.
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1122
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1123
    \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1124
     function transformator @{text f} (named @{text name}) to executable content;
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
  1125
     @{text f} is a transformation of the defining equations belonging
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1126
     to a certain function definition, depending on the
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1127
     current theory context.
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1128
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1129
  \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1130
     function transformator named @{text name} from executable content.
22046
ce84c9887e2d named preprocessors
haftmann
parents: 22015
diff changeset
  1131
24421
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1132
  \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1133
     a datatype to executable content, with generation
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1134
     set @{text cs}.
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1135
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1136
  \item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"}
21189
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1137
     returns type constructor corresponding to
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1138
     constructor @{text const}; returns @{text NONE}
5435ccdb4ea1 (continued)
haftmann
parents: 21178
diff changeset
  1139
     if @{text const} is no constructor.
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1140
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1141
  \end{description}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1142
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1143
22292
3b118010ec08 adjusted to new code generator Isar commands and changes in implementation
haftmann
parents: 22188
diff changeset
  1144
subsection {* Auxiliary *}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1145
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1146
text %mlref {*
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1147
  \begin{mldecls}
24421
acfb2413faa3 updated
haftmann
parents: 24348
diff changeset
  1148
  @{index_ML CodeUnit.read_const: "theory -> string -> string"} \\
26973
6d52187fc2a6 temporary adjustment
haftmann
parents: 26732
diff changeset
  1149
  @{index_ML CodeUnit.head_func: "thm -> string * ((string * sort) list * typ)"} \\
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1150
  @{index_ML CodeUnit.rewrite_func: "MetaSimplifier.simpset -> thm -> thm"} \\
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1151
  \end{mldecls}
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1152
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1153
  \begin{description}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1154
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1155
  \item @{ML CodeUnit.read_const}~@{text thy}~@{text s}
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1156
     reads a constant as a concrete term expression @{text s}.
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1157
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1158
  \item @{ML CodeUnit.head_func}~@{text thm}
22751
1bfd75c1f232 updated
haftmann
parents: 22550
diff changeset
  1159
     extracts the constant and its type from a defining equation @{text thm}.
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1160
27557
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1161
  \item @{ML CodeUnit.rewrite_func}~@{text ss}~@{text thm}
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1162
     rewrites a defining equation @{text thm} with a simpset @{text ss};
151731493264 simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents: 27103
diff changeset
  1163
     only arguments and right hand side are rewritten,
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
  1164
     not the head of the defining equation.
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1165
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1166
  \end{description}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1167
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1168
*}
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1169
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1170
subsection {* Implementing code generator applications *}
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1171
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1172
text {*
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1173
  Implementing code generator applications on top
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1174
  of the framework set out so far usually not only
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1175
  involves using those primitive interfaces
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1176
  but also storing code-dependent data and various
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1177
  other things.
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1178
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1179
  \begin{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1180
    Some interfaces discussed here have not reached
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1181
    a final state yet.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1182
    Changes likely to occur in future.
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1183
  \end{warn}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1184
*}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1185
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1186
subsubsection {* Data depending on the theory's executable content *}
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1187
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1188
text {*
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1189
  Due to incrementality of code generation, changes in the
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1190
  theory's executable content have to be propagated in a
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1191
  certain fashion.  Additionally, such changes may occur
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1192
  not only during theory extension but also during theory
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1193
  merge, which is a little bit nasty from an implementation
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1194
  point of view.  The framework provides a solution
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1195
  to this technical challenge by providing a functorial
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1196
  data slot @{ML_functor CodeDataFun}; on instantiation
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1197
  of this functor, the following types and operations
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1198
  are required:
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1199
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1200
  \medskip
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1201
  \begin{tabular}{l}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1202
  @{text "type T"} \\
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1203
  @{text "val empty: T"} \\
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1204
  @{text "val merge: Pretty.pp \<rightarrow> T * T \<rightarrow> T"} \\
24217
5c4913b478f5 updated
haftmann
parents: 24193
diff changeset
  1205
  @{text "val purge: theory option \<rightarrow> CodeUnit.const list option \<rightarrow> T \<rightarrow> T"}
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1206
  \end{tabular}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1207
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1208
  \begin{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1209
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1210
  \item @{text T} the type of data to store.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1211
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1212
  \item @{text empty} initial (empty) data.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1213
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1214
  \item @{text merge} merging two data slots.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1215
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1216
  \item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content;
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1217
    if possible, the current theory context is handed over
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1218
    as argument @{text thy} (if there is no current theory context (e.g.~during
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1219
    theory merge, @{ML NONE}); @{text consts} indicates the kind
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1220
    of change: @{ML NONE} stands for a fundamental change
22798
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1221
    which invalidates any existing code, @{text "SOME consts"}
e3962371f568 updated doc
haftmann
parents: 22751
diff changeset
  1222
    hints that executable content for constants @{text consts}
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1223
    has changed.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1224
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1225
  \end{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1226
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1227
  An instance of @{ML_functor CodeDataFun} provides the following
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1228
  interface:
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1229
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1230
  \medskip
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1231
  \begin{tabular}{l}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1232
  @{text "get: theory \<rightarrow> T"} \\
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1233
  @{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1234
  @{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1235
  \end{tabular}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1236
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1237
  \begin{description}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1238
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1239
  \item @{text get} retrieval of the current data.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1240
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1241
  \item @{text change} update of current data (cached!)
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1242
    by giving a continuation.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1243
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1244
  \item @{text change_yield} update with side result.
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1245
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1246
  \end{description}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1247
*}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1248
24628
33137422d7fd updated
haftmann
parents: 24421
diff changeset
  1249
(*subsubsection {* Datatype hooks *}
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1250
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1251
text {*
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1252
  Isabelle/HOL's datatype package provides a mechanism to
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1253
  extend theories depending on datatype declarations:
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1254
  \emph{datatype hooks}.  For example, when declaring a new
22060
8a37090726e8 adjusted manual
haftmann
parents: 22046
diff changeset
  1255
  datatype, a hook proves defining equations for equality on
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1256
  that datatype (if possible).
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1257
*}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1258
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1259
text %mlref {*
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1260
  \begin{mldecls}
21323
haftmann
parents: 21223
diff changeset
  1261
  @{index_ML_type DatatypeHooks.hook: "string list -> theory -> theory"} \\
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1262
  @{index_ML DatatypeHooks.add: "DatatypeHooks.hook -> theory -> theory"}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1263
  \end{mldecls}
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1264
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1265
  \begin{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1266
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1267
  \item @{ML_type DatatypeHooks.hook} specifies the interface
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1268
     of \emph{datatype hooks}: a theory update
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1269
     depending on the list of newly introduced
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1270
     datatype names.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1271
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1272
  \item @{ML DatatypeHooks.add} adds a hook to the
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1273
     chain of all hooks.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1274
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1275
  \end{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1276
*}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1277
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1278
subsubsection {* Trivial typedefs -- type copies *}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1279
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1280
text {*
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1281
  Sometimes packages will introduce new types
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1282
  as \emph{marked type copies} similar to Haskell's
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1283
  @{text newtype} declaration (e.g. the HOL record package)
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1284
  \emph{without} tinkering with the overhead of datatypes.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1285
  Technically, these type copies are trivial forms of typedefs.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1286
  Since these type copies in code generation view are nothing
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1287
  else than datatypes, they have been given a own package
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1288
  in order to faciliate code generation:
21147
737a94f047e3 continued tutorial
haftmann
parents: 21089
diff changeset
  1289
*}
21058
a32d357dfd70 started tutorial
haftmann
parents: 20948
diff changeset
  1290
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1291
text %mlref {*
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1292
  \begin{mldecls}
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1293
  @{index_ML_type TypecopyPackage.info} \\
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1294
  @{index_ML TypecopyPackage.add_typecopy: "
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1295
    bstring * string list -> typ -> (bstring * bstring) option
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1296
    -> theory -> (string * TypecopyPackage.info) * theory"} \\
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1297
  @{index_ML TypecopyPackage.get_typecopy_info: "theory
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1298
    -> string -> TypecopyPackage.info option"} \\
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1299
  @{index_ML TypecopyPackage.get_spec: "theory -> string
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1300
    -> (string * sort) list * (string * typ list) list"} \\
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1301
  @{index_ML_type TypecopyPackage.hook: "string * TypecopyPackage.info -> theory -> theory"} \\
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1302
  @{index_ML TypecopyPackage.add_hook:
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1303
     "TypecopyPackage.hook -> theory -> theory"} \\
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1304
  \end{mldecls}
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1305
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1306
  \begin{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1307
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1308
  \item @{ML_type TypecopyPackage.info} a record containing
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1309
     the specification and further data of a type copy.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1310
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1311
  \item @{ML TypecopyPackage.add_typecopy} defines a new
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1312
     type copy.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1313
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1314
  \item @{ML TypecopyPackage.get_typecopy_info} retrieves
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1315
     data of an existing type copy.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1316
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1317
  \item @{ML TypecopyPackage.get_spec} retrieves datatype-like
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1318
     specification of a type copy.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1319
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1320
  \item @{ML_type TypecopyPackage.hook},~@{ML TypecopyPackage.add_hook}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1321
     provide a hook mechanism corresponding to the hook mechanism
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1322
     on datatypes.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1323
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1324
  \end{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1325
*}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1326
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1327
subsubsection {* Unifying type copies and datatypes *}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1328
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1329
text {*
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1330
  Since datatypes and type copies are mapped to the same concept (datatypes)
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1331
  by code generation, the view on both is unified \qt{code types}:
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1332
*}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1333
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1334
text %mlref {*
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1335
  \begin{mldecls}
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1336
  @{index_ML_type DatatypeCodegen.hook: "(string * (bool * ((string * sort) list
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1337
    * (string * typ list) list))) list
21323
haftmann
parents: 21223
diff changeset
  1338
    -> theory -> theory"} \\
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1339
  @{index_ML DatatypeCodegen.add_codetypes_hook_bootstrap: "
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1340
      DatatypeCodegen.hook -> theory -> theory"}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1341
  \end{mldecls}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1342
*}
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1343
21222
6dfdb69e83ef adjusted title
haftmann
parents: 21217
diff changeset
  1344
text {*
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1345
  \begin{description}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1346
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1347
  \item @{ML_type DatatypeCodegen.hook} specifies the code type hook
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1348
     interface: a theory transformation depending on a list of
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1349
     mutual recursive code types; each entry in the list
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1350
     has the structure @{text "(name, (is_data, (vars, cons)))"}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1351
     where @{text name} is the name of the code type, @{text is_data}
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1352
     is true iff @{text name} is a datatype rather then a type copy,
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1353
     and @{text "(vars, cons)"} is the specification of the code type.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1354
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1355
  \item @{ML DatatypeCodegen.add_codetypes_hook_bootstrap} adds a code
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1356
     type hook;  the hook is immediately processed for all already
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1357
     existing datatypes, in blocks of mutual recursive datatypes,
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1358
     where all datatypes a block depends on are processed before
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1359
     the block.
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1360
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1361
  \end{description}
24628
33137422d7fd updated
haftmann
parents: 24421
diff changeset
  1362
*}*)
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1363
24628
33137422d7fd updated
haftmann
parents: 24421
diff changeset
  1364
text {*
21452
f825e0b4d566 final draft
haftmann
parents: 21341
diff changeset
  1365
  \emph{Happy proving, happy hacking!}
21222
6dfdb69e83ef adjusted title
haftmann
parents: 21217
diff changeset
  1366
*}
21217
0425fc57510f continued
haftmann
parents: 21189
diff changeset
  1367
20948
9b9910b82645 initial draft
haftmann
parents:
diff changeset
  1368
end