doc-src/Codegen/Thy/Introduction.thy
author haftmann
Sat, 27 Nov 2010 18:51:04 +0100
changeset 40991 5288144b4358
parent 39978 3aa2bc9c5478
child 42960 9f6652122963
permissions -rw-r--r--
added evaluation section
haftmann@28213
     1
theory Introduction
haftmann@28213
     2
imports Setup
haftmann@28213
     3
begin
haftmann@28213
     4
haftmann@38631
     5
section {* Introduction *}
haftmann@38631
     6
haftmann@38670
     7
text {*
haftmann@38670
     8
  This tutorial introduces the code generator facilities of @{text
haftmann@38670
     9
  "Isabelle/HOL"}.  It allows to turn (a certain class of) HOL
haftmann@38670
    10
  specifications into corresponding executable code in the programming
haftmann@39049
    11
  languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml},
haftmann@39049
    12
  @{text Haskell} \cite{haskell-revised-report} and @{text Scala}
haftmann@39049
    13
  \cite{scala-overview-tech-report}.
haftmann@38631
    14
haftmann@38670
    15
  To profit from this tutorial, some familiarity and experience with
haftmann@38670
    16
  @{theory HOL} \cite{isa-tutorial} and its basic theories is assumed.
haftmann@38670
    17
*}
haftmann@38631
    18
haftmann@38631
    19
haftmann@38670
    20
subsection {* Code generation principle: shallow embedding \label{sec:principle} *}
haftmann@28213
    21
haftmann@28213
    22
text {*
haftmann@38670
    23
  The key concept for understanding Isabelle's code generation is
haftmann@38670
    24
  \emph{shallow embedding}: logical entities like constants, types and
haftmann@38670
    25
  classes are identified with corresponding entities in the target
haftmann@38670
    26
  language.  In particular, the carrier of a generated program's
haftmann@38670
    27
  semantics are \emph{equational theorems} from the logic.  If we view
haftmann@38670
    28
  a generated program as an implementation of a higher-order rewrite
haftmann@38670
    29
  system, then every rewrite step performed by the program can be
haftmann@38670
    30
  simulated in the logic, which guarantees partial correctness
haftmann@38670
    31
  \cite{Haftmann-Nipkow:2010:code}.
haftmann@28213
    32
*}
haftmann@28213
    33
haftmann@38670
    34
haftmann@38670
    35
subsection {* A quick start with the Isabelle/HOL toolbox \label{sec:queue_example} *}
haftmann@28213
    36
haftmann@28419
    37
text {*
haftmann@38730
    38
  In a HOL theory, the @{command_def datatype} and @{command_def
haftmann@38730
    39
  definition}/@{command_def primrec}/@{command_def fun} declarations
haftmann@38730
    40
  form the core of a functional programming language.  By default
haftmann@38730
    41
  equational theorems stemming from those are used for generated code,
haftmann@38730
    42
  therefore \qt{naive} code generation can proceed without further
haftmann@38730
    43
  ado.
haftmann@28419
    44
haftmann@28419
    45
  For example, here a simple \qt{implementation} of amortised queues:
haftmann@28419
    46
*}
haftmann@28419
    47
haftmann@29735
    48
datatype %quote 'a queue = AQueue "'a list" "'a list"
haftmann@28419
    49
haftmann@28564
    50
definition %quote empty :: "'a queue" where
haftmann@29735
    51
  "empty = AQueue [] []"
haftmann@28419
    52
haftmann@28564
    53
primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
haftmann@29735
    54
  "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
haftmann@28419
    55
haftmann@28564
    56
fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
haftmann@29735
    57
    "dequeue (AQueue [] []) = (None, AQueue [] [])"
haftmann@29735
    58
  | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
haftmann@29735
    59
  | "dequeue (AQueue xs []) =
haftmann@38670
    60
      (case rev xs of y # ys \<Rightarrow> (Some y, AQueue [] ys))" (*<*)
haftmann@38670
    61
haftmann@38670
    62
lemma %invisible dequeue_nonempty_Nil [simp]:
haftmann@38670
    63
  "xs \<noteq> [] \<Longrightarrow> dequeue (AQueue xs []) = (case rev xs of y # ys \<Rightarrow> (Some y, AQueue [] ys))"
haftmann@38670
    64
  by (cases xs) (simp_all split: list.splits) (*>*)
haftmann@28419
    65
haftmann@28419
    66
text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *}
haftmann@28419
    67
haftmann@28564
    68
export_code %quote empty dequeue enqueue in SML
haftmann@28447
    69
  module_name Example file "examples/example.ML"
haftmann@28419
    70
haftmann@28419
    71
text {* \noindent resulting in the following code: *}
haftmann@28419
    72
haftmann@39978
    73
text %quotetypewriter {*
haftmann@39906
    74
  @{code_stmts empty enqueue dequeue (SML)}
haftmann@39888
    75
*}
haftmann@28419
    76
haftmann@28419
    77
text {*
haftmann@38730
    78
  \noindent The @{command_def export_code} command takes a
haftmann@38730
    79
  space-separated list of constants for which code shall be generated;
haftmann@38730
    80
  anything else needed for those is added implicitly.  Then follows a
haftmann@38730
    81
  target language identifier and a freely chosen module name.  A file
haftmann@38730
    82
  name denotes the destination to store the generated code.  Note that
haftmann@38730
    83
  the semantics of the destination depends on the target language: for
haftmann@39049
    84
  @{text SML}, @{text OCaml} and @{text Scala} it denotes a \emph{file},
haftmann@39049
    85
  for @{text Haskell} it denotes a \emph{directory} where a file named as the
haftmann@38670
    86
  module name (with extension @{text ".hs"}) is written:
haftmann@28419
    87
*}
haftmann@28419
    88
haftmann@28564
    89
export_code %quote empty dequeue enqueue in Haskell
haftmann@28447
    90
  module_name Example file "examples/"
haftmann@28419
    91
haftmann@28419
    92
text {*
haftmann@38670
    93
  \noindent This is the corresponding code:
haftmann@28419
    94
*}
haftmann@28419
    95
haftmann@39978
    96
text %quotetypewriter {*
haftmann@39906
    97
  @{code_stmts empty enqueue dequeue (Haskell)}
haftmann@39888
    98
*}
haftmann@28419
    99
haftmann@28419
   100
text {*
haftmann@38670
   101
  \noindent For more details about @{command export_code} see
haftmann@38670
   102
  \secref{sec:further}.
haftmann@28419
   103
*}
haftmann@28213
   104
haftmann@38670
   105
haftmann@38670
   106
subsection {* Type classes *}
haftmann@38628
   107
haftmann@38628
   108
text {*
haftmann@38670
   109
  Code can also be generated from type classes in a Haskell-like
haftmann@38670
   110
  manner.  For illustration here an example from abstract algebra:
haftmann@38628
   111
*}
haftmann@38628
   112
haftmann@38670
   113
class %quote semigroup =
haftmann@38670
   114
  fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<otimes>" 70)
haftmann@38670
   115
  assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
haftmann@38670
   116
haftmann@38670
   117
class %quote monoid = semigroup +
haftmann@38670
   118
  fixes neutral :: 'a ("\<one>")
haftmann@38670
   119
  assumes neutl: "\<one> \<otimes> x = x"
haftmann@38670
   120
    and neutr: "x \<otimes> \<one> = x"
haftmann@38670
   121
haftmann@38670
   122
instantiation %quote nat :: monoid
haftmann@38670
   123
begin
haftmann@38670
   124
haftmann@38670
   125
primrec %quote mult_nat where
haftmann@38670
   126
    "0 \<otimes> n = (0\<Colon>nat)"
haftmann@38670
   127
  | "Suc m \<otimes> n = n + m \<otimes> n"
haftmann@38670
   128
haftmann@38670
   129
definition %quote neutral_nat where
haftmann@38670
   130
  "\<one> = Suc 0"
haftmann@38670
   131
haftmann@38670
   132
lemma %quote add_mult_distrib:
haftmann@38670
   133
  fixes n m q :: nat
haftmann@38670
   134
  shows "(n + m) \<otimes> q = n \<otimes> q + m \<otimes> q"
haftmann@38670
   135
  by (induct n) simp_all
haftmann@38670
   136
haftmann@38670
   137
instance %quote proof
haftmann@38670
   138
  fix m n q :: nat
haftmann@38670
   139
  show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)"
haftmann@38670
   140
    by (induct m) (simp_all add: add_mult_distrib)
haftmann@38670
   141
  show "\<one> \<otimes> n = n"
haftmann@38670
   142
    by (simp add: neutral_nat_def)
haftmann@38670
   143
  show "m \<otimes> \<one> = m"
haftmann@38670
   144
    by (induct m) (simp_all add: neutral_nat_def)
haftmann@38670
   145
qed
haftmann@38670
   146
haftmann@38670
   147
end %quote
haftmann@28213
   148
haftmann@28419
   149
text {*
haftmann@38670
   150
  \noindent We define the natural operation of the natural numbers
haftmann@38670
   151
  on monoids:
haftmann@38670
   152
*}
haftmann@28419
   153
haftmann@38670
   154
primrec %quote (in monoid) pow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" where
haftmann@38670
   155
    "pow 0 a = \<one>"
haftmann@38670
   156
  | "pow (Suc n) a = a \<otimes> pow n a"
haftmann@28419
   157
haftmann@38670
   158
text {*
haftmann@38670
   159
  \noindent This we use to define the discrete exponentiation
haftmann@38670
   160
  function:
haftmann@38670
   161
*}
haftmann@38670
   162
haftmann@38670
   163
definition %quote bexp :: "nat \<Rightarrow> nat" where
haftmann@38670
   164
  "bexp n = pow n (Suc (Suc 0))"
haftmann@38670
   165
haftmann@38670
   166
text {*
haftmann@38670
   167
  \noindent The corresponding code in Haskell uses that language's
haftmann@38670
   168
  native classes:
haftmann@38670
   169
*}
haftmann@38670
   170
haftmann@39978
   171
text %quotetypewriter {*
haftmann@39906
   172
  @{code_stmts bexp (Haskell)}
haftmann@39888
   173
*}
haftmann@38670
   174
haftmann@38670
   175
text {*
haftmann@38670
   176
  \noindent This is a convenient place to show how explicit dictionary
haftmann@38670
   177
  construction manifests in generated code -- the same example in
haftmann@38670
   178
  @{text SML}:
haftmann@38670
   179
*}
haftmann@38670
   180
haftmann@39978
   181
text %quotetypewriter {*
haftmann@39906
   182
  @{code_stmts bexp (SML)}
haftmann@39888
   183
*}
haftmann@38670
   184
haftmann@38670
   185
text {*
haftmann@38670
   186
  \noindent Note the parameters with trailing underscore (@{verbatim
haftmann@38670
   187
  "A_"}), which are the dictionary parameters.
haftmann@38670
   188
*}
haftmann@38670
   189
haftmann@38670
   190
haftmann@38670
   191
subsection {* How to continue from here *}
haftmann@38670
   192
haftmann@38670
   193
text {*
haftmann@38670
   194
  What you have seen so far should be already enough in a lot of
haftmann@38670
   195
  cases.  If you are content with this, you can quit reading here.
haftmann@38670
   196
haftmann@38670
   197
  Anyway, to understand situations where problems occur or to increase
haftmann@38670
   198
  the scope of code generation beyond default, it is necessary to gain
haftmann@38670
   199
  some understanding how the code generator actually works:
haftmann@28419
   200
haftmann@28419
   201
  \begin{itemize}
haftmann@28419
   202
haftmann@38670
   203
    \item The foundations of the code generator are described in
haftmann@38670
   204
      \secref{sec:foundations}.
haftmann@28419
   205
haftmann@38670
   206
    \item In particular \secref{sec:utterly_wrong} gives hints how to
haftmann@38670
   207
      debug situations where code generation does not succeed as
haftmann@38670
   208
      expected.
haftmann@28419
   209
haftmann@38670
   210
    \item The scope and quality of generated code can be increased
haftmann@38670
   211
      dramatically by applying refinement techniques, which are
haftmann@38670
   212
      introduced in \secref{sec:refinement}.
haftmann@38670
   213
haftmann@38670
   214
    \item Inductive predicates can be turned executable using an
haftmann@38670
   215
      extension of the code generator \secref{sec:inductive}.
haftmann@38670
   216
haftmann@40991
   217
    \item If you want to utilize code generation to obtain fast
haftmann@40991
   218
      evaluators e.g.~for decision procedures, have a look at
haftmann@40991
   219
      \secref{sec:evaluation}.
haftmann@40991
   220
haftmann@38670
   221
    \item You may want to skim over the more technical sections
haftmann@38670
   222
      \secref{sec:adaptation} and \secref{sec:further}.
haftmann@38670
   223
haftmann@38670
   224
    \item For exhaustive syntax diagrams etc. you should visit the
haftmann@38670
   225
      Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}.
haftmann@28419
   226
haftmann@28419
   227
  \end{itemize}
haftmann@28419
   228
haftmann@38670
   229
  \bigskip
haftmann@38670
   230
haftmann@38670
   231
  \begin{center}\fbox{\fbox{\begin{minipage}{8cm}
haftmann@38670
   232
haftmann@38670
   233
    \begin{center}\textit{Happy proving, happy hacking!}\end{center}
haftmann@38670
   234
haftmann@38670
   235
  \end{minipage}}}\end{center}
haftmann@38670
   236
haftmann@38670
   237
  \begin{warn}
haftmann@38670
   238
    There is also a more ancient code generator in Isabelle by Stefan
haftmann@38670
   239
    Berghofer \cite{Berghofer-Nipkow:2002}.  Although its
haftmann@38670
   240
    functionality is covered by the code generator presented here, it
haftmann@38670
   241
    will sometimes show up as an artifact.  In case of ambiguity, we
haftmann@38670
   242
    will refer to the framework described here as @{text "generic code
haftmann@38670
   243
    generator"}, to the other as @{text "SML code generator"}.
haftmann@38670
   244
  \end{warn}
haftmann@28419
   245
*}
haftmann@28213
   246
haftmann@28213
   247
end