author | haftmann |
Thu, 16 Dec 2010 09:26:46 +0100 | |
changeset 41187 | b0b975e197b5 |
parent 40752 | aae9a020fa77 |
child 42096 | 9f6652122963 |
permissions | -rw-r--r-- |
28213 | 1 |
theory Further |
2 |
imports Setup |
|
3 |
begin |
|
4 |
||
28447 | 5 |
section {* Further issues \label{sec:further} *} |
28419 | 6 |
|
38437 | 7 |
subsection {* Modules namespace *} |
28419 | 8 |
|
9 |
text {* |
|
40752 | 10 |
When invoking the @{command export_code} command it is possible to |
11 |
leave out the @{keyword "module_name"} part; then code is |
|
12 |
distributed over different modules, where the module name space |
|
13 |
roughly is induced by the Isabelle theory name space. |
|
38437 | 14 |
|
40752 | 15 |
Then sometimes the awkward situation occurs that dependencies |
16 |
between definitions introduce cyclic dependencies between modules, |
|
17 |
which in the @{text Haskell} world leaves you to the mercy of the |
|
18 |
@{text Haskell} implementation you are using, while for @{text |
|
19 |
SML}/@{text OCaml} code generation is not possible. |
|
38437 | 20 |
|
40752 | 21 |
A solution is to declare module names explicitly. Let use assume |
22 |
the three cyclically dependent modules are named \emph{A}, \emph{B} |
|
23 |
and \emph{C}. Then, by stating |
|
38437 | 24 |
*} |
25 |
||
26 |
code_modulename %quote SML |
|
27 |
A ABC |
|
28 |
B ABC |
|
29 |
C ABC |
|
30 |
||
40752 | 31 |
text {* |
32 |
\noindent we explicitly map all those modules on \emph{ABC}, |
|
33 |
resulting in an ad-hoc merge of this three modules at serialisation |
|
34 |
time. |
|
28419 | 35 |
*} |
36 |
||
37426 | 37 |
subsection {* Locales and interpretation *} |
38 |
||
39 |
text {* |
|
40 |
A technical issue comes to surface when generating code from |
|
41 |
specifications stemming from locale interpretation. |
|
42 |
||
40752 | 43 |
Let us assume a locale specifying a power operation on arbitrary |
44 |
types: |
|
37426 | 45 |
*} |
46 |
||
47 |
locale %quote power = |
|
48 |
fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" |
|
49 |
assumes power_commute: "power x \<circ> power y = power y \<circ> power x" |
|
50 |
begin |
|
51 |
||
52 |
text {* |
|
40752 | 53 |
\noindent Inside that locale we can lift @{text power} to exponent |
54 |
lists by means of specification relative to that locale: |
|
37426 | 55 |
*} |
56 |
||
57 |
primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where |
|
58 |
"powers [] = id" |
|
59 |
| "powers (x # xs) = power x \<circ> powers xs" |
|
60 |
||
61 |
lemma %quote powers_append: |
|
62 |
"powers (xs @ ys) = powers xs \<circ> powers ys" |
|
63 |
by (induct xs) simp_all |
|
64 |
||
65 |
lemma %quote powers_power: |
|
66 |
"powers xs \<circ> power x = power x \<circ> powers xs" |
|
67 |
by (induct xs) |
|
68 |
(simp_all del: o_apply id_apply add: o_assoc [symmetric], |
|
69 |
simp del: o_apply add: o_assoc power_commute) |
|
70 |
||
71 |
lemma %quote powers_rev: |
|
72 |
"powers (rev xs) = powers xs" |
|
73 |
by (induct xs) (simp_all add: powers_append powers_power) |
|
74 |
||
75 |
end %quote |
|
76 |
||
77 |
text {* |
|
38505 | 78 |
After an interpretation of this locale (say, @{command_def |
40752 | 79 |
interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f |
80 |
:: 'a \<Rightarrow> 'a). f ^^ n)"}), one would expect to have a constant @{text |
|
37426 | 81 |
"fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code |
82 |
can be generated. But this not the case: internally, the term |
|
83 |
@{text "fun_power.powers"} is an abbreviation for the foundational |
|
84 |
term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"} |
|
85 |
(see \cite{isabelle-locale} for the details behind). |
|
86 |
||
40752 | 87 |
Fortunately, with minor effort the desired behaviour can be |
88 |
achieved. First, a dedicated definition of the constant on which |
|
89 |
the local @{text "powers"} after interpretation is supposed to be |
|
90 |
mapped on: |
|
37426 | 91 |
*} |
92 |
||
93 |
definition %quote funpows :: "nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where |
|
94 |
[code del]: "funpows = power.powers (\<lambda>n f. f ^^ n)" |
|
95 |
||
96 |
text {* |
|
40752 | 97 |
\noindent In general, the pattern is @{text "c = t"} where @{text c} |
98 |
is the name of the future constant and @{text t} the foundational |
|
99 |
term corresponding to the local constant after interpretation. |
|
37426 | 100 |
|
101 |
The interpretation itself is enriched with an equation @{text "t = c"}: |
|
102 |
*} |
|
103 |
||
104 |
interpretation %quote fun_power: power "\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n" where |
|
105 |
"power.powers (\<lambda>n f. f ^^ n) = funpows" |
|
106 |
by unfold_locales |
|
39559 | 107 |
(simp_all add: fun_eq_iff funpow_mult mult_commute funpows_def) |
37426 | 108 |
|
109 |
text {* |
|
40752 | 110 |
\noindent This additional equation is trivially proved by the |
111 |
definition itself. |
|
37426 | 112 |
|
113 |
After this setup procedure, code generation can continue as usual: |
|
114 |
*} |
|
115 |
||
39745 | 116 |
text %quotetypewriter {* |
39683 | 117 |
@{code_stmts funpows (consts) Nat.funpow funpows (Haskell)} |
39664
0afaf89ab591
more canonical type setting of type writer code examples
haftmann
parents:
39559
diff
changeset
|
118 |
*} |
37426 | 119 |
|
120 |
||
38437 | 121 |
subsection {* Imperative data structures *} |
28419 | 122 |
|
123 |
text {* |
|
38437 | 124 |
If you consider imperative data structures as inevitable for a |
125 |
specific application, you should consider \emph{Imperative |
|
126 |
Functional Programming with Isabelle/HOL} |
|
127 |
\cite{bulwahn-et-al:2008:imperative}; the framework described there |
|
40752 | 128 |
is available in session @{text Imperative_HOL}, together with a |
129 |
short primer document. |
|
28419 | 130 |
*} |
131 |
||
28213 | 132 |
|
38405 | 133 |
subsection {* ML system interfaces \label{sec:ml} *} |
134 |
||
135 |
text {* |
|
38510 | 136 |
Since the code generator framework not only aims to provide a nice |
137 |
Isar interface but also to form a base for code-generation-based |
|
138 |
applications, here a short description of the most fundamental ML |
|
139 |
interfaces. |
|
38405 | 140 |
*} |
141 |
||
142 |
subsubsection {* Managing executable content *} |
|
143 |
||
144 |
text %mlref {* |
|
145 |
\begin{mldecls} |
|
38510 | 146 |
@{index_ML Code.read_const: "theory -> string -> string"} \\ |
38405 | 147 |
@{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ |
148 |
@{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ |
|
149 |
@{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\ |
|
150 |
@{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\ |
|
38507 | 151 |
@{index_ML Code_Preproc.add_functrans: " |
152 |
string * (theory -> (thm * bool) list -> (thm * bool) list option) |
|
153 |
-> theory -> theory"} \\ |
|
38405 | 154 |
@{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\ |
155 |
@{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ |
|
156 |
@{index_ML Code.get_type: "theory -> string |
|
40752 | 157 |
-> ((string * sort) list * (string * ((string * sort) list * typ list)) list) * bool"} \\ |
38405 | 158 |
@{index_ML Code.get_type_of_constr_or_abstr: "theory -> string -> (string * bool) option"} |
159 |
\end{mldecls} |
|
160 |
||
161 |
\begin{description} |
|
162 |
||
38510 | 163 |
\item @{ML Code.read_const}~@{text thy}~@{text s} |
164 |
reads a constant as a concrete term expression @{text s}. |
|
165 |
||
38405 | 166 |
\item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function |
167 |
theorem @{text "thm"} to executable content. |
|
168 |
||
169 |
\item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function |
|
170 |
theorem @{text "thm"} from executable content, if present. |
|
171 |
||
172 |
\item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes |
|
173 |
the preprocessor simpset. |
|
174 |
||
175 |
\item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds |
|
176 |
function transformer @{text f} (named @{text name}) to executable content; |
|
177 |
@{text f} is a transformer of the code equations belonging |
|
178 |
to a certain function definition, depending on the |
|
179 |
current theory context. Returning @{text NONE} indicates that no |
|
180 |
transformation took place; otherwise, the whole process will be iterated |
|
181 |
with the new code equations. |
|
182 |
||
183 |
\item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes |
|
184 |
function transformer named @{text name} from executable content. |
|
185 |
||
186 |
\item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds |
|
187 |
a datatype to executable content, with generation |
|
188 |
set @{text cs}. |
|
189 |
||
190 |
\item @{ML Code.get_type_of_constr_or_abstr}~@{text "thy"}~@{text "const"} |
|
191 |
returns type constructor corresponding to |
|
192 |
constructor @{text const}; returns @{text NONE} |
|
193 |
if @{text const} is no constructor. |
|
194 |
||
195 |
\end{description} |
|
196 |
*} |
|
197 |
||
198 |
||
199 |
subsubsection {* Data depending on the theory's executable content *} |
|
200 |
||
201 |
text {* |
|
40752 | 202 |
Implementing code generator applications on top of the framework set |
203 |
out so far usually not only involves using those primitive |
|
204 |
interfaces but also storing code-dependent data and various other |
|
205 |
things. |
|
38405 | 206 |
|
40752 | 207 |
Due to incrementality of code generation, changes in the theory's |
208 |
executable content have to be propagated in a certain fashion. |
|
209 |
Additionally, such changes may occur not only during theory |
|
210 |
extension but also during theory merge, which is a little bit nasty |
|
211 |
from an implementation point of view. The framework provides a |
|
212 |
solution to this technical challenge by providing a functorial data |
|
213 |
slot @{ML_functor Code_Data}; on instantiation of this functor, the |
|
214 |
following types and operations are required: |
|
38405 | 215 |
|
216 |
\medskip |
|
217 |
\begin{tabular}{l} |
|
218 |
@{text "type T"} \\ |
|
219 |
@{text "val empty: T"} \\ |
|
220 |
\end{tabular} |
|
221 |
||
222 |
\begin{description} |
|
223 |
||
224 |
\item @{text T} the type of data to store. |
|
225 |
||
226 |
\item @{text empty} initial (empty) data. |
|
227 |
||
228 |
\end{description} |
|
229 |
||
40752 | 230 |
\noindent An instance of @{ML_functor Code_Data} provides the |
231 |
following interface: |
|
38405 | 232 |
|
233 |
\medskip |
|
234 |
\begin{tabular}{l} |
|
235 |
@{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\ |
|
236 |
@{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"} |
|
237 |
\end{tabular} |
|
238 |
||
239 |
\begin{description} |
|
240 |
||
40752 | 241 |
\item @{text change} update of current data (cached!) by giving a |
242 |
continuation. |
|
38405 | 243 |
|
244 |
\item @{text change_yield} update with side result. |
|
245 |
||
246 |
\end{description} |
|
247 |
*} |
|
248 |
||
28213 | 249 |
end |