| author | urbanc |
| Thu, 22 Mar 2007 11:17:32 +0100 | |
| changeset 22495 | c54748fd1f43 |
| parent 22479 | de15ea8fb348 |
| child 22550 | c5039bee2602 |
| permissions | -rw-r--r-- |
| 20948 | 1 |
(* $Id$ *) |
2 |
||
| 21147 | 3 |
(*<*) |
| 20948 | 4 |
theory Codegen |
5 |
imports Main |
|
| 21452 | 6 |
uses "../../../antiquote_setup.ML" |
| 20948 | 7 |
begin |
8 |
||
| 21147 | 9 |
ML {*
|
| 22015 | 10 |
CodegenSerializer.code_width := 74; |
| 21147 | 11 |
*} |
12 |
||
13 |
(*>*) |
|
14 |
||
| 20948 | 15 |
chapter {* Code generation from Isabelle theories *}
|
16 |
||
17 |
section {* Introduction *}
|
|
18 |
||
| 21058 | 19 |
subsection {* Motivation *}
|
20 |
||
| 20948 | 21 |
text {*
|
| 21058 | 22 |
Executing formal specifications as programs is a well-established |
23 |
topic in the theorem proving community. With increasing |
|
24 |
application of theorem proving systems in the area of |
|
25 |
software development and verification, its relevance manifests |
|
26 |
for running test cases and rapid prototyping. In logical |
|
27 |
calculi like constructive type theory, |
|
28 |
a notion of executability is implicit due to the nature |
|
29 |
of the calculus. In contrast, specifications in Isabelle/HOL |
|
30 |
can be highly non-executable. In order to bridge |
|
31 |
the gap between logic and executable specifications, |
|
32 |
an explicit non-trivial transformation has to be applied: |
|
33 |
code generation. |
|
34 |
||
35 |
This tutorial introduces a generic code generator for the |
|
36 |
Isabelle system \cite{isa-tutorial}.
|
|
37 |
Generic in the sense that the |
|
38 |
\qn{target language} for which code shall ultimately be
|
|
39 |
generated is not fixed but may be an arbitrary state-of-the-art |
|
40 |
functional programming language (currently, the implementation |
|
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
41 |
supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell \cite{haskell-revised-report}).
|
| 21058 | 42 |
We aim to provide a |
43 |
versatile environment |
|
44 |
suitable for software development and verification, |
|
45 |
structuring the process |
|
46 |
of code generation into a small set of orthogonal principles |
|
47 |
while achieving a big coverage of application areas |
|
48 |
with maximum flexibility. |
|
| 21189 | 49 |
|
50 |
For readers, some familiarity and experience |
|
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
51 |
with the ingredients |
| 21189 | 52 |
of the HOL \emph{Main} theory is assumed.
|
| 21058 | 53 |
*} |
54 |
||
55 |
||
56 |
subsection {* Overview *}
|
|
57 |
||
58 |
text {*
|
|
59 |
The code generator aims to be usable with no further ado |
|
60 |
in most cases while allowing for detailed customization. |
|
61 |
This manifests in the structure of this tutorial: this introduction |
|
| 21147 | 62 |
continues with a short introduction of concepts. Section |
| 21178 | 63 |
\secref{sec:basics} explains how to use the framework naively,
|
| 21058 | 64 |
presuming a reasonable default setup. Then, section |
65 |
\secref{sec:advanced} deals with advanced topics,
|
|
66 |
introducing further aspects of the code generator framework |
|
67 |
in a motivation-driven manner. Last, section \secref{sec:ml}
|
|
68 |
introduces the framework's internal programming interfaces. |
|
| 20948 | 69 |
|
| 21058 | 70 |
\begin{warn}
|
71 |
Ultimately, the code generator which this tutorial deals with |
|
72 |
is supposed to replace the already established code generator |
|
73 |
by Stefan Berghofer \cite{Berghofer-Nipkow:2002}.
|
|
| 21147 | 74 |
So, for the moment, there are two distinct code generators |
| 21058 | 75 |
in Isabelle. |
76 |
Also note that while the framework itself is largely |
|
77 |
object-logic independent, only HOL provides a reasonable |
|
78 |
framework setup. |
|
79 |
\end{warn}
|
|
80 |
*} |
|
81 |
||
82 |
||
| 22479 | 83 |
subsection {* An exmaple: a simple theory of search trees *}
|
84 |
||
85 |
datatype ('a, 'b) searchtree = Leaf "'a\<Colon>linorder" 'b
|
|
86 |
| Branch "('a, 'b) searchtree" "'a" "('a, 'b) searchtree"
|
|
87 |
||
88 |
fun |
|
89 |
find :: "('a\<Colon>linorder, 'b) searchtree \<Rightarrow> 'a \<Rightarrow> 'b option" where
|
|
90 |
"find (Leaf key val) it = (if it = key then Some val else None)" |
|
91 |
| "find (Branch t1 key t2) it = (if it \<le> key then find t1 it else find t2 it)" |
|
92 |
||
93 |
fun |
|
94 |
update :: "'a\<Colon>linorder \<times> 'b \<Rightarrow> ('a, 'b) searchtree \<Rightarrow> ('a, 'b) searchtree" where
|
|
95 |
"update (it, entry) (Leaf key val) = ( |
|
96 |
if it = key then Leaf key entry |
|
97 |
else if it \<le> key |
|
98 |
then Branch (Leaf it entry) it (Leaf key val) |
|
99 |
else Branch (Leaf key val) it (Leaf it entry) |
|
100 |
)" |
|
101 |
| "update (it, entry) (Branch t1 key t2) = ( |
|
102 |
if it \<le> key |
|
103 |
then (Branch (update (it, entry) t1) key t2) |
|
104 |
else (Branch t1 key (update (it, entry) t2)) |
|
105 |
)" |
|
106 |
||
107 |
fun |
|
108 |
example :: "nat \<Rightarrow> (nat, string) searchtree" where |
|
109 |
"example n = update (n, ''bar'') (Leaf 0 ''foo'')" |
|
110 |
||
111 |
code_gen example (*<*)(SML #)(*>*)(SML "examples/tree.ML") |
|
112 |
||
113 |
text {*
|
|
114 |
\lstsml{Thy/examples/tree.ML}
|
|
115 |
*} |
|
116 |
||
| 21058 | 117 |
subsection {* Code generation process *}
|
118 |
||
119 |
text {*
|
|
| 21189 | 120 |
\begin{figure}[h]
|
121 |
\centering |
|
122 |
\includegraphics[width=0.7\textwidth]{codegen_process}
|
|
123 |
\caption{code generator -- processing overview}
|
|
124 |
\label{fig:process}
|
|
125 |
\end{figure}
|
|
126 |
||
| 21058 | 127 |
The code generator employs a notion of executability |
128 |
for three foundational executable ingredients known |
|
129 |
from functional programming: |
|
| 22060 | 130 |
\emph{defining equations}, \emph{datatypes}, and
|
131 |
\emph{type classes}. A defining equation as a first approximation
|
|
| 21058 | 132 |
is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"}
|
133 |
(an equation headed by a constant @{text f} with arguments
|
|
134 |
@{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right hand side @{text t}.
|
|
| 22060 | 135 |
Code generation aims to turn defining equations |
| 21058 | 136 |
into a functional program by running through |
137 |
a process (see figure \ref{fig:process}):
|
|
138 |
||
139 |
\begin{itemize}
|
|
140 |
||
141 |
\item Out of the vast collection of theorems proven in a |
|
142 |
\qn{theory}, a reasonable subset modeling
|
|
| 22060 | 143 |
defining equations is \qn{selected}.
|
| 21058 | 144 |
|
145 |
\item On those selected theorems, certain |
|
146 |
transformations are carried out |
|
147 |
(\qn{preprocessing}). Their purpose is to turn theorems
|
|
148 |
representing non- or badly executable |
|
149 |
specifications into equivalent but executable counterparts. |
|
150 |
The result is a structured collection of \qn{code theorems}.
|
|
151 |
||
| 22479 | 152 |
\item These \qn{code theorems} then are \qn{translated}
|
| 21058 | 153 |
into an Haskell-like intermediate |
154 |
language. |
|
155 |
||
156 |
\item Finally, out of the intermediate language the final |
|
157 |
code in the desired \qn{target language} is \qn{serialized}.
|
|
158 |
||
159 |
\end{itemize}
|
|
160 |
||
161 |
From these steps, only the two last are carried out |
|
162 |
outside the logic; by keeping this layer as |
|
163 |
thin as possible, the amount of code to trust is |
|
164 |
kept to a minimum. |
|
165 |
*} |
|
166 |
||
167 |
||
168 |
||
169 |
section {* Basics \label{sec:basics} *}
|
|
| 20948 | 170 |
|
171 |
subsection {* Invoking the code generator *}
|
|
172 |
||
| 21058 | 173 |
text {*
|
174 |
Thanks to a reasonable setup of the HOL theories, in |
|
175 |
most cases code generation proceeds without further ado: |
|
176 |
*} |
|
177 |
||
| 22473 | 178 |
fun |
179 |
fac :: "nat \<Rightarrow> nat" where |
|
180 |
"fac 0 = 1" |
|
181 |
| "fac (Suc n) = Suc n * fac n" |
|
| 21058 | 182 |
|
183 |
text {*
|
|
184 |
This executable specification is now turned to SML code: |
|
185 |
*} |
|
186 |
||
| 21545 | 187 |
code_gen fac (*<*)(SML #)(*>*)(SML "examples/fac.ML") |
| 21058 | 188 |
|
189 |
text {*
|
|
| 21323 | 190 |
The @{text "\<CODEGEN>"} command takes a space-separated list of
|
| 21058 | 191 |
constants together with \qn{serialization directives}
|
192 |
in parentheses. These start with a \qn{target language}
|
|
| 21178 | 193 |
identifier, followed by arguments, their semantics |
| 21058 | 194 |
depending on the target. In the SML case, a filename |
195 |
is given where to write the generated code to. |
|
196 |
||
| 22060 | 197 |
Internally, the defining equations for all selected |
| 21178 | 198 |
constants are taken, including any transitively required |
| 21058 | 199 |
constants, datatypes and classes, resulting in the following |
200 |
code: |
|
201 |
||
202 |
\lstsml{Thy/examples/fac.ML}
|
|
203 |
||
204 |
The code generator will complain when a required |
|
205 |
ingredient does not provide a executable counterpart. |
|
206 |
This is the case if an involved type is not a datatype: |
|
207 |
*} |
|
208 |
||
209 |
(*<*) |
|
210 |
setup {* Sign.add_path "foo" *}
|
|
211 |
(*>*) |
|
212 |
||
213 |
typedecl 'a foo |
|
214 |
||
215 |
definition |
|
| 21993 | 216 |
bar :: "'a foo \<Rightarrow> 'a \<Rightarrow> 'a" where |
| 21058 | 217 |
"bar x y = y" |
218 |
||
219 |
(*<*) |
|
220 |
hide type foo |
|
221 |
hide const bar |
|
222 |
||
223 |
setup {* Sign.parent_path *}
|
|
224 |
||
225 |
datatype 'a foo = Foo |
|
226 |
||
227 |
definition |
|
| 21993 | 228 |
bar :: "'a foo \<Rightarrow> 'a \<Rightarrow> 'a" where |
| 21058 | 229 |
"bar x y = y" |
230 |
(*>*) |
|
231 |
||
232 |
code_gen bar (SML "examples/fail_type.ML") |
|
233 |
||
234 |
text {*
|
|
235 |
\noindent will result in an error. Likewise, generating code |
|
236 |
for constants not yielding |
|
| 22060 | 237 |
a defining equation will fail, e.g.~the Hilbert choice |
| 21058 | 238 |
operation @{text "SOME"}:
|
239 |
*} |
|
240 |
||
241 |
(*<*) |
|
242 |
setup {* Sign.add_path "foo" *}
|
|
243 |
(*>*) |
|
244 |
||
245 |
definition |
|
| 21993 | 246 |
pick_some :: "'a list \<Rightarrow> 'a" where |
| 21058 | 247 |
"pick_some xs = (SOME x. x \<in> set xs)" |
248 |
||
249 |
(*<*) |
|
250 |
hide const pick_some |
|
251 |
||
252 |
setup {* Sign.parent_path *}
|
|
253 |
||
254 |
definition |
|
| 21993 | 255 |
pick_some :: "'a list \<Rightarrow> 'a" where |
| 21058 | 256 |
"pick_some = hd" |
257 |
(*>*) |
|
258 |
||
259 |
code_gen pick_some (SML "examples/fail_const.ML") |
|
260 |
||
| 20948 | 261 |
subsection {* Theorem selection *}
|
262 |
||
| 21058 | 263 |
text {*
|
| 22060 | 264 |
The list of all defining equations in a theory may be inspected |
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
265 |
using the @{text "\<PRINTCODESETUP>"} command:
|
| 21058 | 266 |
*} |
267 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
268 |
print_codesetup |
| 21058 | 269 |
|
270 |
text {*
|
|
271 |
\noindent which displays a table of constant with corresponding |
|
| 22060 | 272 |
defining equations (the additional stuff displayed |
| 21058 | 273 |
shall not bother us for the moment). If this table does |
274 |
not provide at least one function |
|
| 21178 | 275 |
equation, the table of primitive definitions is searched |
| 21058 | 276 |
whether it provides one. |
277 |
||
278 |
The typical HOL tools are already set up in a way that |
|
| 21323 | 279 |
function definitions introduced by @{text "\<FUN>"},
|
280 |
@{text "\<FUNCTION>"}, @{text "\<PRIMREC>"}
|
|
281 |
@{text "\<RECDEF>"} are implicitly propagated
|
|
| 22060 | 282 |
to this defining equation table. Specific theorems may be |
| 21058 | 283 |
selected using an attribute: \emph{code func}. As example,
|
284 |
a weight selector function: |
|
285 |
*} |
|
286 |
||
287 |
consts |
|
288 |
pick :: "(nat \<times> 'a) list \<Rightarrow> nat \<Rightarrow> 'a" |
|
289 |
||
290 |
primrec |
|
291 |
"pick (x#xs) n = (let (k, v) = x in |
|
292 |
if n < k then v else pick xs (n - k))" |
|
293 |
||
294 |
text {*
|
|
295 |
We want to eliminate the explicit destruction |
|
296 |
of @{term x} to @{term "(k, v)"}:
|
|
297 |
*} |
|
298 |
||
299 |
lemma [code func]: |
|
300 |
"pick ((k, v)#xs) n = (if n < k then v else pick xs (n - k))" |
|
301 |
by simp |
|
302 |
||
| 21545 | 303 |
code_gen pick (*<*)(SML #)(*>*)(SML "examples/pick1.ML") |
| 21058 | 304 |
|
305 |
text {*
|
|
| 22060 | 306 |
This theorem is now added to the defining equation table: |
| 21058 | 307 |
|
308 |
\lstsml{Thy/examples/pick1.ML}
|
|
309 |
||
310 |
It might be convenient to remove the pointless original |
|
311 |
equation, using the \emph{nofunc} attribute:
|
|
312 |
*} |
|
313 |
||
314 |
lemmas [code nofunc] = pick.simps |
|
315 |
||
| 21545 | 316 |
code_gen pick (*<*)(SML #)(*>*)(SML "examples/pick2.ML") |
| 21058 | 317 |
|
318 |
text {*
|
|
319 |
\lstsml{Thy/examples/pick2.ML}
|
|
320 |
||
321 |
Syntactic redundancies are implicitly dropped. For example, |
|
322 |
using a modified version of the @{const fac} function
|
|
| 22060 | 323 |
as defining equation, the then redundant (since |
324 |
syntactically subsumed) original defining equations |
|
| 21058 | 325 |
are dropped, resulting in a warning: |
326 |
*} |
|
327 |
||
328 |
lemma [code func]: |
|
329 |
"fac n = (case n of 0 \<Rightarrow> 1 | Suc m \<Rightarrow> n * fac m)" |
|
330 |
by (cases n) simp_all |
|
331 |
||
| 21545 | 332 |
code_gen fac (*<*)(SML #)(*>*)(SML "examples/fac_case.ML") |
| 21058 | 333 |
|
334 |
text {*
|
|
335 |
\lstsml{Thy/examples/fac_case.ML}
|
|
336 |
||
337 |
\begin{warn}
|
|
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
338 |
The attributes \emph{code} and \emph{code del}
|
| 21058 | 339 |
associated with the existing code generator also apply to |
340 |
the new one: \emph{code} implies \emph{code func},
|
|
341 |
and \emph{code del} implies \emph{code nofunc}.
|
|
342 |
\end{warn}
|
|
343 |
*} |
|
| 20948 | 344 |
|
345 |
subsection {* Type classes *}
|
|
346 |
||
| 21058 | 347 |
text {*
|
348 |
Type classes enter the game via the Isar class package. |
|
| 21075 | 349 |
For a short introduction how to use it, see \cite{isabelle-classes};
|
| 21147 | 350 |
here we just illustrate its impact on code generation. |
| 21058 | 351 |
|
352 |
In a target language, type classes may be represented |
|
| 21178 | 353 |
natively (as in the case of Haskell). For languages |
| 21058 | 354 |
like SML, they are implemented using \emph{dictionaries}.
|
| 21178 | 355 |
Our following example specifies a class \qt{null},
|
| 21058 | 356 |
assigning to each of its inhabitants a \qt{null} value:
|
357 |
*} |
|
358 |
||
| 22473 | 359 |
class null = type + |
| 21058 | 360 |
fixes null :: 'a |
361 |
||
362 |
consts |
|
363 |
head :: "'a\<Colon>null list \<Rightarrow> 'a" |
|
364 |
||
365 |
primrec |
|
366 |
"head [] = null" |
|
367 |
"head (x#xs) = x" |
|
368 |
||
369 |
text {*
|
|
370 |
We provide some instances for our @{text null}:
|
|
371 |
*} |
|
372 |
||
373 |
instance option :: (type) null |
|
374 |
"null \<equiv> None" .. |
|
375 |
||
376 |
instance list :: (type) null |
|
377 |
"null \<equiv> []" .. |
|
378 |
||
379 |
text {*
|
|
380 |
Constructing a dummy example: |
|
381 |
*} |
|
382 |
||
383 |
definition |
|
384 |
"dummy = head [Some (Suc 0), None]" |
|
385 |
||
386 |
text {*
|
|
| 21178 | 387 |
Type classes offer a suitable occasion to introduce |
| 21058 | 388 |
the Haskell serializer. Its usage is almost the same |
| 21075 | 389 |
as SML, but, in accordance with conventions |
390 |
some Haskell systems enforce, each module ends |
|
391 |
up in a single file. The module hierarchy is reflected in |
|
392 |
the file system, with root given by the user. |
|
| 21058 | 393 |
*} |
394 |
||
| 21075 | 395 |
code_gen dummy (Haskell "examples/") |
| 21147 | 396 |
(* NOTE: you may use Haskell only once in this document, otherwise |
397 |
you have to work in distinct subdirectories *) |
|
| 21058 | 398 |
|
399 |
text {*
|
|
400 |
\lsthaskell{Thy/examples/Codegen.hs}
|
|
401 |
||
402 |
(we have left out all other modules). |
|
403 |
||
404 |
The whole code in SML with explicit dictionary passing: |
|
405 |
*} |
|
406 |
||
| 21545 | 407 |
code_gen dummy (*<*)(SML #)(*>*)(SML "examples/class.ML") |
| 21058 | 408 |
|
409 |
text {*
|
|
410 |
\lstsml{Thy/examples/class.ML}
|
|
411 |
*} |
|
412 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
413 |
text {*
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
414 |
or in OCaml: |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
415 |
*} |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
416 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
417 |
code_gen dummy (OCaml "examples/class.ocaml") |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
418 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
419 |
text {*
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
420 |
\lstsml{Thy/examples/class.ocaml}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
421 |
*} |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
422 |
|
| 21058 | 423 |
subsection {* Incremental code generation *}
|
424 |
||
| 21075 | 425 |
text {*
|
426 |
Code generation is \emph{incremental}: theorems
|
|
427 |
and abstract intermediate code are cached and extended on demand. |
|
428 |
The cache may be partially or fully dropped if the underlying |
|
429 |
executable content of the theory changes. |
|
430 |
Implementation of caching is supposed to transparently |
|
431 |
hid away the details from the user. Anyway, caching |
|
432 |
reaches the surface by using a slightly more general form |
|
| 21323 | 433 |
of the @{text "\<CODEGEN>"}: either the list of constants or the
|
| 21075 | 434 |
list of serialization expressions may be dropped. If no |
435 |
serialization expressions are given, only abstract code |
|
436 |
is generated and cached; if no constants are given, the |
|
437 |
current cache is serialized. |
|
438 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
439 |
For explorative purpose, the |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
440 |
@{text "\<CODETHMS>"} command may prove useful:
|
| 21075 | 441 |
*} |
442 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
443 |
code_thms |
| 21075 | 444 |
|
445 |
text {*
|
|
| 22060 | 446 |
\noindent print all cached defining equations (i.e.~\emph{after}
|
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
447 |
any applied transformation). A |
| 21075 | 448 |
list of constants may be given; their function |
| 21178 | 449 |
equations are added to the cache if not already present. |
| 21075 | 450 |
*} |
451 |
||
| 21058 | 452 |
|
453 |
section {* Recipes and advanced topics \label{sec:advanced} *}
|
|
454 |
||
| 21089 | 455 |
text {*
|
456 |
In this tutorial, we do not attempt to give an exhaustive |
|
457 |
description of the code generator framework; instead, |
|
458 |
we cast a light on advanced topics by introducing |
|
459 |
them together with practically motivated examples. Concerning |
|
460 |
further reading, see |
|
461 |
||
462 |
\begin{itemize}
|
|
463 |
||
464 |
\item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}
|
|
465 |
for exhaustive syntax diagrams. |
|
| 21222 | 466 |
\item or \fixme[ref] which deals with foundational issues |
| 21089 | 467 |
of the code generator framework. |
468 |
||
469 |
\end{itemize}
|
|
470 |
*} |
|
| 21058 | 471 |
|
472 |
subsection {* Library theories *}
|
|
473 |
||
| 21089 | 474 |
text {*
|
475 |
The HOL \emph{Main} theory already provides a code generator setup
|
|
476 |
which should be suitable for most applications. Common extensions |
|
477 |
and modifications are available by certain theories of the HOL |
|
478 |
library; beside being useful in applications, they may serve |
|
| 21178 | 479 |
as a tutorial for customizing the code generator setup. |
| 21089 | 480 |
|
481 |
\begin{description}
|
|
482 |
||
| 21452 | 483 |
\item[@{text "ExecutableSet"}] allows to generate code
|
| 21089 | 484 |
for finite sets using lists. |
| 21452 | 485 |
\item[@{text "ExecutableRat"}] \label{exec_rat} implements rational
|
| 21089 | 486 |
numbers as triples @{text "(sign, enumerator, denominator)"}.
|
| 21452 | 487 |
\item[@{text "EfficientNat"}] \label{eff_nat} implements natural numbers by integers,
|
| 21178 | 488 |
which in general will result in higher efficency; pattern |
| 21089 | 489 |
matching with @{const "0\<Colon>nat"} / @{const "Suc"}
|
| 21189 | 490 |
is eliminated. |
| 21452 | 491 |
\item[@{text "MLString"}] provides an additional datatype @{text "mlstring"};
|
| 21089 | 492 |
in the HOL default setup, strings in HOL are mapped to list |
493 |
of chars in SML; values of type @{text "mlstring"} are
|
|
494 |
mapped to strings in SML. |
|
495 |
||
496 |
\end{description}
|
|
497 |
*} |
|
498 |
||
| 20948 | 499 |
subsection {* Preprocessing *}
|
500 |
||
| 21089 | 501 |
text {*
|
| 21147 | 502 |
Before selected function theorems are turned into abstract |
503 |
code, a chain of definitional transformation steps is carried |
|
| 21178 | 504 |
out: \emph{preprocessing}. There are three possibilities
|
| 21147 | 505 |
to customize preprocessing: \emph{inline theorems},
|
506 |
\emph{inline procedures} and \emph{generic preprocessors}.
|
|
507 |
||
508 |
\emph{Inline theorems} are rewriting rules applied to each
|
|
| 22060 | 509 |
defining equation. Due to the interpretation of theorems |
510 |
of defining equations, rewrites are applied to the right |
|
| 21147 | 511 |
hand side and the arguments of the left hand side of an |
512 |
equation, but never to the constant heading the left hand side. |
|
513 |
Inline theorems may be declared an undeclared using the |
|
| 21178 | 514 |
\emph{code inline} or \emph{code noinline} attribute respectively.
|
| 21147 | 515 |
|
516 |
Some common applications: |
|
517 |
*} |
|
518 |
||
519 |
text_raw {*
|
|
520 |
\begin{itemize}
|
|
521 |
\item replacing non-executable constructs by executable ones: \\ |
|
522 |
*} |
|
523 |
||
524 |
lemma [code inline]: |
|
525 |
"x \<in> set xs \<longleftrightarrow> x mem xs" by (induct xs) simp_all |
|
526 |
||
527 |
text_raw {*
|
|
528 |
\item eliminating superfluous constants: \\ |
|
529 |
*} |
|
530 |
||
531 |
lemma [code inline]: |
|
532 |
"1 = Suc 0" by simp |
|
533 |
||
534 |
text_raw {*
|
|
535 |
\item replacing executable but inconvenient constructs: \\ |
|
| 21089 | 536 |
*} |
537 |
||
| 21147 | 538 |
lemma [code inline]: |
539 |
"xs = [] \<longleftrightarrow> List.null xs" by (induct xs) simp_all |
|
540 |
||
541 |
text_raw {*
|
|
542 |
\end{itemize}
|
|
543 |
*} |
|
544 |
||
545 |
text {*
|
|
546 |
The current set of inline theorems may be inspected using |
|
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
547 |
the @{text "\<PRINTCODESETUP>"} command.
|
| 21147 | 548 |
|
549 |
\emph{Inline procedures} are a generalized version of inline
|
|
550 |
theorems written in ML -- rewrite rules are generated dependent |
|
551 |
on the function theorems for a certain function. One |
|
552 |
application is the implicit expanding of @{typ nat} numerals
|
|
553 |
to @{const "0\<Colon>nat"} / @{const Suc} representation. See further
|
|
554 |
\secref{sec:ml}
|
|
555 |
||
556 |
\emph{Generic preprocessors} provide a most general interface,
|
|
557 |
transforming a list of function theorems to another |
|
558 |
list of function theorems, provided that neither the heading |
|
559 |
constant nor its type change. The @{const "0\<Colon>nat"} / @{const Suc}
|
|
| 21323 | 560 |
pattern elimination implemented in |
| 21452 | 561 |
theory @{text "EfficientNat"} (\secref{eff_nat}) uses this
|
| 21147 | 562 |
interface. |
563 |
||
564 |
\begin{warn}
|
|
565 |
The order in which single preprocessing steps are carried |
|
566 |
out currently is not specified; in particular, preprocessing |
|
| 21178 | 567 |
is \emph{no} fix point process. Keep this in mind when
|
| 21147 | 568 |
setting up the preprocessor. |
569 |
||
570 |
Further, the attribute \emph{code unfold}
|
|
571 |
associated with the existing code generator also applies to |
|
572 |
the new one: \emph{code unfold} implies \emph{code inline}.
|
|
573 |
\end{warn}
|
|
574 |
*} |
|
| 20948 | 575 |
|
| 21058 | 576 |
subsection {* Customizing serialization *}
|
| 20948 | 577 |
|
| 21147 | 578 |
text {*
|
579 |
Consider the following function and its corresponding |
|
580 |
SML code: |
|
581 |
*} |
|
582 |
||
583 |
fun |
|
584 |
in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where |
|
585 |
"in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l" |
|
586 |
||
587 |
(*<*) |
|
| 21323 | 588 |
code_type %tt bool |
| 21147 | 589 |
(SML) |
| 21323 | 590 |
code_const %tt True and False and "op \<and>" and Not |
| 21147 | 591 |
(SML and and and) |
592 |
(*>*) |
|
593 |
||
| 21545 | 594 |
code_gen in_interval (*<*)(SML #)(*>*)(SML "examples/bool_literal.ML") |
| 21147 | 595 |
|
596 |
text {*
|
|
| 21323 | 597 |
\lstsml{Thy/examples/bool_literal.ML}
|
| 21147 | 598 |
|
599 |
Though this is correct code, it is a little bit unsatisfactory: |
|
600 |
boolean values and operators are materialized as distinguished |
|
601 |
entities with have nothing to do with the SML-builtin notion |
|
602 |
of \qt{bool}. This results in less readable code;
|
|
603 |
additionally, eager evaluation may cause programs to |
|
604 |
loop or break which would perfectly terminate when |
|
605 |
the existing SML \qt{bool} would be used. To map
|
|
606 |
the HOL \qt{bool} on SML \qt{bool}, we may use
|
|
607 |
\qn{custom serializations}:
|
|
608 |
*} |
|
609 |
||
| 21323 | 610 |
code_type %tt bool |
| 21147 | 611 |
(SML "bool") |
| 21323 | 612 |
code_const %tt True and False and "op \<and>" |
| 21147 | 613 |
(SML "true" and "false" and "_ andalso _") |
614 |
||
615 |
text {*
|
|
| 21323 | 616 |
The @{text "\<CODETYPE>"} commad takes a type constructor
|
| 21147 | 617 |
as arguments together with a list of custom serializations. |
618 |
Each custom serialization starts with a target language |
|
619 |
identifier followed by an expression, which during |
|
620 |
code serialization is inserted whenever the type constructor |
|
| 21323 | 621 |
would occur. For constants, @{text "\<CODECONST>"} implements
|
622 |
the corresponding mechanism. Each ``@{verbatim "_"}'' in
|
|
| 21147 | 623 |
a serialization expression is treated as a placeholder |
624 |
for the type constructor's (the constant's) arguments. |
|
625 |
*} |
|
626 |
||
627 |
code_reserved SML |
|
628 |
bool true false |
|
629 |
||
630 |
text {*
|
|
631 |
To assert that the existing \qt{bool}, \qt{true} and \qt{false}
|
|
| 21323 | 632 |
is not used for generated code, we use @{text "\<CODERESERVED>"}.
|
| 21147 | 633 |
|
634 |
After this setup, code looks quite more readable: |
|
635 |
*} |
|
636 |
||
| 21545 | 637 |
code_gen in_interval (*<*)(SML #)(*>*)(SML "examples/bool_mlbool.ML") |
| 21147 | 638 |
|
639 |
text {*
|
|
| 21323 | 640 |
\lstsml{Thy/examples/bool_mlbool.ML}
|
| 21147 | 641 |
|
642 |
This still is not perfect: the parentheses |
|
| 21323 | 643 |
around the \qt{andalso} expression are superfluous.
|
644 |
Though the serializer |
|
| 21147 | 645 |
by no means attempts to imitate the rich Isabelle syntax |
646 |
framework, it provides some common idioms, notably |
|
647 |
associative infixes with precedences which may be used here: |
|
648 |
*} |
|
649 |
||
| 21323 | 650 |
code_const %tt "op \<and>" |
| 21147 | 651 |
(SML infixl 1 "andalso") |
652 |
||
| 21545 | 653 |
code_gen in_interval (*<*)(SML #)(*>*)(SML "examples/bool_infix.ML") |
| 21147 | 654 |
|
655 |
text {*
|
|
| 21323 | 656 |
\lstsml{Thy/examples/bool_infix.ML}
|
| 21147 | 657 |
|
658 |
Next, we try to map HOL pairs to SML pairs, using the |
|
| 21323 | 659 |
infix ``@{verbatim "*"}'' type constructor and parentheses:
|
| 21147 | 660 |
*} |
661 |
||
662 |
(*<*) |
|
663 |
code_type * |
|
664 |
(SML) |
|
665 |
code_const Pair |
|
666 |
(SML) |
|
667 |
(*>*) |
|
668 |
||
| 21323 | 669 |
code_type %tt * |
| 21147 | 670 |
(SML infix 2 "*") |
671 |
||
| 21323 | 672 |
code_const %tt Pair |
| 21147 | 673 |
(SML "!((_),/ (_))") |
674 |
||
675 |
text {*
|
|
| 21323 | 676 |
The initial bang ``@{verbatim "!"}'' tells the serializer to never put
|
| 21147 | 677 |
parentheses around the whole expression (they are already present), |
678 |
while the parentheses around argument place holders |
|
679 |
tell not to put parentheses around the arguments. |
|
| 21323 | 680 |
The slash ``@{verbatim "/"}'' (followed by arbitrary white space)
|
| 21147 | 681 |
inserts a space which may be used as a break if necessary |
682 |
during pretty printing. |
|
683 |
||
| 21178 | 684 |
So far, we did only provide more idiomatic serializations for |
685 |
constructs which would be executable on their own. Target-specific |
|
686 |
serializations may also be used to \emph{implement} constructs
|
|
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
687 |
which have no explicit notion of executability. For example, |
| 21178 | 688 |
take the HOL integers: |
689 |
*} |
|
690 |
||
691 |
definition |
|
| 21993 | 692 |
double_inc :: "int \<Rightarrow> int" where |
| 21178 | 693 |
"double_inc k = 2 * k + 1" |
694 |
||
695 |
code_gen double_inc (SML "examples/integers.ML") |
|
696 |
||
697 |
text {*
|
|
698 |
will fail: @{typ int} in HOL is implemented using a quotient
|
|
699 |
type, which does not provide any notion of executability. |
|
700 |
\footnote{Eventually, we also want to provide executability
|
|
701 |
for quotients.}. However, we could use the SML builtin |
|
702 |
integers: |
|
| 21147 | 703 |
*} |
704 |
||
| 21323 | 705 |
code_type %tt int |
| 21147 | 706 |
(SML "IntInf.int") |
707 |
||
| 21323 | 708 |
code_const %tt "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int" |
| 21147 | 709 |
and "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int" |
| 21178 | 710 |
(SML "IntInf.+ (_, _)" and "IntInf.* (_, _)") |
711 |
||
| 21545 | 712 |
code_gen double_inc (*<*)(SML #)(*>*)(SML "examples/integers.ML") |
| 21147 | 713 |
|
| 21178 | 714 |
text {*
|
715 |
resulting in: |
|
| 21147 | 716 |
|
| 21178 | 717 |
\lstsml{Thy/examples/integers.ML}
|
718 |
*} |
|
| 21147 | 719 |
|
| 21178 | 720 |
text {*
|
721 |
These examples give a glimpse what powerful mechanisms |
|
722 |
custom serializations provide; however their usage |
|
723 |
requires careful thinking in order not to introduce |
|
724 |
inconsistencies -- or, in other words: |
|
725 |
custom serializations are completely axiomatic. |
|
| 21147 | 726 |
|
| 21178 | 727 |
A further noteworthy details is that any special |
728 |
character in a custom serialization may be quoted |
|
| 21323 | 729 |
using ``@{verbatim "'"}''; thus, in
|
730 |
``@{verbatim "fn '_ => _"}'' the first
|
|
731 |
``@{verbatim "_"}'' is a proper underscore while the
|
|
732 |
second ``@{verbatim "_"}'' is a placeholder.
|
|
| 21147 | 733 |
|
| 21178 | 734 |
The HOL theories provide further |
735 |
examples for custom serializations and form |
|
736 |
a recommended tutorial on how to use them properly. |
|
737 |
*} |
|
| 21147 | 738 |
|
739 |
subsection {* Concerning operational equality *}
|
|
740 |
||
741 |
text {*
|
|
742 |
Surely you have already noticed how equality is treated |
|
743 |
by the code generator: |
|
744 |
*} |
|
745 |
||
746 |
fun |
|
747 |
collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
| 22473 | 748 |
"collect_duplicates xs ys [] = xs" |
749 |
| "collect_duplicates xs ys (z#zs) = (if z \<in> set xs |
|
750 |
then if z \<in> set ys |
|
751 |
then collect_duplicates xs ys zs |
|
752 |
else collect_duplicates xs (z#ys) zs |
|
753 |
else collect_duplicates (z#xs) (z#ys) zs)" |
|
| 21147 | 754 |
|
755 |
text {*
|
|
| 21217 | 756 |
The membership test during preprocessing is rewritten, |
| 21147 | 757 |
resulting in @{const List.memberl}, which itself
|
758 |
performs an explicit equality check. |
|
759 |
*} |
|
760 |
||
| 21545 | 761 |
code_gen collect_duplicates (*<*)(SML #)(*>*)(SML "examples/collect_duplicates.ML") |
| 21147 | 762 |
|
763 |
text {*
|
|
764 |
\lstsml{Thy/examples/collect_duplicates.ML}
|
|
765 |
*} |
|
766 |
||
767 |
text {*
|
|
768 |
Obviously, polymorphic equality is implemented the Haskell |
|
769 |
way using a type class. How is this achieved? By an |
|
770 |
almost trivial definition in the HOL setup: |
|
771 |
*} |
|
772 |
||
773 |
(*<*) |
|
774 |
setup {* Sign.add_path "foo" *}
|
|
| 21452 | 775 |
consts "op =" :: "'a" |
| 21147 | 776 |
(*>*) |
777 |
||
| 22473 | 778 |
class eq (attach "op =") = type |
| 21147 | 779 |
|
780 |
text {*
|
|
781 |
This merely introduces a class @{text eq} with corresponding
|
|
| 21993 | 782 |
operation @{text "op ="};
|
| 21452 | 783 |
the preprocessing framework does the rest. |
| 21147 | 784 |
*} |
785 |
||
786 |
(*<*) |
|
787 |
hide (open) "class" eq |
|
| 21452 | 788 |
hide (open) const "op =" |
| 21147 | 789 |
setup {* Sign.parent_path *}
|
790 |
(*>*) |
|
791 |
||
792 |
text {*
|
|
793 |
For datatypes, instances of @{text eq} are implicitly derived
|
|
794 |
when possible. |
|
795 |
||
796 |
Though this class is designed to get rarely in the way, there |
|
797 |
are some cases when it suddenly comes to surface: |
|
798 |
*} |
|
799 |
||
| 21223 | 800 |
subsubsection {* typedecls interpreted by customary serializations *}
|
| 21178 | 801 |
|
802 |
text {*
|
|
803 |
A common idiom is to use unspecified types for formalizations |
|
804 |
and interpret them for a specific target language: |
|
| 21147 | 805 |
*} |
806 |
||
807 |
typedecl key |
|
808 |
||
809 |
fun |
|
810 |
lookup :: "(key \<times> 'a) list \<Rightarrow> key \<Rightarrow> 'a option" where |
|
| 22473 | 811 |
"lookup [] l = None" |
812 |
| "lookup ((k, v) # xs) l = (if k = l then Some v else lookup xs l)" |
|
| 21147 | 813 |
|
| 21323 | 814 |
code_type %tt key |
| 21147 | 815 |
(SML "string") |
816 |
||
| 21178 | 817 |
text {*
|
818 |
This, though, is not sufficient: @{typ key} is no instance
|
|
| 21147 | 819 |
of @{text eq} since @{typ key} is no datatype; the instance
|
820 |
has to be declared manually, including a serialization |
|
| 21452 | 821 |
for the particular instance of @{const "op ="}:
|
| 21147 | 822 |
*} |
823 |
||
824 |
instance key :: eq .. |
|
825 |
||
| 21452 | 826 |
code_const %tt "op = \<Colon> key \<Rightarrow> key \<Rightarrow> bool" |
827 |
(SML "!((_ : string) = _)") |
|
| 21147 | 828 |
|
| 21178 | 829 |
text {*
|
830 |
Then everything goes fine: |
|
| 21147 | 831 |
*} |
832 |
||
| 21545 | 833 |
code_gen lookup (*<*)(SML #)(*>*)(SML "examples/lookup.ML") |
| 21147 | 834 |
|
835 |
text {*
|
|
836 |
\lstsml{Thy/examples/lookup.ML}
|
|
837 |
*} |
|
838 |
||
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
839 |
subsubsection {* lexicographic orderings *}
|
| 21178 | 840 |
|
841 |
text {*
|
|
842 |
Another subtlety |
|
| 21147 | 843 |
enters the stage when definitions of overloaded constants |
844 |
are dependent on operational equality. For example, let |
|
| 21189 | 845 |
us define a lexicographic ordering on tuples: |
| 21147 | 846 |
*} |
847 |
||
| 21178 | 848 |
instance * :: (ord, ord) ord |
849 |
"p1 < p2 \<equiv> let (x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) = p1; (x2, y2) = p2 in |
|
| 21147 | 850 |
x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)" |
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
851 |
"p1 \<le> p2 \<equiv> let (x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) = p1; (x2, y2) = p2 in |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
852 |
x1 < x2 \<or> (x1 = x2 \<and> y1 \<le> y2)" .. |
| 21147 | 853 |
|
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
854 |
lemma ord_prod [code func]: |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
855 |
"(x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)" |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
856 |
"(x1 \<Colon> 'a\<Colon>ord, y1 \<Colon> 'b\<Colon>ord) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> (x1 = x2 \<and> y1 \<le> y2)" |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
857 |
unfolding "less_eq_*_def" "less_*_def" by simp_all |
| 21147 | 858 |
|
| 21178 | 859 |
text {*
|
860 |
Then code generation will fail. Why? The definition |
|
| 21147 | 861 |
of @{const "op \<le>"} depends on equality on both arguments,
|
| 21178 | 862 |
which are polymorphic and impose an additional @{text eq}
|
| 21147 | 863 |
class constraint, thus violating the type discipline |
864 |
for class operations. |
|
865 |
||
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
866 |
The solution is to add @{text eq} explicitly to the first sort arguments in the
|
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
867 |
code theorems: |
| 21147 | 868 |
*} |
869 |
||
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
870 |
(*<*) |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
871 |
declare ord_prod [code del] |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
872 |
(*>*) |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
873 |
|
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
874 |
lemma ord_prod_code [code func]: |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
875 |
"(x1 \<Colon> 'a\<Colon>{ord, eq}, y1 \<Colon> 'b\<Colon>ord) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> (x1 = x2 \<and> y1 < y2)"
|
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
876 |
"(x1 \<Colon> 'a\<Colon>{ord, eq}, y1 \<Colon> 'b\<Colon>ord) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> (x1 = x2 \<and> y1 \<le> y2)"
|
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
877 |
unfolding ord_prod by rule+ |
| 21147 | 878 |
|
| 21178 | 879 |
text {*
|
880 |
Then code generation succeeds: |
|
| 21147 | 881 |
*} |
882 |
||
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
883 |
code_gen "op \<le> \<Colon> 'a\<Colon>{eq, ord} \<times> 'b\<Colon>ord \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
|
| 21545 | 884 |
(*<*)(SML #)(*>*)(SML "examples/lexicographic.ML") |
| 21147 | 885 |
|
886 |
text {*
|
|
887 |
\lstsml{Thy/examples/lexicographic.ML}
|
|
888 |
*} |
|
889 |
||
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
890 |
text {*
|
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
891 |
In general, code theorems for overloaded constants may have more |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
892 |
restrictive sort constraints than the underlying instance relation |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
893 |
between class and type constructor as long as the whole system of |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
894 |
constraints is coregular; code theorems violating coregularity |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
895 |
are rejected immediately. |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
896 |
*} |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
897 |
|
| 21178 | 898 |
subsubsection {* Haskell serialization *}
|
899 |
||
900 |
text {*
|
|
901 |
For convenience, the default |
|
902 |
HOL setup for Haskell maps the @{text eq} class to
|
|
903 |
its counterpart in Haskell, giving custom serializations |
|
| 21323 | 904 |
for the class (@{text "\<CODECLASS>"}) and its operation:
|
| 21178 | 905 |
*} |
906 |
||
907 |
(*<*) |
|
908 |
setup {* Sign.add_path "bar" *}
|
|
| 22473 | 909 |
class eq = type + fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
| 21178 | 910 |
(*>*) |
911 |
||
| 21323 | 912 |
code_class %tt eq |
| 21178 | 913 |
(Haskell "Eq" where eq \<equiv> "(==)") |
914 |
||
| 21323 | 915 |
code_const %tt eq |
| 21178 | 916 |
(Haskell infixl 4 "==") |
917 |
||
918 |
(*<*) |
|
919 |
hide "class" eq |
|
920 |
hide const eq |
|
921 |
setup {* Sign.parent_path *}
|
|
922 |
(*>*) |
|
923 |
||
924 |
text {*
|
|
925 |
A problem now occurs whenever a type which |
|
926 |
is an instance of @{text eq} in HOL is mapped
|
|
927 |
on a Haskell-builtin type which is also an instance |
|
928 |
of Haskell @{text Eq}:
|
|
| 21147 | 929 |
*} |
930 |
||
| 21178 | 931 |
typedecl bar |
932 |
||
933 |
instance bar :: eq .. |
|
934 |
||
| 21323 | 935 |
code_type %tt bar |
| 21178 | 936 |
(Haskell "Integer") |
937 |
||
938 |
text {*
|
|
|
22188
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
939 |
The code generator would produce |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
940 |
an additional instance, which of course is rejected. |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
941 |
To suppress this additional instance, use |
|
a63889770d57
adjusted manual to improved treatment of overloaded constants
haftmann
parents:
22175
diff
changeset
|
942 |
@{text "\<CODEINSTANCE>"}:
|
| 21147 | 943 |
*} |
944 |
||
| 21323 | 945 |
code_instance %tt bar :: eq |
| 21178 | 946 |
(Haskell -) |
947 |
||
948 |
subsection {* Types matter *}
|
|
949 |
||
950 |
text {*
|
|
951 |
Imagine the following quick-and-dirty setup for implementing |
|
| 21189 | 952 |
some kind of sets as lists in SML: |
| 21178 | 953 |
*} |
954 |
||
| 21323 | 955 |
code_type %tt set |
| 21178 | 956 |
(SML "_ list") |
957 |
||
| 21323 | 958 |
code_const %tt "{}" and insert
|
| 21178 | 959 |
(SML "![]" and infixl 7 "::") |
960 |
||
961 |
definition |
|
| 21993 | 962 |
dummy_set :: "(nat \<Rightarrow> nat) set" where |
| 21189 | 963 |
"dummy_set = {Suc}"
|
964 |
||
965 |
text {*
|
|
966 |
Then code generation for @{const dummy_set} will fail.
|
|
| 22060 | 967 |
Why? A glimpse at the defining equations will offer: |
| 21189 | 968 |
*} |
969 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
970 |
code_thms insert |
| 21189 | 971 |
|
972 |
text {*
|
|
| 22060 | 973 |
This reveals the defining equation @{thm insert_def}
|
| 21223 | 974 |
for @{const insert}, which is operationally meaningless
|
| 21189 | 975 |
but forces an equality constraint on the set members |
| 21223 | 976 |
(which is not satisfiable if the set members are functions). |
| 21189 | 977 |
Even when using set of natural numbers (which are an instance |
978 |
of \emph{eq}), we run into a problem:
|
|
979 |
*} |
|
980 |
||
981 |
definition |
|
| 21993 | 982 |
foobar_set :: "nat set" where |
| 21189 | 983 |
"foobar_set = {0, 1, 2}"
|
984 |
||
985 |
text {*
|
|
986 |
In this case the serializer would complain that @{const insert}
|
|
987 |
expects dictionaries (namely an \emph{eq} dictionary) but
|
|
988 |
has also been given a customary serialization. |
|
989 |
||
990 |
The solution to this dilemma: |
|
991 |
*} |
|
992 |
||
993 |
lemma [code func]: |
|
994 |
"insert = insert" .. |
|
995 |
||
| 21545 | 996 |
code_gen dummy_set foobar_set (*<*)(SML #)(*>*)(SML "examples/dirty_set.ML") |
| 21189 | 997 |
|
998 |
text {*
|
|
999 |
\lstsml{Thy/examples/dirty_set.ML}
|
|
| 21178 | 1000 |
|
| 22060 | 1001 |
Reflexive defining equations by convention are dropped. |
| 21189 | 1002 |
But their presence prevents primitive definitions to be |
| 22060 | 1003 |
used as defining equations: |
| 21189 | 1004 |
*} |
1005 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1006 |
code_thms insert |
| 21189 | 1007 |
|
1008 |
text {*
|
|
| 22060 | 1009 |
will show \emph{no} defining equations for insert.
|
| 21178 | 1010 |
|
| 21189 | 1011 |
Note that the sort constraints of reflexive equations |
1012 |
are considered; so |
|
1013 |
*} |
|
1014 |
||
1015 |
lemma [code func]: |
|
1016 |
"(insert \<Colon> 'a\<Colon>eq \<Rightarrow> 'a set \<Rightarrow> 'a set) = insert" .. |
|
1017 |
||
1018 |
text {*
|
|
1019 |
would mean nothing else than to introduce the evil |
|
1020 |
sort constraint by hand. |
|
1021 |
*} |
|
1022 |
||
1023 |
subsection {* Cyclic module dependencies *}
|
|
| 21178 | 1024 |
|
| 21189 | 1025 |
text {*
|
1026 |
Sometimes the awkward situation occurs that dependencies |
|
1027 |
between definitions introduce cyclic dependencies |
|
1028 |
between modules, which in the Haskell world leaves |
|
1029 |
you to the mercy of the Haskell implementation you are using, |
|
1030 |
while for SML code generation is not possible. |
|
| 21178 | 1031 |
|
| 21189 | 1032 |
A solution is to declare module names explicitly. |
1033 |
Let use assume the three cyclically dependent |
|
1034 |
modules are named \emph{A}, \emph{B} and \emph{C}.
|
|
1035 |
Then, by stating |
|
1036 |
*} |
|
1037 |
||
1038 |
code_modulename SML |
|
1039 |
A ABC |
|
1040 |
B ABC |
|
1041 |
C ABC |
|
1042 |
||
1043 |
text {*
|
|
1044 |
we explicitly map all those modules on \emph{ABC},
|
|
1045 |
resulting in an ad-hoc merge of this three modules |
|
1046 |
at serialization time. |
|
1047 |
*} |
|
| 21147 | 1048 |
|
1049 |
subsection {* Axiomatic extensions *}
|
|
1050 |
||
1051 |
text {*
|
|
1052 |
\begin{warn}
|
|
1053 |
The extensions introduced in this section, though working |
|
| 21189 | 1054 |
in practice, are not the cream of the crop, as you |
1055 |
will notice during reading. They will |
|
| 21147 | 1056 |
eventually be replaced by more mature approaches. |
1057 |
\end{warn}
|
|
| 21189 | 1058 |
|
1059 |
Sometimes equalities are taken for granted which are |
|
1060 |
not derivable inside the HOL logic but are silently assumed |
|
1061 |
to hold for executable code. For example, we may want |
|
1062 |
to identify the famous HOL constant @{const arbitrary}
|
|
1063 |
of type @{typ "'a option"} with @{const None}.
|
|
1064 |
By brute force: |
|
1065 |
*} |
|
1066 |
||
| 21323 | 1067 |
axiomatization where |
1068 |
"arbitrary = None" |
|
| 21189 | 1069 |
|
1070 |
text {*
|
|
1071 |
However this has to be considered harmful since this axiom, |
|
1072 |
though probably justifiable for generated code, could |
|
1073 |
introduce serious inconsistencies into the logic. |
|
1074 |
||
1075 |
So, there is a distinguished construct for stating axiomatic |
|
1076 |
equalities of constants which apply only for code generation. |
|
1077 |
Before introducing this, here is a convenient place to describe |
|
1078 |
shortly how to deal with some restrictions the type discipline |
|
1079 |
imposes. |
|
1080 |
||
1081 |
By itself, the constant @{const arbitrary} is a non-overloaded
|
|
1082 |
polymorphic constant. So, there is no way to distinguish |
|
1083 |
different versions of @{const arbitrary} for different types
|
|
1084 |
inside the code generator framework. However, inlining |
|
1085 |
theorems together with auxiliary constants provide a solution: |
|
| 21147 | 1086 |
*} |
1087 |
||
| 21189 | 1088 |
definition |
| 21993 | 1089 |
arbitrary_option :: "'a option" where |
| 21189 | 1090 |
[symmetric, code inline]: "arbitrary_option = arbitrary" |
1091 |
||
1092 |
text {*
|
|
1093 |
By that, we replace any @{const arbitrary} with option type
|
|
| 22060 | 1094 |
by @{const arbitrary_option} in defining equations.
|
| 21189 | 1095 |
|
1096 |
For technical reasons, we further have to provide a |
|
1097 |
synonym for @{const None} which in code generator view
|
|
| 22175 | 1098 |
is a function rather than a term constructor: |
| 21189 | 1099 |
*} |
1100 |
||
1101 |
definition |
|
1102 |
"None' = None" |
|
1103 |
||
1104 |
text {*
|
|
| 21323 | 1105 |
Then finally we are enabled to use @{text "\<CODEAXIOMS>"}:
|
| 21189 | 1106 |
*} |
1107 |
||
1108 |
code_axioms |
|
1109 |
arbitrary_option \<equiv> None' |
|
1110 |
||
1111 |
text {*
|
|
1112 |
A dummy example: |
|
1113 |
*} |
|
1114 |
||
1115 |
fun |
|
1116 |
dummy_option :: "'a list \<Rightarrow> 'a option" where |
|
| 22473 | 1117 |
"dummy_option (x#xs) = Some x" |
1118 |
| "dummy_option [] = arbitrary" |
|
| 21189 | 1119 |
|
| 21545 | 1120 |
code_gen dummy_option (*<*)(SML #)(*>*)(SML "examples/arbitrary.ML") |
| 21189 | 1121 |
|
1122 |
text {*
|
|
1123 |
\lstsml{Thy/examples/arbitrary.ML}
|
|
1124 |
||
1125 |
Another axiomatic extension is code generation |
|
1126 |
for abstracted types. For this, the |
|
| 21452 | 1127 |
@{text "ExecutableRat"} (see \secref{exec_rat})
|
| 21189 | 1128 |
forms a good example. |
1129 |
*} |
|
1130 |
||
| 20948 | 1131 |
|
| 21058 | 1132 |
section {* ML interfaces \label{sec:ml} *}
|
| 20948 | 1133 |
|
| 21189 | 1134 |
text {*
|
1135 |
Since the code generator framework not only aims to provide |
|
1136 |
a nice Isar interface but also to form a base for |
|
1137 |
code-generation-based applications, here a short |
|
1138 |
description of the most important ML interfaces. |
|
1139 |
*} |
|
1140 |
||
| 21147 | 1141 |
subsection {* Constants with type discipline: codegen\_consts.ML *}
|
1142 |
||
| 21189 | 1143 |
text {*
|
1144 |
This Pure module manages identification of (probably overloaded) |
|
1145 |
constants by unique identifiers. |
|
1146 |
*} |
|
1147 |
||
| 21147 | 1148 |
text %mlref {*
|
1149 |
\begin{mldecls}
|
|
| 21323 | 1150 |
@{index_ML_type CodegenConsts.const: "string * typ list"} \\
|
| 21189 | 1151 |
@{index_ML CodegenConsts.norm_of_typ: "theory -> string * typ -> CodegenConsts.const"} \\
|
| 21147 | 1152 |
@{index_ML CodegenConsts.typ_of_inst: "theory -> CodegenConsts.const -> string * typ"} \\
|
| 21189 | 1153 |
\end{mldecls}
|
1154 |
||
1155 |
\begin{description}
|
|
1156 |
||
1157 |
\item @{ML_type CodegenConsts.const} is the identifier type:
|
|
1158 |
the product of a \emph{string} with a list of \emph{typs}.
|
|
1159 |
The \emph{string} is the constant name as represented inside Isabelle;
|
|
| 21223 | 1160 |
the \emph{typs} are a type instantiation in the sense of System F,
|
| 21189 | 1161 |
with canonical names for type variables. |
1162 |
||
1163 |
\item @{ML CodegenConsts.norm_of_typ}~@{text thy}~@{text "(constname, typ)"}
|
|
1164 |
maps a constant expression @{text "(constname, typ)"} to its canonical identifier.
|
|
1165 |
||
1166 |
\item @{ML CodegenConsts.typ_of_inst}~@{text thy}~@{text const}
|
|
1167 |
maps a canonical identifier @{text const} to a constant
|
|
1168 |
expression with appropriate type. |
|
1169 |
||
1170 |
\end{description}
|
|
| 21147 | 1171 |
*} |
1172 |
||
1173 |
subsection {* Executable theory content: codegen\_data.ML *}
|
|
1174 |
||
1175 |
text {*
|
|
1176 |
This Pure module implements the core notions of |
|
1177 |
executable content of a theory. |
|
1178 |
*} |
|
1179 |
||
1180 |
subsubsection {* Suspended theorems *}
|
|
1181 |
||
1182 |
text %mlref {*
|
|
1183 |
\begin{mldecls}
|
|
| 21341 | 1184 |
@{index_ML CodegenData.lazy: "(unit -> thm list) -> thm list Susp.T"}
|
| 21147 | 1185 |
\end{mldecls}
|
| 21189 | 1186 |
|
1187 |
\begin{description}
|
|
1188 |
||
1189 |
\item @{ML CodegenData.lazy}~@{text f} turns an abstract
|
|
| 21323 | 1190 |
theorem computation @{text f} into a suspension of theorems.
|
| 21189 | 1191 |
|
1192 |
\end{description}
|
|
| 21147 | 1193 |
*} |
1194 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1195 |
subsubsection {* Managing executable content *}
|
| 20948 | 1196 |
|
| 21147 | 1197 |
text %mlref {*
|
1198 |
\begin{mldecls}
|
|
1199 |
@{index_ML CodegenData.add_func: "thm -> theory -> theory"} \\
|
|
1200 |
@{index_ML CodegenData.del_func: "thm -> theory -> theory"} \\
|
|
| 21341 | 1201 |
@{index_ML CodegenData.add_funcl: "CodegenConsts.const * thm list Susp.T -> theory -> theory"} \\
|
| 21147 | 1202 |
@{index_ML CodegenData.add_inline: "thm -> theory -> theory"} \\
|
1203 |
@{index_ML CodegenData.del_inline: "thm -> theory -> theory"} \\
|
|
| 22046 | 1204 |
@{index_ML CodegenData.add_inline_proc: "string * (theory -> cterm list -> thm list)
|
| 21189 | 1205 |
-> theory -> theory"} \\ |
| 22046 | 1206 |
@{index_ML CodegenData.del_inline_proc: "string -> theory -> theory"} \\
|
1207 |
@{index_ML CodegenData.add_preproc: "string * (theory -> thm list -> thm list)
|
|
| 21189 | 1208 |
-> theory -> theory"} \\ |
| 22046 | 1209 |
@{index_ML CodegenData.del_preproc: "string -> theory -> theory"} \\
|
| 22423 | 1210 |
@{index_ML CodegenData.add_datatype: "string * ((string * sort) list * (string * typ list) list)
|
1211 |
-> theory -> theory"} \\ |
|
| 21189 | 1212 |
@{index_ML CodegenData.get_datatype: "theory -> string
|
| 22479 | 1213 |
-> (string * sort) list * (string * typ list) list"} \\ |
| 21147 | 1214 |
@{index_ML CodegenData.get_datatype_of_constr: "theory -> CodegenConsts.const -> string option"}
|
1215 |
\end{mldecls}
|
|
1216 |
||
1217 |
\begin{description}
|
|
1218 |
||
| 21189 | 1219 |
\item @{ML CodegenData.add_func}~@{text "thm"}~@{text "thy"} adds function
|
1220 |
theorem @{text "thm"} to executable content.
|
|
1221 |
||
1222 |
\item @{ML CodegenData.del_func}~@{text "thm"}~@{text "thy"} removes function
|
|
1223 |
theorem @{text "thm"} from executable content, if present.
|
|
1224 |
||
1225 |
\item @{ML CodegenData.add_funcl}~@{text "(const, lthms)"}~@{text "thy"} adds
|
|
| 22060 | 1226 |
suspended defining equations @{text lthms} for constant
|
| 21189 | 1227 |
@{text const} to executable content.
|
1228 |
||
1229 |
\item @{ML CodegenData.add_inline}~@{text "thm"}~@{text "thy"} adds
|
|
| 21223 | 1230 |
inlining theorem @{text thm} to executable content.
|
| 21189 | 1231 |
|
1232 |
\item @{ML CodegenData.del_inline}~@{text "thm"}~@{text "thy"} remove
|
|
1233 |
inlining theorem @{text thm} from executable content, if present.
|
|
1234 |
||
| 22046 | 1235 |
\item @{ML CodegenData.add_inline_proc}~@{text "(name, f)"}~@{text "thy"} adds
|
1236 |
inline procedure @{text f} (named @{text name}) to executable content;
|
|
| 21189 | 1237 |
@{text f} is a computation of rewrite rules dependent on
|
1238 |
the current theory context and the list of all arguments |
|
| 22060 | 1239 |
and right hand sides of the defining equations belonging |
| 21189 | 1240 |
to a certain function definition. |
1241 |
||
| 22046 | 1242 |
\item @{ML CodegenData.del_inline_proc}~@{text "name"}~@{text "thy"} removes
|
1243 |
inline procedure named @{text name} from executable content.
|
|
1244 |
||
1245 |
\item @{ML CodegenData.add_preproc}~@{text "(name, f)"}~@{text "thy"} adds
|
|
1246 |
generic preprocessor @{text f} (named @{text name}) to executable content;
|
|
| 22060 | 1247 |
@{text f} is a transformation of the defining equations belonging
|
| 21189 | 1248 |
to a certain function definition, depending on the |
1249 |
current theory context. |
|
1250 |
||
| 22060 | 1251 |
\item @{ML CodegenData.del_preproc}~@{text "name"}~@{text "thy"} removes
|
| 22046 | 1252 |
generic preprcoessor named @{text name} from executable content.
|
1253 |
||
| 22423 | 1254 |
\item @{ML CodegenData.add_datatype}~@{text "(name, spec)"}~@{text "thy"} adds
|
| 21189 | 1255 |
a datatype to executable content, with type constructor |
1256 |
@{text name} and specification @{text spec}; @{text spec} is
|
|
1257 |
a pair consisting of a list of type variable with sort |
|
| 21223 | 1258 |
constraints and a list of constructors with name |
| 22423 | 1259 |
and types of arguments. |
| 21189 | 1260 |
|
1261 |
\item @{ML CodegenData.get_datatype_of_constr}~@{text "thy"}~@{text "const"}
|
|
1262 |
returns type constructor corresponding to |
|
1263 |
constructor @{text const}; returns @{text NONE}
|
|
1264 |
if @{text const} is no constructor.
|
|
| 21147 | 1265 |
|
1266 |
\end{description}
|
|
1267 |
*} |
|
1268 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1269 |
subsection {* Auxiliary *}
|
| 21147 | 1270 |
|
1271 |
text %mlref {*
|
|
1272 |
\begin{mldecls}
|
|
1273 |
@{index_ML CodegenConsts.const_ord: "CodegenConsts.const * CodegenConsts.const -> order"} \\
|
|
1274 |
@{index_ML CodegenConsts.eq_const: "CodegenConsts.const * CodegenConsts.const -> bool"} \\
|
|
1275 |
@{index_ML CodegenConsts.consts_of: "theory -> term -> CodegenConsts.const list"} \\
|
|
1276 |
@{index_ML CodegenConsts.read_const: "theory -> string -> CodegenConsts.const"} \\
|
|
1277 |
@{index_ML_structure CodegenConsts.Consttab} \\
|
|
| 22060 | 1278 |
@{index_ML CodegenFunc.typ_func: "thm -> typ"} \\
|
1279 |
@{index_ML CodegenFunc.rewrite_func: "thm list -> thm -> thm"} \\
|
|
| 21147 | 1280 |
\end{mldecls}
|
| 21217 | 1281 |
|
1282 |
\begin{description}
|
|
1283 |
||
1284 |
\item @{ML CodegenConsts.const_ord},~@{ML CodegenConsts.eq_const}
|
|
1285 |
provide order and equality on constant identifiers. |
|
1286 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1287 |
\item @{ML_struct CodegenConsts.Consttab}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1288 |
provides table structures with constant identifiers as keys. |
| 21217 | 1289 |
|
1290 |
\item @{ML CodegenConsts.consts_of}~@{text thy}~@{text t}
|
|
1291 |
returns all constant identifiers mentioned in a term @{text t}.
|
|
1292 |
||
1293 |
\item @{ML CodegenConsts.read_const}~@{text thy}~@{text s}
|
|
1294 |
reads a constant as a concrete term expression @{text s}.
|
|
1295 |
||
| 22060 | 1296 |
\item @{ML CodegenFunc.typ_func}~@{text thm}
|
1297 |
extracts the type of a constant in a defining equation @{text thm}.
|
|
| 21217 | 1298 |
|
| 22060 | 1299 |
\item @{ML CodegenFunc.rewrite_func}~@{text rews}~@{text thm}
|
1300 |
rewrites a defining equation @{text thm} with a set of rewrite
|
|
| 21217 | 1301 |
rules @{text rews}; only arguments and right hand side are rewritten,
|
| 22060 | 1302 |
not the head of the defining equation. |
| 21217 | 1303 |
|
1304 |
\end{description}
|
|
1305 |
||
| 21147 | 1306 |
*} |
| 20948 | 1307 |
|
1308 |
subsection {* Implementing code generator applications *}
|
|
1309 |
||
| 21147 | 1310 |
text {*
|
| 21217 | 1311 |
Implementing code generator applications on top |
1312 |
of the framework set out so far usually not only |
|
1313 |
involves using those primitive interfaces |
|
1314 |
but also storing code-dependent data and various |
|
1315 |
other things. |
|
1316 |
||
| 21147 | 1317 |
\begin{warn}
|
1318 |
Some interfaces discussed here have not reached |
|
1319 |
a final state yet. |
|
1320 |
Changes likely to occur in future. |
|
1321 |
\end{warn}
|
|
1322 |
*} |
|
1323 |
||
1324 |
subsubsection {* Data depending on the theory's executable content *}
|
|
1325 |
||
| 21217 | 1326 |
text {*
|
| 21452 | 1327 |
Due to incrementality of code generation, changes in the |
1328 |
theory's executable content have to be propagated in a |
|
1329 |
certain fashion. Additionally, such changes may occur |
|
1330 |
not only during theory extension but also during theory |
|
1331 |
merge, which is a little bit nasty from an implementation |
|
1332 |
point of view. The framework provides a solution |
|
1333 |
to this technical challenge by providing a functorial |
|
1334 |
data slot @{ML_functor CodeDataFun}; on instantiation
|
|
1335 |
of this functor, the following types and operations |
|
1336 |
are required: |
|
1337 |
||
| 21217 | 1338 |
\medskip |
1339 |
\begin{tabular}{l}
|
|
1340 |
@{text "val name: string"} \\
|
|
1341 |
@{text "type T"} \\
|
|
1342 |
@{text "val empty: T"} \\
|
|
1343 |
@{text "val merge: Pretty.pp \<rightarrow> T * T \<rightarrow> T"} \\
|
|
1344 |
@{text "val purge: theory option \<rightarrow> CodegenConsts.const list option \<rightarrow> T \<rightarrow> T"}
|
|
1345 |
\end{tabular}
|
|
1346 |
||
| 21452 | 1347 |
\begin{description}
|
1348 |
||
1349 |
\item @{text name} is a system-wide unique name identifying the data.
|
|
1350 |
||
1351 |
\item @{text T} the type of data to store.
|
|
1352 |
||
1353 |
\item @{text empty} initial (empty) data.
|
|
1354 |
||
1355 |
\item @{text merge} merging two data slots.
|
|
1356 |
||
1357 |
\item @{text purge}~@{text thy}~@{text cs} propagates changes in executable content;
|
|
1358 |
if possible, the current theory context is handed over |
|
1359 |
as argument @{text thy} (if there is no current theory context (e.g.~during
|
|
1360 |
theory merge, @{ML NONE}); @{text cs} indicates the kind
|
|
1361 |
of change: @{ML NONE} stands for a fundamental change
|
|
1362 |
which invalidates any existing code, @{text "SOME cs"}
|
|
1363 |
hints that executable content for constants @{text cs}
|
|
1364 |
has changed. |
|
1365 |
||
1366 |
\end{description}
|
|
1367 |
||
1368 |
An instance of @{ML_functor CodeDataFun} provides the following
|
|
1369 |
interface: |
|
1370 |
||
| 21217 | 1371 |
\medskip |
1372 |
\begin{tabular}{l}
|
|
1373 |
@{text "init: theory \<rightarrow> theory"} \\
|
|
1374 |
@{text "get: theory \<rightarrow> T"} \\
|
|
1375 |
@{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
|
|
1376 |
@{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
|
|
1377 |
\end{tabular}
|
|
1378 |
||
1379 |
\begin{description}
|
|
1380 |
||
| 21452 | 1381 |
\item @{text init} initialization during theory setup.
|
1382 |
||
1383 |
\item @{text get} retrieval of the current data.
|
|
1384 |
||
1385 |
\item @{text change} update of current data (cached!)
|
|
1386 |
by giving a continuation. |
|
1387 |
||
1388 |
\item @{text change_yield} update with side result.
|
|
| 21217 | 1389 |
|
1390 |
\end{description}
|
|
1391 |
*} |
|
1392 |
||
|
22292
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1393 |
(* subsubsection {* Building implementable systems fo defining equations *}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1394 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1395 |
text {*
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1396 |
Out of the executable content of a theory, a normalized |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1397 |
defining equation systems may be constructed containing |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1398 |
function definitions for constants. The system is cached |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1399 |
until its underlying executable content changes. |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1400 |
*} |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1401 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1402 |
text %mlref {*
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1403 |
\begin{mldecls}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1404 |
@{index_ML_type CodegenFuncgr.T} \\
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1405 |
@{index_ML CodegenFuncgr.make: "theory -> CodegenConsts.const list -> CodegenFuncgr.T"} \\
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1406 |
@{index_ML CodegenFuncgr.funcs: "CodegenFuncgr.T -> CodegenConsts.const -> thm list"} \\
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1407 |
@{index_ML CodegenFuncgr.typ: "CodegenFuncgr.T -> CodegenConsts.const -> typ"} \\
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1408 |
@{index_ML CodegenFuncgr.deps: "CodegenFuncgr.T
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1409 |
-> CodegenConsts.const list -> CodegenConsts.const list list"} \\ |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1410 |
@{index_ML CodegenFuncgr.all: "CodegenFuncgr.T -> CodegenConsts.const list"}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1411 |
\end{mldecls}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1412 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1413 |
\begin{description}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1414 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1415 |
\item @{ML_type CodegenFuncgr.T} represents
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1416 |
a normalized defining equation system. |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1417 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1418 |
\item @{ML CodegenFuncgr.make}~@{text thy}~@{text cs}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1419 |
returns a normalized defining equation system, |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1420 |
with the assertion that it contains any function |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1421 |
definition for constants @{text cs} (if existing).
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1422 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1423 |
\item @{ML CodegenFuncgr.funcs}~@{text funcgr}~@{text c}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1424 |
retrieves function definition for constant @{text c}.
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1425 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1426 |
\item @{ML CodegenFuncgr.typ}~@{text funcgr}~@{text c}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1427 |
retrieves function type for constant @{text c}.
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1428 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1429 |
\item @{ML CodegenFuncgr.deps}~@{text funcgr}~@{text cs}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1430 |
returns the transitive closure of dependencies for |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1431 |
constants @{text cs} as a partitioning where each partition
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1432 |
corresponds to a strongly connected component of |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1433 |
dependencies and any partition does \emph{not}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1434 |
depend on partitions further left. |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1435 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1436 |
\item @{ML CodegenFuncgr.all}~@{text funcgr}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1437 |
returns all currently represented constants. |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1438 |
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1439 |
\end{description}
|
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1440 |
*} *) |
|
3b118010ec08
adjusted to new code generator Isar commands and changes in implementation
haftmann
parents:
22188
diff
changeset
|
1441 |
|
| 21147 | 1442 |
subsubsection {* Datatype hooks *}
|
1443 |
||
| 21452 | 1444 |
text {*
|
1445 |
Isabelle/HOL's datatype package provides a mechanism to |
|
1446 |
extend theories depending on datatype declarations: |
|
1447 |
\emph{datatype hooks}. For example, when declaring a new
|
|
| 22060 | 1448 |
datatype, a hook proves defining equations for equality on |
| 21452 | 1449 |
that datatype (if possible). |
1450 |
*} |
|
1451 |
||
| 21217 | 1452 |
text %mlref {*
|
1453 |
\begin{mldecls}
|
|
| 21323 | 1454 |
@{index_ML_type DatatypeHooks.hook: "string list -> theory -> theory"} \\
|
| 21217 | 1455 |
@{index_ML DatatypeHooks.add: "DatatypeHooks.hook -> theory -> theory"}
|
1456 |
\end{mldecls}
|
|
| 21452 | 1457 |
|
1458 |
\begin{description}
|
|
1459 |
||
1460 |
\item @{ML_type DatatypeHooks.hook} specifies the interface
|
|
1461 |
of \emph{datatype hooks}: a theory update
|
|
1462 |
depending on the list of newly introduced |
|
1463 |
datatype names. |
|
1464 |
||
1465 |
\item @{ML DatatypeHooks.add} adds a hook to the
|
|
1466 |
chain of all hooks. |
|
1467 |
||
1468 |
\end{description}
|
|
1469 |
*} |
|
1470 |
||
1471 |
subsubsection {* Trivial typedefs -- type copies *}
|
|
1472 |
||
1473 |
text {*
|
|
1474 |
Sometimes packages will introduce new types |
|
1475 |
as \emph{marked type copies} similar to Haskell's
|
|
1476 |
@{text newtype} declaration (e.g. the HOL record package)
|
|
1477 |
\emph{without} tinkering with the overhead of datatypes.
|
|
1478 |
Technically, these type copies are trivial forms of typedefs. |
|
1479 |
Since these type copies in code generation view are nothing |
|
1480 |
else than datatypes, they have been given a own package |
|
1481 |
in order to faciliate code generation: |
|
| 21147 | 1482 |
*} |
| 21058 | 1483 |
|
| 21217 | 1484 |
text %mlref {*
|
1485 |
\begin{mldecls}
|
|
| 21452 | 1486 |
@{index_ML_type TypecopyPackage.info} \\
|
| 21217 | 1487 |
@{index_ML TypecopyPackage.add_typecopy: "
|
1488 |
bstring * string list -> typ -> (bstring * bstring) option |
|
1489 |
-> theory -> (string * TypecopyPackage.info) * theory"} \\ |
|
1490 |
@{index_ML TypecopyPackage.get_typecopy_info: "theory
|
|
1491 |
-> string -> TypecopyPackage.info option"} \\ |
|
1492 |
@{index_ML TypecopyPackage.get_spec: "theory -> string
|
|
| 21452 | 1493 |
-> (string * sort) list * (string * typ list) list"} \\ |
1494 |
@{index_ML_type TypecopyPackage.hook: "string * TypecopyPackage.info -> theory -> theory"} \\
|
|
1495 |
@{index_ML TypecopyPackage.add_hook:
|
|
1496 |
"TypecopyPackage.hook -> theory -> theory"} \\ |
|
| 21217 | 1497 |
\end{mldecls}
|
| 21452 | 1498 |
|
1499 |
\begin{description}
|
|
1500 |
||
1501 |
\item @{ML_type TypecopyPackage.info} a record containing
|
|
1502 |
the specification and further data of a type copy. |
|
1503 |
||
1504 |
\item @{ML TypecopyPackage.add_typecopy} defines a new
|
|
1505 |
type copy. |
|
1506 |
||
1507 |
\item @{ML TypecopyPackage.get_typecopy_info} retrieves
|
|
1508 |
data of an existing type copy. |
|
1509 |
||
1510 |
\item @{ML TypecopyPackage.get_spec} retrieves datatype-like
|
|
1511 |
specification of a type copy. |
|
1512 |
||
1513 |
\item @{ML_type TypecopyPackage.hook},~@{ML TypecopyPackage.add_hook}
|
|
1514 |
provide a hook mechanism corresponding to the hook mechanism |
|
1515 |
on datatypes. |
|
1516 |
||
1517 |
\end{description}
|
|
1518 |
*} |
|
1519 |
||
1520 |
subsubsection {* Unifying type copies and datatypes *}
|
|
1521 |
||
1522 |
text {*
|
|
1523 |
Since datatypes and type copies are mapped to the same concept (datatypes) |
|
1524 |
by code generation, the view on both is unified \qt{code types}:
|
|
| 21217 | 1525 |
*} |
1526 |
||
1527 |
text %mlref {*
|
|
1528 |
\begin{mldecls}
|
|
| 21452 | 1529 |
@{index_ML_type DatatypeCodegen.hook: "(string * (bool * ((string * sort) list
|
1530 |
* (string * typ list) list))) list |
|
| 21323 | 1531 |
-> theory -> theory"} \\ |
| 21217 | 1532 |
@{index_ML DatatypeCodegen.add_codetypes_hook_bootstrap: "
|
1533 |
DatatypeCodegen.hook -> theory -> theory"} |
|
1534 |
\end{mldecls}
|
|
1535 |
*} |
|
1536 |
||
| 21222 | 1537 |
text {*
|
| 21452 | 1538 |
\begin{description}
|
1539 |
||
1540 |
\item @{ML_type DatatypeCodegen.hook} specifies the code type hook
|
|
1541 |
interface: a theory transformation depending on a list of |
|
1542 |
mutual recursive code types; each entry in the list |
|
1543 |
has the structure @{text "(name, (is_data, (vars, cons)))"}
|
|
1544 |
where @{text name} is the name of the code type, @{text is_data}
|
|
1545 |
is true iff @{text name} is a datatype rather then a type copy,
|
|
1546 |
and @{text "(vars, cons)"} is the specification of the code type.
|
|
1547 |
||
1548 |
\item @{ML DatatypeCodegen.add_codetypes_hook_bootstrap} adds a code
|
|
1549 |
type hook; the hook is immediately processed for all already |
|
1550 |
existing datatypes, in blocks of mutual recursive datatypes, |
|
1551 |
where all datatypes a block depends on are processed before |
|
1552 |
the block. |
|
1553 |
||
1554 |
\end{description}
|
|
1555 |
||
1556 |
\emph{Happy proving, happy hacking!}
|
|
| 21222 | 1557 |
*} |
| 21217 | 1558 |
|
| 20948 | 1559 |
end |