author | wenzelm |
Fri, 23 Jun 2017 14:38:32 +0200 | |
changeset 66176 | b51a40281016 |
parent 65981 | e2c25346b156 |
child 66251 | cd935b7cb3fb |
permissions | -rw-r--r-- |
28213 | 1 |
theory Further |
59378
065f349852e6
separate image for prerequisites of codegen tutorial
haftmann
parents:
59377
diff
changeset
|
2 |
imports Setup |
28213 | 3 |
begin |
4 |
||
59377 | 5 |
section \<open>Further issues \label{sec:further}\<close> |
28419 | 6 |
|
65041 | 7 |
subsection \<open>Incorporating generated code directly into the system runtime -- @{text code_reflect}\<close> |
8 |
||
9 |
subsubsection \<open>Static embedding of generated code into the system runtime\<close> |
|
10 |
||
11 |
text \<open> |
|
12 |
The @{ML_antiquotation code} antiquotation is lightweight, but the generated code |
|
13 |
is only accessible while the ML section is processed. Sometimes this |
|
14 |
is not appropriate, especially if the generated code contains datatype |
|
15 |
declarations which are shared with other parts of the system. In these |
|
16 |
cases, @{command_def code_reflect} can be used: |
|
17 |
\<close> |
|
18 |
||
19 |
code_reflect %quote Sum_Type |
|
20 |
datatypes sum = Inl | Inr |
|
21 |
functions "Sum_Type.sum.projl" "Sum_Type.sum.projr" |
|
22 |
||
23 |
text \<open> |
|
24 |
\noindent @{command code_reflect} takes a structure name and |
|
25 |
references to datatypes and functions; for these code is compiled |
|
26 |
into the named ML structure and the \emph{Eval} target is modified |
|
27 |
in a way that future code generation will reference these |
|
28 |
precompiled versions of the given datatypes and functions. This |
|
29 |
also allows to refer to the referenced datatypes and functions from |
|
30 |
arbitrary ML code as well. |
|
31 |
||
32 |
A typical example for @{command code_reflect} can be found in the |
|
33 |
@{theory Predicate} theory. |
|
34 |
\<close> |
|
35 |
||
36 |
||
37 |
subsubsection \<open>Separate compilation\<close> |
|
38 |
||
39 |
text \<open> |
|
40 |
For technical reasons it is sometimes necessary to separate |
|
41 |
generation and compilation of code which is supposed to be used in |
|
42 |
the system runtime. For this @{command code_reflect} with an |
|
43 |
optional \<^theory_text>\<open>file\<close> argument can be used: |
|
44 |
\<close> |
|
45 |
||
46 |
code_reflect %quote Rat |
|
47 |
datatypes rat |
|
48 |
functions Fract |
|
49 |
"(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)" |
|
50 |
"(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)" |
|
65981
e2c25346b156
more robust -- avoid race condition wrt. Haskell output in $ISABELLE_TMP/examples/
wenzelm
parents:
65041
diff
changeset
|
51 |
file "$ISABELLE_TMP/rat.ML" |
65041 | 52 |
|
53 |
text \<open> |
|
54 |
\noindent This merely generates the referenced code to the given |
|
55 |
file which can be included into the system runtime later on. |
|
56 |
\<close> |
|
57 |
||
59377 | 58 |
subsection \<open>Specialities of the @{text Scala} target language \label{sec:scala}\<close> |
42096 | 59 |
|
59377 | 60 |
text \<open> |
42096 | 61 |
@{text Scala} deviates from languages of the ML family in a couple |
62 |
of aspects; those which affect code generation mainly have to do with |
|
63 |
@{text Scala}'s type system: |
|
64 |
||
65 |
\begin{itemize} |
|
66 |
||
67 |
\item @{text Scala} prefers tupled syntax over curried syntax. |
|
68 |
||
69 |
\item @{text Scala} sacrifices Hindely-Milner type inference for a |
|
70 |
much more rich type system with subtyping etc. For this reason |
|
71 |
type arguments sometimes have to be given explicitly in square |
|
46520 | 72 |
brackets (mimicking System F syntax). |
42096 | 73 |
|
74 |
\item In contrast to @{text Haskell} where most specialities of |
|
75 |
the type system are implemented using \emph{type classes}, |
|
76 |
@{text Scala} provides a sophisticated system of \emph{implicit |
|
77 |
arguments}. |
|
78 |
||
79 |
\end{itemize} |
|
80 |
||
81 |
\noindent Concerning currying, the @{text Scala} serializer counts |
|
82 |
arguments in code equations to determine how many arguments |
|
83 |
shall be tupled; remaining arguments and abstractions in terms |
|
84 |
rather than function definitions are always curried. |
|
85 |
||
86 |
The second aspect affects user-defined adaptations with @{command |
|
55372 | 87 |
code_printing}. For regular terms, the @{text Scala} serializer prints |
42096 | 88 |
all type arguments explicitly. For user-defined term adaptations |
89 |
this is only possible for adaptations which take no arguments: here |
|
90 |
the type arguments are just appended. Otherwise they are ignored; |
|
91 |
hence user-defined adaptations for polymorphic constants have to be |
|
92 |
designed very carefully to avoid ambiguity. |
|
59377 | 93 |
\<close> |
42096 | 94 |
|
95 |
||
59377 | 96 |
subsection \<open>Modules namespace\<close> |
28419 | 97 |
|
59377 | 98 |
text \<open> |
40752 | 99 |
When invoking the @{command export_code} command it is possible to |
100 |
leave out the @{keyword "module_name"} part; then code is |
|
101 |
distributed over different modules, where the module name space |
|
102 |
roughly is induced by the Isabelle theory name space. |
|
38437 | 103 |
|
40752 | 104 |
Then sometimes the awkward situation occurs that dependencies |
105 |
between definitions introduce cyclic dependencies between modules, |
|
106 |
which in the @{text Haskell} world leaves you to the mercy of the |
|
107 |
@{text Haskell} implementation you are using, while for @{text |
|
108 |
SML}/@{text OCaml} code generation is not possible. |
|
38437 | 109 |
|
40752 | 110 |
A solution is to declare module names explicitly. Let use assume |
111 |
the three cyclically dependent modules are named \emph{A}, \emph{B} |
|
112 |
and \emph{C}. Then, by stating |
|
59377 | 113 |
\<close> |
38437 | 114 |
|
52378
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
51717
diff
changeset
|
115 |
code_identifier %quote |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
51717
diff
changeset
|
116 |
code_module A \<rightharpoonup> (SML) ABC |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
51717
diff
changeset
|
117 |
| code_module B \<rightharpoonup> (SML) ABC |
08dbf9ff2140
documentation on code_printing and code_identifier
haftmann
parents:
51717
diff
changeset
|
118 |
| code_module C \<rightharpoonup> (SML) ABC |
38437 | 119 |
|
59377 | 120 |
text \<open> |
40752 | 121 |
\noindent we explicitly map all those modules on \emph{ABC}, |
122 |
resulting in an ad-hoc merge of this three modules at serialisation |
|
123 |
time. |
|
59377 | 124 |
\<close> |
28419 | 125 |
|
59377 | 126 |
subsection \<open>Locales and interpretation\<close> |
37426 | 127 |
|
59377 | 128 |
text \<open> |
37426 | 129 |
A technical issue comes to surface when generating code from |
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61779
diff
changeset
|
130 |
specifications stemming from locale interpretation into global |
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61779
diff
changeset
|
131 |
theories. |
37426 | 132 |
|
40752 | 133 |
Let us assume a locale specifying a power operation on arbitrary |
134 |
types: |
|
59377 | 135 |
\<close> |
37426 | 136 |
|
137 |
locale %quote power = |
|
138 |
fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" |
|
139 |
assumes power_commute: "power x \<circ> power y = power y \<circ> power x" |
|
140 |
begin |
|
141 |
||
59377 | 142 |
text \<open> |
40752 | 143 |
\noindent Inside that locale we can lift @{text power} to exponent |
61779 | 144 |
lists by means of a specification relative to that locale: |
59377 | 145 |
\<close> |
37426 | 146 |
|
147 |
primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where |
|
148 |
"powers [] = id" |
|
149 |
| "powers (x # xs) = power x \<circ> powers xs" |
|
150 |
||
151 |
lemma %quote powers_append: |
|
152 |
"powers (xs @ ys) = powers xs \<circ> powers ys" |
|
153 |
by (induct xs) simp_all |
|
154 |
||
155 |
lemma %quote powers_power: |
|
156 |
"powers xs \<circ> power x = power x \<circ> powers xs" |
|
157 |
by (induct xs) |
|
49739 | 158 |
(simp_all del: o_apply id_apply add: comp_assoc, |
37426 | 159 |
simp del: o_apply add: o_assoc power_commute) |
160 |
||
161 |
lemma %quote powers_rev: |
|
162 |
"powers (rev xs) = powers xs" |
|
163 |
by (induct xs) (simp_all add: powers_append powers_power) |
|
164 |
||
165 |
end %quote |
|
166 |
||
59377 | 167 |
text \<open> |
38505 | 168 |
After an interpretation of this locale (say, @{command_def |
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61779
diff
changeset
|
169 |
global_interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f |
61779 | 170 |
:: 'a \<Rightarrow> 'a). f ^^ n)"}), one could naively expect to have a constant @{text |
37426 | 171 |
"fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code |
172 |
can be generated. But this not the case: internally, the term |
|
173 |
@{text "fun_power.powers"} is an abbreviation for the foundational |
|
174 |
term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"} |
|
58620 | 175 |
(see @{cite "isabelle-locale"} for the details behind). |
37426 | 176 |
|
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61779
diff
changeset
|
177 |
Fortunately, a succint solution is available: a dedicated |
61779 | 178 |
rewrite definition: |
59377 | 179 |
\<close> |
37426 | 180 |
|
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61779
diff
changeset
|
181 |
global_interpretation %quote fun_power: power "(\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)" |
61779 | 182 |
defines funpows = fun_power.powers |
183 |
by unfold_locales |
|
184 |
(simp_all add: fun_eq_iff funpow_mult mult.commute) |
|
37426 | 185 |
|
59377 | 186 |
text \<open> |
61779 | 187 |
\noindent This amends the interpretation morphisms such that |
188 |
occurrences of the foundational term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"} |
|
189 |
are folded to a newly defined constant @{const funpows}. |
|
37426 | 190 |
|
191 |
After this setup procedure, code generation can continue as usual: |
|
59377 | 192 |
\<close> |
37426 | 193 |
|
59377 | 194 |
text %quotetypewriter \<open> |
39683 | 195 |
@{code_stmts funpows (consts) Nat.funpow funpows (Haskell)} |
61779 | 196 |
\<close> |
37426 | 197 |
|
198 |
||
59377 | 199 |
subsection \<open>Parallel computation\<close> |
51172 | 200 |
|
59377 | 201 |
text \<open> |
63680 | 202 |
Theory @{text Parallel} in \<^dir>\<open>~~/src/HOL/Library\<close> contains |
51172 | 203 |
operations to exploit parallelism inside the Isabelle/ML |
204 |
runtime engine. |
|
59377 | 205 |
\<close> |
51172 | 206 |
|
59377 | 207 |
subsection \<open>Imperative data structures\<close> |
28419 | 208 |
|
59377 | 209 |
text \<open> |
38437 | 210 |
If you consider imperative data structures as inevitable for a |
211 |
specific application, you should consider \emph{Imperative |
|
212 |
Functional Programming with Isabelle/HOL} |
|
58620 | 213 |
@{cite "bulwahn-et-al:2008:imperative"}; the framework described there |
40752 | 214 |
is available in session @{text Imperative_HOL}, together with a |
215 |
short primer document. |
|
59377 | 216 |
\<close> |
28419 | 217 |
|
28213 | 218 |
|
59377 | 219 |
subsection \<open>ML system interfaces \label{sec:ml}\<close> |
38405 | 220 |
|
59377 | 221 |
text \<open> |
38510 | 222 |
Since the code generator framework not only aims to provide a nice |
223 |
Isar interface but also to form a base for code-generation-based |
|
224 |
applications, here a short description of the most fundamental ML |
|
225 |
interfaces. |
|
59377 | 226 |
\<close> |
38405 | 227 |
|
59377 | 228 |
subsubsection \<open>Managing executable content\<close> |
38405 | 229 |
|
59377 | 230 |
text %mlref \<open> |
38405 | 231 |
\begin{mldecls} |
38510 | 232 |
@{index_ML Code.read_const: "theory -> string -> string"} \\ |
63239
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
61891
diff
changeset
|
233 |
@{index_ML Code.add_eqn: "Code.kind * bool -> thm -> theory -> theory"} \\ |
38405 | 234 |
@{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51172
diff
changeset
|
235 |
@{index_ML Code_Preproc.map_pre: "(Proof.context -> Proof.context) -> theory -> theory"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51172
diff
changeset
|
236 |
@{index_ML Code_Preproc.map_post: "(Proof.context -> Proof.context) -> theory -> theory"} \\ |
38507 | 237 |
@{index_ML Code_Preproc.add_functrans: " |
55757 | 238 |
string * (Proof.context -> (thm * bool) list -> (thm * bool) list option) |
38507 | 239 |
-> theory -> theory"} \\ |
38405 | 240 |
@{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\ |
241 |
@{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ |
|
242 |
@{index_ML Code.get_type: "theory -> string |
|
40752 | 243 |
-> ((string * sort) list * (string * ((string * sort) list * typ list)) list) * bool"} \\ |
38405 | 244 |
@{index_ML Code.get_type_of_constr_or_abstr: "theory -> string -> (string * bool) option"} |
245 |
\end{mldecls} |
|
246 |
||
247 |
\begin{description} |
|
248 |
||
38510 | 249 |
\item @{ML Code.read_const}~@{text thy}~@{text s} |
250 |
reads a constant as a concrete term expression @{text s}. |
|
251 |
||
63239
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
61891
diff
changeset
|
252 |
\item @{ML Code.add_eqn}~@{text "(kind, default)"}~@{text "thm"}~@{text "thy"} adds code equation |
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
61891
diff
changeset
|
253 |
@{text "thm"} to executable content. |
38405 | 254 |
|
63239
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
61891
diff
changeset
|
255 |
\item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes code equation |
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
61891
diff
changeset
|
256 |
@{text "thm"} from executable content, if present. |
38405 | 257 |
|
258 |
\item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes |
|
259 |
the preprocessor simpset. |
|
260 |
||
261 |
\item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds |
|
262 |
function transformer @{text f} (named @{text name}) to executable content; |
|
263 |
@{text f} is a transformer of the code equations belonging |
|
264 |
to a certain function definition, depending on the |
|
265 |
current theory context. Returning @{text NONE} indicates that no |
|
266 |
transformation took place; otherwise, the whole process will be iterated |
|
267 |
with the new code equations. |
|
268 |
||
269 |
\item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes |
|
270 |
function transformer named @{text name} from executable content. |
|
271 |
||
272 |
\item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds |
|
273 |
a datatype to executable content, with generation |
|
274 |
set @{text cs}. |
|
275 |
||
276 |
\item @{ML Code.get_type_of_constr_or_abstr}~@{text "thy"}~@{text "const"} |
|
277 |
returns type constructor corresponding to |
|
278 |
constructor @{text const}; returns @{text NONE} |
|
279 |
if @{text const} is no constructor. |
|
280 |
||
281 |
\end{description} |
|
59377 | 282 |
\<close> |
38405 | 283 |
|
284 |
||
59377 | 285 |
subsubsection \<open>Data depending on the theory's executable content\<close> |
38405 | 286 |
|
59377 | 287 |
text \<open> |
40752 | 288 |
Implementing code generator applications on top of the framework set |
289 |
out so far usually not only involves using those primitive |
|
290 |
interfaces but also storing code-dependent data and various other |
|
291 |
things. |
|
38405 | 292 |
|
40752 | 293 |
Due to incrementality of code generation, changes in the theory's |
294 |
executable content have to be propagated in a certain fashion. |
|
295 |
Additionally, such changes may occur not only during theory |
|
296 |
extension but also during theory merge, which is a little bit nasty |
|
297 |
from an implementation point of view. The framework provides a |
|
298 |
solution to this technical challenge by providing a functorial data |
|
299 |
slot @{ML_functor Code_Data}; on instantiation of this functor, the |
|
300 |
following types and operations are required: |
|
38405 | 301 |
|
302 |
\medskip |
|
303 |
\begin{tabular}{l} |
|
304 |
@{text "type T"} \\ |
|
305 |
@{text "val empty: T"} \\ |
|
306 |
\end{tabular} |
|
307 |
||
308 |
\begin{description} |
|
309 |
||
310 |
\item @{text T} the type of data to store. |
|
311 |
||
312 |
\item @{text empty} initial (empty) data. |
|
313 |
||
314 |
\end{description} |
|
315 |
||
40752 | 316 |
\noindent An instance of @{ML_functor Code_Data} provides the |
317 |
following interface: |
|
38405 | 318 |
|
319 |
\medskip |
|
320 |
\begin{tabular}{l} |
|
321 |
@{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\ |
|
322 |
@{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"} |
|
323 |
\end{tabular} |
|
324 |
||
325 |
\begin{description} |
|
326 |
||
40752 | 327 |
\item @{text change} update of current data (cached!) by giving a |
328 |
continuation. |
|
38405 | 329 |
|
330 |
\item @{text change_yield} update with side result. |
|
331 |
||
332 |
\end{description} |
|
59377 | 333 |
\<close> |
38405 | 334 |
|
28213 | 335 |
end |