| author | haftmann |
| Mon, 20 Apr 2009 16:28:13 +0200 | |
| changeset 30957 | 20d01210b9b1 |
| parent 30928 | 983dfcce45ad |
| child 31088 | 36a011423fcc |
| permissions | -rw-r--r-- |
| 24219 | 1 |
(* Title: Pure/Isar/code.ML |
2 |
Author: Florian Haftmann, TU Muenchen |
|
3 |
||
4 |
Abstract executable content of theory. Management of data dependent on |
|
| 25485 | 5 |
executable content. Cache assumes non-concurrent processing of a single theory. |
| 24219 | 6 |
*) |
7 |
||
8 |
signature CODE = |
|
9 |
sig |
|
| 28368 | 10 |
val add_eqn: thm -> theory -> theory |
11 |
val add_nonlinear_eqn: thm -> theory -> theory |
|
12 |
val add_default_eqn: thm -> theory -> theory |
|
| 28703 | 13 |
val add_default_eqn_attribute: attribute |
14 |
val add_default_eqn_attrib: Attrib.src |
|
| 28368 | 15 |
val del_eqn: thm -> theory -> theory |
16 |
val del_eqns: string -> theory -> theory |
|
| 28971 | 17 |
val add_eqnl: string * (thm * bool) list lazy -> theory -> theory |
| 29302 | 18 |
val map_pre: (simpset -> simpset) -> theory -> theory |
19 |
val map_post: (simpset -> simpset) -> theory -> theory |
|
| 24219 | 20 |
val add_inline: thm -> theory -> theory |
| 28423 | 21 |
val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory |
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
22 |
val del_functrans: string -> theory -> theory |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
23 |
val add_datatype: (string * typ) list -> theory -> theory |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
24 |
val add_datatype_cmd: string list -> theory -> theory |
| 25485 | 25 |
val type_interpretation: |
26 |
(string * ((string * sort) list * (string * typ list) list) |
|
27 |
-> theory -> theory) -> theory -> theory |
|
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
28 |
val add_case: thm -> theory -> theory |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
29 |
val add_undefined: string -> theory -> theory |
| 27675 | 30 |
val purge_data: theory -> theory |
| 24219 | 31 |
|
| 28368 | 32 |
val these_eqns: theory -> string -> (thm * bool) list |
| 28525 | 33 |
val these_raw_eqns: theory -> string -> (thm * bool) list |
| 24219 | 34 |
val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list) |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
35 |
val get_datatype_of_constr: theory -> string -> string option |
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
36 |
val get_case_scheme: theory -> string -> (int * (int * string list)) option |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
37 |
val is_undefined: theory -> string -> bool |
| 28423 | 38 |
val default_typscheme: theory -> string -> (string * sort) list * typ |
| 24219 | 39 |
|
| 28423 | 40 |
val preprocess_conv: theory -> cterm -> thm |
| 24837 | 41 |
val preprocess_term: theory -> term -> term |
| 28423 | 42 |
val postprocess_conv: theory -> cterm -> thm |
| 24837 | 43 |
val postprocess_term: theory -> term -> term |
| 24219 | 44 |
|
| 30513 | 45 |
val add_attribute: string * attribute parser -> theory -> theory |
| 24219 | 46 |
|
47 |
val print_codesetup: theory -> unit |
|
48 |
end; |
|
49 |
||
50 |
signature CODE_DATA_ARGS = |
|
51 |
sig |
|
52 |
type T |
|
53 |
val empty: T |
|
| 27609 | 54 |
val purge: theory -> string list -> T -> T |
| 24219 | 55 |
end; |
56 |
||
57 |
signature CODE_DATA = |
|
58 |
sig |
|
59 |
type T |
|
60 |
val get: theory -> T |
|
61 |
val change: theory -> (T -> T) -> T |
|
62 |
val change_yield: theory -> (T -> 'a * T) -> 'a * T |
|
63 |
end; |
|
64 |
||
65 |
signature PRIVATE_CODE = |
|
66 |
sig |
|
67 |
include CODE |
|
| 27609 | 68 |
val declare_data: Object.T -> (theory -> string list -> Object.T -> Object.T) |
69 |
-> serial |
|
| 24219 | 70 |
val get_data: serial * ('a -> Object.T) * (Object.T -> 'a)
|
71 |
-> theory -> 'a |
|
72 |
val change_data: serial * ('a -> Object.T) * (Object.T -> 'a)
|
|
73 |
-> theory -> ('a -> 'a) -> 'a
|
|
74 |
val change_yield_data: serial * ('a -> Object.T) * (Object.T -> 'a)
|
|
75 |
-> theory -> ('a -> 'b * 'a) -> 'b * 'a
|
|
76 |
end; |
|
77 |
||
78 |
structure Code : PRIVATE_CODE = |
|
79 |
struct |
|
80 |
||
| 25312 | 81 |
(** code attributes **) |
82 |
||
83 |
structure CodeAttr = TheoryDataFun ( |
|
| 30513 | 84 |
type T = (string * attribute parser) list; |
| 25312 | 85 |
val empty = []; |
86 |
val copy = I; |
|
87 |
val extend = I; |
|
| 26970 | 88 |
fun merge _ = AList.merge (op = : string * string -> bool) (K true); |
| 25312 | 89 |
); |
| 24219 | 90 |
|
| 25312 | 91 |
fun add_attribute (attr as (name, _)) = |
92 |
let |
|
| 28562 | 93 |
fun add_parser ("", parser) attrs = attrs |> rev |> AList.update (op =) ("", parser) |> rev
|
| 25312 | 94 |
| add_parser (name, parser) attrs = (name, Args.$$$ name |-- parser) :: attrs; |
| 28562 | 95 |
in CodeAttr.map (fn attrs => if not (name = "") andalso AList.defined (op =) attrs name |
96 |
then error ("Code attribute " ^ name ^ " already declared") else add_parser attr attrs)
|
|
| 25312 | 97 |
end; |
98 |
||
|
30722
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
99 |
val _ = Context.>> (Context.map_theory |
|
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
100 |
(Attrib.setup (Binding.name "code") |
|
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
101 |
(Scan.peek (fn context => |
|
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
102 |
List.foldr op || Scan.fail (map snd (CodeAttr.get (Context.theory_of context))))) |
|
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
103 |
"declare theorems for code generation")); |
|
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
104 |
|
| 25312 | 105 |
|
106 |
||
| 28143 | 107 |
(** logical and syntactical specification of executable code **) |
| 24219 | 108 |
|
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
109 |
(* code equations *) |
| 28695 | 110 |
|
| 28971 | 111 |
type eqns = bool * (thm * bool) list lazy; |
| 28695 | 112 |
(*default flag, theorems with linear flag (perhaps lazy)*) |
| 24219 | 113 |
|
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
114 |
fun pretty_lthms ctxt r = case Lazy.peek r |
| 28672 | 115 |
of SOME thms => map (ProofContext.pretty_thm ctxt o fst) (Exn.release thms) |
| 24219 | 116 |
| NONE => [Pretty.str "[...]"]; |
117 |
||
118 |
fun certificate thy f r = |
|
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
119 |
case Lazy.peek r |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
120 |
of SOME thms => (Lazy.value o f thy) (Exn.release thms) |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
121 |
| NONE => let |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
122 |
val thy_ref = Theory.check_thy thy; |
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
123 |
in Lazy.lazy (fn () => (f (Theory.deref thy_ref) o Lazy.force) r) end; |
| 24219 | 124 |
|
| 28423 | 125 |
fun add_drop_redundant thy (thm, linear) thms = |
| 24219 | 126 |
let |
127 |
val args_of = snd o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; |
|
128 |
val args = args_of thm; |
|
| 28403 | 129 |
val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1); |
| 28350 | 130 |
fun matches_args args' = length args <= length args' andalso |
| 28403 | 131 |
Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args'); |
| 28359 | 132 |
fun drop (thm', linear') = if (linear orelse not linear') |
133 |
andalso matches_args (args_of thm') then |
|
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
134 |
(warning ("Code generator: dropping redundant code equation\n" ^ Display.string_of_thm thm'); true)
|
| 28350 | 135 |
else false; |
136 |
in (thm, linear) :: filter_out drop thms end; |
|
| 24219 | 137 |
|
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
138 |
fun add_thm thy _ thm (false, thms) = (false, Lazy.map_force (add_drop_redundant thy thm) thms) |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
139 |
| add_thm thy true thm (true, thms) = (true, Lazy.map_force (fn thms => thms @ [thm]) thms) |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
140 |
| add_thm thy false thm (true, thms) = (false, Lazy.value [thm]); |
| 24219 | 141 |
|
| 28143 | 142 |
fun add_lthms lthms _ = (false, lthms); |
| 25312 | 143 |
|
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
28672
diff
changeset
|
144 |
fun del_thm thm = (apsnd o Lazy.map_force) (remove (eq_fst Thm.eq_thm_prop) (thm, true)); |
| 25312 | 145 |
|
| 28143 | 146 |
|
147 |
(* specification data *) |
|
148 |
||
| 24219 | 149 |
datatype spec = Spec of {
|
| 28695 | 150 |
concluded_history: bool, |
151 |
eqns: ((bool * eqns) * (serial * eqns) list) Symtab.table |
|
152 |
(*with explicit history*), |
|
153 |
dtyps: ((serial * ((string * sort) list * (string * typ list) list)) list) Symtab.table |
|
154 |
(*with explicit history*), |
|
| 30076 | 155 |
cases: (int * (int * string list)) Symtab.table * unit Symtab.table |
| 24219 | 156 |
}; |
157 |
||
| 28695 | 158 |
fun mk_spec ((concluded_history, eqns), (dtyps, cases)) = |
159 |
Spec { concluded_history = concluded_history, eqns = eqns, dtyps = dtyps, cases = cases };
|
|
160 |
fun map_spec f (Spec { concluded_history = concluded_history, eqns = eqns,
|
|
161 |
dtyps = dtyps, cases = cases }) = |
|
162 |
mk_spec (f ((concluded_history, eqns), (dtyps, cases))); |
|
163 |
fun merge_spec (Spec { concluded_history = _, eqns = eqns1, dtyps = dtyps1, cases = (cases1, undefs1) },
|
|
164 |
Spec { concluded_history = _, eqns = eqns2, dtyps = dtyps2, cases = (cases2, undefs2) }) =
|
|
| 24219 | 165 |
let |
| 28695 | 166 |
fun merge_eqns ((_, history1), (_, history2)) = |
167 |
let |
|
168 |
val raw_history = AList.merge (op =) (K true) (history1, history2) |
|
169 |
val filtered_history = filter_out (fst o snd) raw_history |
|
170 |
val history = if null filtered_history |
|
171 |
then raw_history else filtered_history; |
|
172 |
in ((false, (snd o hd) history), history) end; |
|
173 |
val eqns = Symtab.join (K merge_eqns) (eqns1, eqns2); |
|
174 |
val dtyps = Symtab.join (K (AList.merge (op =) (K true))) (dtyps1, dtyps2); |
|
| 27609 | 175 |
val cases = (Symtab.merge (K true) (cases1, cases2), |
176 |
Symtab.merge (K true) (undefs1, undefs2)); |
|
| 28695 | 177 |
in mk_spec ((false, eqns), (dtyps, cases)) end; |
| 24219 | 178 |
|
| 25312 | 179 |
|
180 |
(* pre- and postprocessor *) |
|
181 |
||
182 |
datatype thmproc = Thmproc of {
|
|
| 29302 | 183 |
pre: simpset, |
184 |
post: simpset, |
|
| 28423 | 185 |
functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list |
| 25312 | 186 |
}; |
187 |
||
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
188 |
fun mk_thmproc ((pre, post), functrans) = |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
189 |
Thmproc { pre = pre, post = post, functrans = functrans };
|
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
190 |
fun map_thmproc f (Thmproc { pre, post, functrans }) =
|
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
191 |
mk_thmproc (f ((pre, post), functrans)); |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
192 |
fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
|
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
193 |
Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
|
| 25312 | 194 |
let |
| 29302 | 195 |
val pre = Simplifier.merge_ss (pre1, pre2); |
196 |
val post = Simplifier.merge_ss (post1, post2); |
|
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
197 |
val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2); |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
198 |
in mk_thmproc ((pre, post), functrans) end; |
| 25312 | 199 |
|
| 24219 | 200 |
datatype exec = Exec of {
|
201 |
thmproc: thmproc, |
|
202 |
spec: spec |
|
203 |
}; |
|
204 |
||
| 28143 | 205 |
|
206 |
(* code setup data *) |
|
207 |
||
| 24219 | 208 |
fun mk_exec (thmproc, spec) = |
209 |
Exec { thmproc = thmproc, spec = spec };
|
|
210 |
fun map_exec f (Exec { thmproc = thmproc, spec = spec }) =
|
|
211 |
mk_exec (f (thmproc, spec)); |
|
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
212 |
fun merge_exec (Exec { thmproc = thmproc1, spec = spec1 },
|
| 24219 | 213 |
Exec { thmproc = thmproc2, spec = spec2 }) =
|
214 |
let |
|
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
215 |
val thmproc = merge_thmproc (thmproc1, thmproc2); |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
216 |
val spec = merge_spec (spec1, spec2); |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
217 |
in mk_exec (thmproc, spec) end; |
| 29302 | 218 |
val empty_exec = mk_exec (mk_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []), |
| 28695 | 219 |
mk_spec ((false, Symtab.empty), (Symtab.empty, (Symtab.empty, Symtab.empty)))); |
| 24219 | 220 |
|
| 25312 | 221 |
fun the_thmproc (Exec { thmproc = Thmproc x, ...}) = x;
|
| 24219 | 222 |
fun the_spec (Exec { spec = Spec x, ...}) = x;
|
| 28368 | 223 |
val the_eqns = #eqns o the_spec; |
| 24219 | 224 |
val the_dtyps = #dtyps o the_spec; |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
225 |
val the_cases = #cases o the_spec; |
| 24219 | 226 |
val map_thmproc = map_exec o apfst o map_thmproc; |
| 28695 | 227 |
val map_concluded_history = map_exec o apsnd o map_spec o apfst o apfst; |
228 |
val map_eqns = map_exec o apsnd o map_spec o apfst o apsnd; |
|
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
229 |
val map_dtyps = map_exec o apsnd o map_spec o apsnd o apfst; |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
230 |
val map_cases = map_exec o apsnd o map_spec o apsnd o apsnd; |
| 24219 | 231 |
|
232 |
||
233 |
(* data slots dependent on executable content *) |
|
234 |
||
235 |
(*private copy avoids potential conflict of table exceptions*) |
|
236 |
structure Datatab = TableFun(type key = int val ord = int_ord); |
|
237 |
||
238 |
local |
|
239 |
||
240 |
type kind = {
|
|
241 |
empty: Object.T, |
|
| 27609 | 242 |
purge: theory -> string list -> Object.T -> Object.T |
| 24219 | 243 |
}; |
244 |
||
245 |
val kinds = ref (Datatab.empty: kind Datatab.table); |
|
246 |
val kind_keys = ref ([]: serial list); |
|
247 |
||
248 |
fun invoke f k = case Datatab.lookup (! kinds) k |
|
249 |
of SOME kind => f kind |
|
250 |
| NONE => sys_error "Invalid code data identifier"; |
|
251 |
||
252 |
in |
|
253 |
||
| 27609 | 254 |
fun declare_data empty purge = |
| 24219 | 255 |
let |
256 |
val k = serial (); |
|
| 27609 | 257 |
val kind = {empty = empty, purge = purge};
|
| 24219 | 258 |
val _ = change kinds (Datatab.update (k, kind)); |
259 |
val _ = change kind_keys (cons k); |
|
260 |
in k end; |
|
261 |
||
| 27609 | 262 |
fun invoke_init k = invoke (fn kind => #empty kind) k; |
| 24219 | 263 |
|
| 27609 | 264 |
fun invoke_purge_all thy cs = |
| 24219 | 265 |
fold (fn k => Datatab.map_entry k |
| 27609 | 266 |
(invoke (fn kind => #purge kind thy cs) k)) (! kind_keys); |
| 24219 | 267 |
|
268 |
end; (*local*) |
|
269 |
||
270 |
||
| 25312 | 271 |
(** theory store **) |
| 24219 | 272 |
|
273 |
local |
|
274 |
||
275 |
type data = Object.T Datatab.table; |
|
| 27609 | 276 |
val empty_data = Datatab.empty : data; |
| 24219 | 277 |
|
278 |
structure CodeData = TheoryDataFun |
|
279 |
( |
|
280 |
type T = exec * data ref; |
|
| 27609 | 281 |
val empty = (empty_exec, ref empty_data); |
| 24219 | 282 |
fun copy (exec, data) = (exec, ref (! data)); |
283 |
val extend = copy; |
|
284 |
fun merge pp ((exec1, data1), (exec2, data2)) = |
|
| 27609 | 285 |
(merge_exec (exec1, exec2), ref empty_data); |
| 24219 | 286 |
); |
287 |
||
288 |
fun thy_data f thy = f ((snd o CodeData.get) thy); |
|
289 |
||
290 |
fun get_ensure_init kind data_ref = |
|
291 |
case Datatab.lookup (! data_ref) kind |
|
292 |
of SOME x => x |
|
| 27609 | 293 |
| NONE => let val y = invoke_init kind |
| 24219 | 294 |
in (change data_ref (Datatab.update (kind, y)); y) end; |
295 |
||
296 |
in |
|
297 |
||
298 |
(* access to executable content *) |
|
299 |
||
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
300 |
val the_exec = fst o CodeData.get; |
| 24219 | 301 |
|
| 27983 | 302 |
fun complete_class_params thy cs = |
303 |
fold (fn c => case AxClass.inst_of_param thy c |
|
304 |
of NONE => insert (op =) c |
|
305 |
| SOME (c', _) => insert (op =) c' #> insert (op =) c) cs []; |
|
306 |
||
| 24219 | 307 |
fun map_exec_purge touched f thy = |
| 27609 | 308 |
CodeData.map (fn (exec, data) => (f exec, ref (case touched |
| 27983 | 309 |
of SOME cs => invoke_purge_all thy (complete_class_params thy cs) (! data) |
| 27609 | 310 |
| NONE => empty_data))) thy; |
| 24219 | 311 |
|
| 27675 | 312 |
val purge_data = (CodeData.map o apsnd) (K (ref empty_data)); |
313 |
||
| 24219 | 314 |
|
| 28695 | 315 |
(* tackling equation history *) |
316 |
||
317 |
fun get_eqns thy c = |
|
318 |
Symtab.lookup ((the_eqns o the_exec) thy) c |
|
319 |
|> Option.map (Lazy.force o snd o snd o fst) |
|
320 |
|> these; |
|
321 |
||
322 |
fun continue_history thy = if (#concluded_history o the_spec o the_exec) thy |
|
323 |
then thy |
|
324 |
|> (CodeData.map o apfst o map_concluded_history) (K false) |
|
325 |
|> SOME |
|
326 |
else NONE; |
|
327 |
||
328 |
fun conclude_history thy = if (#concluded_history o the_spec o the_exec) thy |
|
329 |
then NONE |
|
330 |
else thy |
|
331 |
|> (CodeData.map o apfst) |
|
332 |
((map_eqns o Symtab.map) (fn ((changed, current), history) => |
|
333 |
((false, current), |
|
334 |
if changed then (serial (), current) :: history else history)) |
|
335 |
#> map_concluded_history (K true)) |
|
336 |
|> SOME; |
|
337 |
||
338 |
val _ = Context.>> (Context.map_theory (CodeData.init |
|
339 |
#> Theory.at_begin continue_history |
|
340 |
#> Theory.at_end conclude_history)); |
|
341 |
||
342 |
||
| 24219 | 343 |
(* access to data dependent on abstract executable content *) |
344 |
||
345 |
fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest); |
|
346 |
||
347 |
fun change_data (kind, mk, dest) = |
|
348 |
let |
|
349 |
fun chnge data_ref f = |
|
350 |
let |
|
351 |
val data = get_ensure_init kind data_ref; |
|
352 |
val data' = f (dest data); |
|
353 |
in (change data_ref (Datatab.update (kind, mk data')); data') end; |
|
354 |
in thy_data chnge end; |
|
355 |
||
356 |
fun change_yield_data (kind, mk, dest) = |
|
357 |
let |
|
358 |
fun chnge data_ref f = |
|
359 |
let |
|
360 |
val data = get_ensure_init kind data_ref; |
|
361 |
val (x, data') = f (dest data); |
|
362 |
in (x, (change data_ref (Datatab.update (kind, mk data')); data')) end; |
|
363 |
in thy_data chnge end; |
|
364 |
||
365 |
end; (*local*) |
|
366 |
||
367 |
||
368 |
(* print executable content *) |
|
369 |
||
370 |
fun print_codesetup thy = |
|
371 |
let |
|
372 |
val ctxt = ProofContext.init thy; |
|
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
373 |
val exec = the_exec thy; |
| 28368 | 374 |
fun pretty_eqn (s, (_, lthms)) = |
| 24219 | 375 |
(Pretty.block o Pretty.fbreaks) ( |
| 28143 | 376 |
Pretty.str s :: pretty_lthms ctxt lthms |
| 24219 | 377 |
); |
378 |
fun pretty_dtyp (s, []) = |
|
379 |
Pretty.str s |
|
380 |
| pretty_dtyp (s, cos) = |
|
381 |
(Pretty.block o Pretty.breaks) ( |
|
382 |
Pretty.str s |
|
383 |
:: Pretty.str "=" |
|
| 28403 | 384 |
:: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (Code_Unit.string_of_const thy c) |
| 24219 | 385 |
| (c, tys) => |
386 |
(Pretty.block o Pretty.breaks) |
|
| 28054 | 387 |
(Pretty.str (Code_Unit.string_of_const thy c) |
|
26947
133905a0c493
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26928
diff
changeset
|
388 |
:: Pretty.str "of" |
|
133905a0c493
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26928
diff
changeset
|
389 |
:: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) |
| 24219 | 390 |
); |
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
391 |
val pre = (#pre o the_thmproc) exec; |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
392 |
val post = (#post o the_thmproc) exec; |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
393 |
val functrans = (map fst o #functrans o the_thmproc) exec; |
| 28368 | 394 |
val eqns = the_eqns exec |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
395 |
|> Symtab.dest |
| 28054 | 396 |
|> (map o apfst) (Code_Unit.string_of_const thy) |
| 28695 | 397 |
|> (map o apsnd) (snd o fst) |
| 24219 | 398 |
|> sort (string_ord o pairself fst); |
399 |
val dtyps = the_dtyps exec |
|
400 |
|> Symtab.dest |
|
| 28695 | 401 |
|> map (fn (dtco, (_, (vs, cos)) :: _) => |
|
26947
133905a0c493
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26928
diff
changeset
|
402 |
(Syntax.string_of_typ_global thy (Type (dtco, map TFree vs)), cos)) |
| 24219 | 403 |
|> sort (string_ord o pairself fst) |
404 |
in |
|
405 |
(Pretty.writeln o Pretty.chunks) [ |
|
406 |
Pretty.block ( |
|
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
407 |
Pretty.str "code equations:" |
| 24219 | 408 |
:: Pretty.fbrk |
| 28368 | 409 |
:: (Pretty.fbreaks o map pretty_eqn) eqns |
| 24219 | 410 |
), |
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
411 |
Pretty.block [ |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
412 |
Pretty.str "preprocessing simpset:", |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
413 |
Pretty.fbrk, |
| 30357 | 414 |
Simplifier.pretty_ss ctxt pre |
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
415 |
], |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
416 |
Pretty.block [ |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
417 |
Pretty.str "postprocessing simpset:", |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
418 |
Pretty.fbrk, |
| 30357 | 419 |
Simplifier.pretty_ss ctxt post |
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
420 |
], |
| 24219 | 421 |
Pretty.block ( |
| 27609 | 422 |
Pretty.str "function transformers:" |
| 24219 | 423 |
:: Pretty.fbrk |
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
424 |
:: (Pretty.fbreaks o map Pretty.str) functrans |
| 25968 | 425 |
), |
426 |
Pretty.block ( |
|
| 24219 | 427 |
Pretty.str "datatypes:" |
428 |
:: Pretty.fbrk |
|
429 |
:: (Pretty.fbreaks o map pretty_dtyp) dtyps |
|
430 |
) |
|
431 |
] |
|
432 |
end; |
|
433 |
||
434 |
||
435 |
(** theorem transformation and certification **) |
|
436 |
||
| 28423 | 437 |
fun common_typ_eqns thy [] = [] |
438 |
| common_typ_eqns thy [thm] = [thm] |
|
439 |
| common_typ_eqns thy (thms as thm :: _) = (*FIXME is too general*) |
|
| 24219 | 440 |
let |
441 |
fun incr_thm thm max = |
|
442 |
let |
|
443 |
val thm' = incr_indexes max thm; |
|
444 |
val max' = Thm.maxidx_of thm' + 1; |
|
445 |
in (thm', max') end; |
|
446 |
val (thms', maxidx) = fold_map incr_thm thms 0; |
|
| 28423 | 447 |
val ty1 :: tys = map (snd o Code_Unit.const_typ_eqn) thms'; |
| 24219 | 448 |
fun unify ty env = Sign.typ_unify thy (ty1, ty) env |
449 |
handle Type.TUNIFY => |
|
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
450 |
error ("Type unificaton failed, while unifying code equations\n"
|
| 24219 | 451 |
^ (cat_lines o map Display.string_of_thm) thms |
452 |
^ "\nwith types\n" |
|
| 28054 | 453 |
^ (cat_lines o map (Code_Unit.string_of_typ thy)) (ty1 :: tys)); |
| 24219 | 454 |
val (env, _) = fold unify tys (Vartab.empty, maxidx) |
455 |
val instT = Vartab.fold (fn (x_i, (sort, ty)) => |
|
456 |
cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env []; |
|
457 |
in map (Thm.instantiate (instT, [])) thms' end; |
|
458 |
||
| 28423 | 459 |
fun check_linear (eqn as (thm, linear)) = |
460 |
if linear then eqn else Code_Unit.bad_thm |
|
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
461 |
("Duplicate variables on left hand side of code equation:\n"
|
| 28423 | 462 |
^ Display.string_of_thm thm); |
463 |
||
464 |
fun mk_eqn thy linear = |
|
465 |
Code_Unit.error_thm ((if linear then check_linear else I) o Code_Unit.mk_eqn thy); |
|
466 |
fun mk_syntactic_eqn thy = Code_Unit.warning_thm (Code_Unit.mk_eqn thy); |
|
467 |
fun mk_default_eqn thy = Code_Unit.try_thm (check_linear o Code_Unit.mk_eqn thy); |
|
| 24219 | 468 |
|
469 |
||
470 |
(** interfaces and attributes **) |
|
471 |
||
|
24624
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24585
diff
changeset
|
472 |
fun delete_force msg key xs = |
|
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24585
diff
changeset
|
473 |
if AList.defined (op =) xs key then AList.delete (op =) key xs |
|
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24585
diff
changeset
|
474 |
else error ("No such " ^ msg ^ ": " ^ quote key);
|
|
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24585
diff
changeset
|
475 |
|
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
476 |
fun get_datatype thy tyco = |
| 28695 | 477 |
case these (Symtab.lookup ((the_dtyps o the_exec) thy) tyco) |
478 |
of (_, spec) :: _ => spec |
|
479 |
| [] => Sign.arity_number thy tyco |
|
| 24848 | 480 |
|> Name.invents Name.context Name.aT |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
481 |
|> map (rpair []) |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
482 |
|> rpair []; |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
483 |
|
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
484 |
fun get_datatype_of_constr thy c = |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
485 |
case (snd o strip_type o Sign.the_const_type thy) c |
| 28695 | 486 |
of Type (tyco, _) => if member (op =) ((map fst o snd o get_datatype thy) tyco) c |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
487 |
then SOME tyco else NONE |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
488 |
| _ => NONE; |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
489 |
|
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
490 |
fun recheck_eqn thy = Code_Unit.error_thm |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
491 |
(Code_Unit.assert_linear (is_some o get_datatype_of_constr thy) o apfst (Code_Unit.assert_eqn thy)); |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
492 |
|
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
493 |
fun recheck_eqns_const thy c eqns = |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
494 |
let |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
495 |
fun cert (eqn as (thm, _)) = if c = Code_Unit.const_eqn thm |
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
496 |
then eqn else error ("Wrong head of code equation,\nexpected constant "
|
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
497 |
^ Code_Unit.string_of_const thy c ^ "\n" ^ Display.string_of_thm thm) |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
498 |
in map (cert o recheck_eqn thy) eqns end; |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
499 |
|
| 28695 | 500 |
fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns |
501 |
o (if delete then Symtab.map_entry c else Symtab.map_default (c, ((false, (true, Lazy.value [])), []))) |
|
502 |
o apfst) (fn (_, eqns) => (true, f eqns)); |
|
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
503 |
|
| 28695 | 504 |
fun gen_add_eqn linear default thm thy = |
505 |
case (if default then mk_default_eqn thy else SOME o mk_eqn thy linear) thm |
|
| 28368 | 506 |
of SOME (thm, _) => |
| 28143 | 507 |
let |
| 28423 | 508 |
val c = Code_Unit.const_eqn thm; |
| 28695 | 509 |
val _ = if not default andalso (is_some o AxClass.class_of_param thy) c |
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
510 |
then error ("Rejected polymorphic code equation for overloaded constant:\n"
|
| 28143 | 511 |
^ Display.string_of_thm thm) |
512 |
else (); |
|
| 28695 | 513 |
val _ = if not default andalso (is_some o get_datatype_of_constr thy) c |
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
514 |
then error ("Rejected code equation for datatype constructor:\n"
|
| 28368 | 515 |
^ Display.string_of_thm thm) |
| 28143 | 516 |
else (); |
| 28695 | 517 |
in change_eqns false c (add_thm thy default (thm, linear)) thy end |
|
24624
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24585
diff
changeset
|
518 |
| NONE => thy; |
|
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24585
diff
changeset
|
519 |
|
| 28695 | 520 |
val add_eqn = gen_add_eqn true false; |
521 |
val add_default_eqn = gen_add_eqn true true; |
|
522 |
val add_nonlinear_eqn = gen_add_eqn false false; |
|
| 26021 | 523 |
|
| 28368 | 524 |
fun add_eqnl (c, lthms) thy = |
| 24219 | 525 |
let |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
526 |
val lthms' = certificate thy (fn thy => recheck_eqns_const thy c) lthms; |
| 28695 | 527 |
in change_eqns false c (add_lthms lthms') thy end; |
| 24219 | 528 |
|
| 28703 | 529 |
val add_default_eqn_attribute = Thm.declaration_attribute |
530 |
(fn thm => Context.mapping (add_default_eqn thm) I); |
|
531 |
val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute); |
|
| 24219 | 532 |
|
| 28695 | 533 |
fun del_eqn thm thy = case mk_syntactic_eqn thy thm |
534 |
of SOME (thm, _) => change_eqns true (Code_Unit.const_eqn thm) (del_thm thm) thy |
|
535 |
| NONE => thy; |
|
536 |
||
537 |
fun del_eqns c = change_eqns true c (K (false, Lazy.value [])); |
|
538 |
||
| 30076 | 539 |
fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy); |
| 28695 | 540 |
|
541 |
val is_undefined = Symtab.defined o snd o the_cases o the_exec; |
|
542 |
||
| 25485 | 543 |
structure TypeInterpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool); |
| 25462 | 544 |
|
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
545 |
fun add_datatype raw_cs thy = |
| 24219 | 546 |
let |
|
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25501
diff
changeset
|
547 |
val cs = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) raw_cs; |
| 28054 | 548 |
val (tyco, vs_cos) = Code_Unit.constrset_of_consts thy cs; |
| 30076 | 549 |
val old_cs = (map fst o snd o get_datatype thy) tyco; |
550 |
fun drop_outdated_cases cases = fold Symtab.delete_safe |
|
551 |
(Symtab.fold (fn (c, (_, (_, cos))) => |
|
552 |
if exists (member (op =) old_cs) cos |
|
553 |
then insert (op =) c else I) cases []) cases; |
|
| 24219 | 554 |
in |
555 |
thy |
|
| 28703 | 556 |
|> map_exec_purge NONE |
| 28695 | 557 |
((map_dtyps o Symtab.map_default (tyco, [])) (cons (serial (), vs_cos)) |
| 30076 | 558 |
#> map_eqns (fold (Symtab.delete_safe o fst) cs) |
559 |
#> (map_cases o apfst) drop_outdated_cases) |
|
| 25485 | 560 |
|> TypeInterpretation.data (tyco, serial ()) |
| 24219 | 561 |
end; |
562 |
||
| 25485 | 563 |
fun type_interpretation f = TypeInterpretation.interpretation |
564 |
(fn (tyco, _) => fn thy => f (tyco, get_datatype thy tyco) thy); |
|
565 |
||
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
566 |
fun add_datatype_cmd raw_cs thy = |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
567 |
let |
| 28054 | 568 |
val cs = map (Code_Unit.read_bare_const thy) raw_cs; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
569 |
in add_datatype cs thy end; |
| 24219 | 570 |
|
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
571 |
fun add_case thm thy = |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
572 |
let |
| 30076 | 573 |
val (c, (k, case_pats)) = Code_Unit.case_cert thm; |
574 |
val _ = case filter (is_none o get_datatype_of_constr thy) case_pats |
|
575 |
of [] => () |
|
576 |
| cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs));
|
|
577 |
val entry = (1 + Int.max (1, length case_pats), (k, case_pats)) |
|
578 |
in (map_exec_purge (SOME [c]) o map_cases o apfst) (Symtab.update (c, entry)) thy end; |
|
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
579 |
|
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
580 |
fun add_undefined c thy = |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
581 |
(map_exec_purge (SOME [c]) o map_cases o apsnd) (Symtab.update (c, ())) thy; |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24837
diff
changeset
|
582 |
|
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
583 |
val map_pre = map_exec_purge NONE o map_thmproc o apfst o apfst; |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
584 |
val map_post = map_exec_purge NONE o map_thmproc o apfst o apsnd; |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
585 |
|
| 28525 | 586 |
val add_inline = map_pre o MetaSimplifier.add_simp; |
587 |
val del_inline = map_pre o MetaSimplifier.del_simp; |
|
588 |
val add_post = map_post o MetaSimplifier.add_simp; |
|
589 |
val del_post = map_post o MetaSimplifier.del_simp; |
|
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
590 |
|
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
591 |
fun add_functrans (name, f) = |
|
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
592 |
(map_exec_purge NONE o map_thmproc o apsnd) |
| 24219 | 593 |
(AList.update (op =) (name, (serial (), f))); |
594 |
||
|
27557
151731493264
simpsets as pre/postprocessors; generic preprocessor now named function transformators
haftmann
parents:
26970
diff
changeset
|
595 |
fun del_functrans name = |
| 24219 | 596 |
(map_exec_purge NONE o map_thmproc o apsnd) |
| 27609 | 597 |
(delete_force "function transformer" name); |
| 24219 | 598 |
|
| 26463 | 599 |
val _ = Context.>> (Context.map_theory |
| 24219 | 600 |
(let |
601 |
fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); |
|
602 |
fun add_simple_attribute (name, f) = |
|
603 |
add_attribute (name, Scan.succeed (mk_attribute f)); |
|
604 |
fun add_del_attribute (name, (add, del)) = |
|
605 |
add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del) |
|
606 |
|| Scan.succeed (mk_attribute add)) |
|
607 |
in |
|
| 25462 | 608 |
TypeInterpretation.init |
| 28562 | 609 |
#> add_del_attribute ("", (add_eqn, del_eqn))
|
| 28368 | 610 |
#> add_simple_attribute ("nbe", add_nonlinear_eqn)
|
| 24219 | 611 |
#> add_del_attribute ("inline", (add_inline, del_inline))
|
612 |
#> add_del_attribute ("post", (add_post, del_post))
|
|
| 26463 | 613 |
end)); |
| 24219 | 614 |
|
615 |
||
616 |
(** post- and preprocessing **) |
|
617 |
||
618 |
local |
|
619 |
||
| 28423 | 620 |
fun apply_functrans thy c _ [] = [] |
621 |
| apply_functrans thy c [] eqns = eqns |
|
622 |
| apply_functrans thy c functrans eqns = eqns |
|
623 |
|> perhaps (perhaps_loop (perhaps_apply functrans)) |
|
| 28525 | 624 |
|> (map o apfst) (AxClass.unoverload thy) |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
625 |
|> recheck_eqns_const thy c |
| 28525 | 626 |
|> (map o apfst) (AxClass.overload thy); |
| 24219 | 627 |
|
| 28423 | 628 |
fun rhs_conv conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm); |
| 24219 | 629 |
|
| 24837 | 630 |
fun term_of_conv thy f = |
631 |
Thm.cterm_of thy |
|
632 |
#> f |
|
633 |
#> Thm.prop_of |
|
634 |
#> Logic.dest_equals |
|
635 |
#> snd; |
|
636 |
||
| 28423 | 637 |
fun preprocess thy functrans c eqns = |
| 27582 | 638 |
let |
639 |
val pre = (Simplifier.theory_context thy o #pre o the_thmproc o the_exec) thy; |
|
640 |
in |
|
| 28423 | 641 |
eqns |
| 28525 | 642 |
|> (map o apfst) (AxClass.overload thy) |
| 28423 | 643 |
|> apply_functrans thy c functrans |
| 28368 | 644 |
|> (map o apfst) (Code_Unit.rewrite_eqn pre) |
| 28525 | 645 |
|> (map o apfst) (AxClass.unoverload thy) |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28703
diff
changeset
|
646 |
|> map (recheck_eqn thy) |
| 28423 | 647 |
|> burrow_fst (common_typ_eqns thy) |
| 27582 | 648 |
end; |
| 26970 | 649 |
|
| 28423 | 650 |
in |
651 |
||
652 |
fun preprocess_conv thy ct = |
|
| 24219 | 653 |
let |
| 27582 | 654 |
val pre = (Simplifier.theory_context thy o #pre o the_thmproc o the_exec) thy; |
| 24219 | 655 |
in |
656 |
ct |
|
| 27582 | 657 |
|> Simplifier.rewrite pre |
|
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25501
diff
changeset
|
658 |
|> rhs_conv (AxClass.unoverload_conv thy) |
| 24219 | 659 |
end; |
660 |
||
| 28423 | 661 |
fun preprocess_term thy = term_of_conv thy (preprocess_conv thy); |
| 24837 | 662 |
|
| 28423 | 663 |
fun postprocess_conv thy ct = |
| 24219 | 664 |
let |
| 27582 | 665 |
val post = (Simplifier.theory_context thy o #post o the_thmproc o the_exec) thy; |
| 24219 | 666 |
in |
667 |
ct |
|
|
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25501
diff
changeset
|
668 |
|> AxClass.overload_conv thy |
| 27582 | 669 |
|> rhs_conv (Simplifier.rewrite post) |
| 24219 | 670 |
end; |
671 |
||
| 28423 | 672 |
fun postprocess_term thy = term_of_conv thy (postprocess_conv thy); |
| 24219 | 673 |
|
| 28525 | 674 |
fun these_raw_eqns thy c = |
675 |
get_eqns thy c |
|
| 28695 | 676 |
|> (map o apfst) (Thm.transfer thy) |
| 28525 | 677 |
|> burrow_fst (common_typ_eqns thy); |
678 |
||
| 28423 | 679 |
fun these_eqns thy c = |
| 24219 | 680 |
let |
| 28423 | 681 |
val functrans = (map (fn (_, (_, f)) => f thy) o #functrans |
682 |
o the_thmproc o the_exec) thy; |
|
| 24219 | 683 |
in |
| 28423 | 684 |
get_eqns thy c |
| 28695 | 685 |
|> (map o apfst) (Thm.transfer thy) |
| 28423 | 686 |
|> preprocess thy functrans c |
| 24219 | 687 |
end; |
688 |
||
| 28525 | 689 |
fun default_typscheme thy c = |
690 |
let |
|
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
691 |
fun the_const_typscheme c = (curry (Code_Unit.typscheme thy) c o snd o dest_Const |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
692 |
o TermSubst.zero_var_indexes o curry Const "" o Sign.the_const_type thy) c; |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
693 |
fun strip_sorts (vs, ty) = (map (fn (v, _) => (v, [])) vs, ty); |
| 28423 | 694 |
in case AxClass.class_of_param thy c |
|
30023
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
695 |
of SOME class => ([(Name.aT, [class])], snd (the_const_typscheme c)) |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
696 |
| NONE => if is_some (get_datatype_of_constr thy c) |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
697 |
then strip_sorts (the_const_typscheme c) |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
698 |
else case get_eqns thy c |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
699 |
of (thm, _) :: _ => snd (Code_Unit.head_eqn thy (Drule.zero_var_indexes thm)) |
|
55954f726043
permissive check for pattern discipline in case schemes
haftmann
parents:
29970
diff
changeset
|
700 |
| [] => strip_sorts (the_const_typscheme c) end; |
| 24219 | 701 |
|
702 |
end; (*local*) |
|
703 |
||
704 |
end; (*struct*) |
|
705 |
||
706 |
||
707 |
(** type-safe interfaces for data depedent on executable content **) |
|
708 |
||
709 |
functor CodeDataFun(Data: CODE_DATA_ARGS): CODE_DATA = |
|
710 |
struct |
|
711 |
||
712 |
type T = Data.T; |
|
713 |
exception Data of T; |
|
714 |
fun dest (Data x) = x |
|
715 |
||
716 |
val kind = Code.declare_data (Data Data.empty) |
|
| 27609 | 717 |
(fn thy => fn cs => fn Data x => Data (Data.purge thy cs x)); |
| 24219 | 718 |
|
719 |
val data_op = (kind, Data, dest); |
|
720 |
||
721 |
val get = Code.get_data data_op; |
|
722 |
val change = Code.change_data data_op; |
|
723 |
fun change_yield thy = Code.change_yield_data data_op thy; |
|
724 |
||
725 |
end; |
|
726 |
||
| 28143 | 727 |
structure Code : CODE = struct open Code; end; |