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