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