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