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