author | wenzelm |
Sun, 28 Dec 2008 23:20:57 +0100 | |
changeset 29186 | 3d25e96ceb98 |
parent 28673 | d746a8c12c43 |
child 29296 | 96cf8edb6249 |
permissions | -rw-r--r-- |
28213 | 1 |
theory "ML" |
2 |
imports Setup |
|
3 |
begin |
|
4 |
||
28419 | 5 |
section {* ML system interfaces \label{sec:ml} *} |
6 |
||
7 |
text {* |
|
8 |
Since the code generator framework not only aims to provide |
|
9 |
a nice Isar interface but also to form a base for |
|
10 |
code-generation-based applications, here a short |
|
11 |
description of the most important ML interfaces. |
|
12 |
*} |
|
13 |
||
14 |
subsection {* Executable theory content: @{text Code} *} |
|
15 |
||
16 |
text {* |
|
17 |
This Pure module implements the core notions of |
|
18 |
executable content of a theory. |
|
19 |
*} |
|
20 |
||
21 |
subsubsection {* Managing executable content *} |
|
22 |
||
23 |
text %mlref {* |
|
24 |
\begin{mldecls} |
|
25 |
@{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ |
|
26 |
@{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ |
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28635
diff
changeset
|
27 |
@{index_ML Code.add_eqnl: "string * (thm * bool) list Lazy.T -> theory -> theory"} \\ |
28419 | 28 |
@{index_ML Code.map_pre: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\ |
29 |
@{index_ML Code.map_post: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\ |
|
30 |
@{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option) |
|
31 |
-> theory -> theory"} \\ |
|
32 |
@{index_ML Code.del_functrans: "string -> theory -> theory"} \\ |
|
33 |
@{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ |
|
34 |
@{index_ML Code.get_datatype: "theory -> string |
|
35 |
-> (string * sort) list * (string * typ list) list"} \\ |
|
36 |
@{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"} |
|
37 |
\end{mldecls} |
|
38 |
||
39 |
\begin{description} |
|
40 |
||
41 |
\item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function |
|
42 |
theorem @{text "thm"} to executable content. |
|
43 |
||
44 |
\item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function |
|
45 |
theorem @{text "thm"} from executable content, if present. |
|
46 |
||
47 |
\item @{ML Code.add_eqnl}~@{text "(const, lthms)"}~@{text "thy"} adds |
|
48 |
suspended defining equations @{text lthms} for constant |
|
49 |
@{text const} to executable content. |
|
50 |
||
51 |
\item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes |
|
52 |
the preprocessor simpset. |
|
53 |
||
54 |
\item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds |
|
55 |
function transformer @{text f} (named @{text name}) to executable content; |
|
56 |
@{text f} is a transformer of the defining equations belonging |
|
57 |
to a certain function definition, depending on the |
|
58 |
current theory context. Returning @{text NONE} indicates that no |
|
59 |
transformation took place; otherwise, the whole process will be iterated |
|
60 |
with the new defining equations. |
|
61 |
||
62 |
\item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes |
|
63 |
function transformer named @{text name} from executable content. |
|
64 |
||
65 |
\item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds |
|
66 |
a datatype to executable content, with generation |
|
67 |
set @{text cs}. |
|
68 |
||
69 |
\item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"} |
|
70 |
returns type constructor corresponding to |
|
71 |
constructor @{text const}; returns @{text NONE} |
|
72 |
if @{text const} is no constructor. |
|
73 |
||
74 |
\end{description} |
|
75 |
*} |
|
76 |
||
77 |
subsection {* Auxiliary *} |
|
78 |
||
79 |
text %mlref {* |
|
80 |
\begin{mldecls} |
|
81 |
@{index_ML Code_Unit.read_const: "theory -> string -> string"} \\ |
|
82 |
@{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\ |
|
83 |
@{index_ML Code_Unit.rewrite_eqn: "MetaSimplifier.simpset -> thm -> thm"} \\ |
|
84 |
\end{mldecls} |
|
85 |
||
86 |
\begin{description} |
|
87 |
||
88 |
\item @{ML Code_Unit.read_const}~@{text thy}~@{text s} |
|
89 |
reads a constant as a concrete term expression @{text s}. |
|
90 |
||
91 |
\item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm} |
|
92 |
extracts the constant and its type from a defining equation @{text thm}. |
|
93 |
||
94 |
\item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm} |
|
95 |
rewrites a defining equation @{text thm} with a simpset @{text ss}; |
|
96 |
only arguments and right hand side are rewritten, |
|
97 |
not the head of the defining equation. |
|
98 |
||
99 |
\end{description} |
|
100 |
||
101 |
*} |
|
102 |
||
103 |
subsection {* Implementing code generator applications *} |
|
104 |
||
105 |
text {* |
|
106 |
Implementing code generator applications on top |
|
107 |
of the framework set out so far usually not only |
|
108 |
involves using those primitive interfaces |
|
109 |
but also storing code-dependent data and various |
|
110 |
other things. |
|
111 |
*} |
|
112 |
||
113 |
subsubsection {* Data depending on the theory's executable content *} |
|
114 |
||
115 |
text {* |
|
116 |
Due to incrementality of code generation, changes in the |
|
117 |
theory's executable content have to be propagated in a |
|
118 |
certain fashion. Additionally, such changes may occur |
|
119 |
not only during theory extension but also during theory |
|
120 |
merge, which is a little bit nasty from an implementation |
|
121 |
point of view. The framework provides a solution |
|
122 |
to this technical challenge by providing a functorial |
|
123 |
data slot @{ML_functor CodeDataFun}; on instantiation |
|
124 |
of this functor, the following types and operations |
|
125 |
are required: |
|
126 |
||
127 |
\medskip |
|
128 |
\begin{tabular}{l} |
|
129 |
@{text "type T"} \\ |
|
130 |
@{text "val empty: T"} \\ |
|
28635 | 131 |
@{text "val purge: theory \<rightarrow> string list option \<rightarrow> T \<rightarrow> T"} |
28419 | 132 |
\end{tabular} |
133 |
||
134 |
\begin{description} |
|
135 |
||
136 |
\item @{text T} the type of data to store. |
|
137 |
||
138 |
\item @{text empty} initial (empty) data. |
|
139 |
||
140 |
\item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content; |
|
141 |
@{text consts} indicates the kind |
|
142 |
of change: @{ML NONE} stands for a fundamental change |
|
143 |
which invalidates any existing code, @{text "SOME consts"} |
|
144 |
hints that executable content for constants @{text consts} |
|
145 |
has changed. |
|
146 |
||
147 |
\end{description} |
|
148 |
||
28447 | 149 |
\noindent An instance of @{ML_functor CodeDataFun} provides the following |
28419 | 150 |
interface: |
151 |
||
152 |
\medskip |
|
153 |
\begin{tabular}{l} |
|
154 |
@{text "get: theory \<rightarrow> T"} \\ |
|
155 |
@{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\ |
|
156 |
@{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"} |
|
157 |
\end{tabular} |
|
158 |
||
159 |
\begin{description} |
|
160 |
||
161 |
\item @{text get} retrieval of the current data. |
|
162 |
||
163 |
\item @{text change} update of current data (cached!) |
|
164 |
by giving a continuation. |
|
165 |
||
166 |
\item @{text change_yield} update with side result. |
|
167 |
||
168 |
\end{description} |
|
169 |
*} |
|
170 |
||
171 |
text {* |
|
28447 | 172 |
\bigskip |
173 |
||
28419 | 174 |
\emph{Happy proving, happy hacking!} |
175 |
*} |
|
28213 | 176 |
|
177 |
end |