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