author | haftmann |
Tue, 20 Mar 2007 15:52:41 +0100 | |
changeset 22484 | 25dfebd7b4c8 |
parent 22423 | c1836b14c63a |
child 22507 | 3572bc633d9a |
permissions | -rw-r--r-- |
20600 | 1 |
(* Title: Pure/Tools/codegen_data.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
22020 | 5 |
Abstract executable content of theory. Management of data dependent on |
6 |
executable content. |
|
20600 | 7 |
*) |
8 |
||
9 |
signature CODEGEN_DATA = |
|
10 |
sig |
|
21338 | 11 |
val lazy: (unit -> thm list) -> thm list Susp.T |
20600 | 12 |
val eval_always: bool ref |
13 |
||
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
14 |
val add_func: bool -> thm -> theory -> theory |
20600 | 15 |
val del_func: thm -> theory -> theory |
21338 | 16 |
val add_funcl: CodegenConsts.const * thm list Susp.T -> theory -> theory |
22423 | 17 |
val add_datatype: string * ((string * sort) list * (string * typ list) list) |
20600 | 18 |
-> theory -> theory |
19 |
val add_inline: thm -> theory -> theory |
|
20 |
val del_inline: thm -> theory -> theory |
|
22050 | 21 |
val add_inline_proc: string * (theory -> cterm list -> thm list) -> theory -> theory |
22 |
val del_inline_proc: string -> theory -> theory |
|
23 |
val add_preproc: string * (theory -> thm list -> thm list) -> theory -> theory |
|
24 |
val del_preproc: string -> theory -> theory |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
25 |
val coregular_algebra: theory -> Sorts.algebra |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
26 |
val operational_algebra: theory -> (sort -> sort) * Sorts.algebra |
20600 | 27 |
val these_funcs: theory -> CodegenConsts.const -> thm list |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
28 |
val tap_typ: theory -> CodegenConsts.const -> typ option |
22423 | 29 |
val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list) |
20600 | 30 |
val get_datatype_of_constr: theory -> CodegenConsts.const -> string option |
31 |
||
22033 | 32 |
val preprocess_cterm: cterm -> thm |
20600 | 33 |
|
20844 | 34 |
val trace: bool ref |
20600 | 35 |
end; |
36 |
||
37 |
signature PRIVATE_CODEGEN_DATA = |
|
38 |
sig |
|
39 |
include CODEGEN_DATA |
|
40 |
type data |
|
41 |
structure CodeData: THEORY_DATA |
|
42 |
val declare: string -> Object.T -> (Pretty.pp -> Object.T * Object.T -> Object.T) |
|
20937 | 43 |
-> (theory option -> CodegenConsts.const list option -> Object.T -> Object.T) -> serial |
20600 | 44 |
val init: serial -> theory -> theory |
45 |
val get: serial -> (Object.T -> 'a) -> data -> 'a |
|
46 |
val put: serial -> ('a -> Object.T) -> 'a -> data -> data |
|
47 |
end; |
|
48 |
||
49 |
structure CodegenData : PRIVATE_CODEGEN_DATA = |
|
50 |
struct |
|
51 |
||
52 |
(** diagnostics **) |
|
53 |
||
20844 | 54 |
val trace = ref false; |
55 |
fun tracing f x = (if !trace then Output.tracing (f x) else (); x); |
|
20600 | 56 |
|
57 |
||
58 |
||
59 |
(** lazy theorems, certificate theorems **) |
|
60 |
||
61 |
val eval_always = ref false; |
|
62 |
||
63 |
fun lazy f = if !eval_always |
|
64 |
then Susp.value (f ()) |
|
65 |
else Susp.delay f; |
|
66 |
||
67 |
fun string_of_lthms r = case Susp.peek r |
|
68 |
of SOME thms => (map string_of_thm o rev) thms |
|
69 |
| NONE => ["[...]"]; |
|
70 |
||
71 |
fun pretty_lthms ctxt r = case Susp.peek r |
|
21066
ce6759d1d0b4
code nofunc now permits theorems violating typing discipline
haftmann
parents:
20937
diff
changeset
|
72 |
of SOME thms => map (ProofContext.pretty_thm ctxt) thms |
20600 | 73 |
| NONE => [Pretty.str "[...]"]; |
74 |
||
20844 | 75 |
fun certificate thy f r = |
20600 | 76 |
case Susp.peek r |
20844 | 77 |
of SOME thms => (Susp.value o f thy) thms |
78 |
| NONE => let |
|
79 |
val thy_ref = Theory.self_ref thy; |
|
80 |
in lazy (fn () => (f (Theory.deref thy_ref) o Susp.force) r) end; |
|
20600 | 81 |
|
82 |
fun merge' _ ([], []) = (false, []) |
|
83 |
| merge' _ ([], ys) = (true, ys) |
|
84 |
| merge' eq (xs, ys) = fold_rev |
|
85 |
(fn y => fn (t, xs) => (t orelse not (member eq xs y), insert eq y xs)) ys (false, xs); |
|
86 |
||
87 |
fun merge_alist eq_key eq (xys as (xs, ys)) = |
|
88 |
if eq_list (eq_pair eq_key eq) (xs, ys) |
|
89 |
then (false, xs) |
|
90 |
else (true, AList.merge eq_key eq xys); |
|
91 |
||
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22319
diff
changeset
|
92 |
val merge_thms = merge' Thm.eq_thm; |
20600 | 93 |
|
94 |
fun merge_lthms (r1, r2) = |
|
21119 | 95 |
if Susp.same (r1, r2) |
96 |
then (false, r1) |
|
20600 | 97 |
else case Susp.peek r1 |
98 |
of SOME [] => (true, r2) |
|
99 |
| _ => case Susp.peek r2 |
|
100 |
of SOME [] => (true, r1) |
|
101 |
| _ => (apsnd (lazy o K)) (merge_thms (Susp.force r1, Susp.force r2)); |
|
102 |
||
103 |
||
104 |
||
105 |
(** code theorems **) |
|
106 |
||
22197 | 107 |
(* pairs of (selected, deleted) defining equations *) |
20600 | 108 |
|
21338 | 109 |
type sdthms = thm list Susp.T * thm list; |
20600 | 110 |
|
22006
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
111 |
fun add_drop_redundant thm (sels, dels) = |
20600 | 112 |
let |
22006
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
113 |
val thy = Thm.theory_of_thm thm; |
21284 | 114 |
val args_of = snd o strip_comb o fst o Logic.dest_equals o Drule.plain_prop_of; |
115 |
val args = args_of thm; |
|
116 |
fun matches [] _ = true |
|
117 |
| matches (Var _ :: xs) [] = matches xs [] |
|
118 |
| matches (_ :: _) [] = false |
|
119 |
| matches (x :: xs) (y :: ys) = Pattern.matches thy (x, y) andalso matches xs ys; |
|
22006
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
120 |
fun drop thm' = not (matches args (args_of thm')) |
22197 | 121 |
orelse (warning ("Dropping redundant defining equation\n" ^ string_of_thm thm'); false); |
22006
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
122 |
val (keeps, drops) = List.partition drop sels; |
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22319
diff
changeset
|
123 |
in (thm :: keeps, dels |> fold (insert Thm.eq_thm) drops |> remove Thm.eq_thm thm) end; |
20600 | 124 |
|
125 |
fun add_thm thm (sels, dels) = |
|
22006
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
126 |
apfst Susp.value (add_drop_redundant thm (Susp.force sels, dels)); |
20600 | 127 |
|
128 |
fun add_lthms lthms (sels, []) = |
|
129 |
(lazy (fn () => fold add_drop_redundant |
|
22006
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
130 |
(Susp.force lthms) (Susp.force sels, []) |> fst), []) |
9dc365b03573
dropped function theorems are considered as deleted
haftmann
parents:
21988
diff
changeset
|
131 |
(*FIXME*) |
20600 | 132 |
| add_lthms lthms (sels, dels) = |
133 |
fold add_thm (Susp.force lthms) (sels, dels); |
|
134 |
||
135 |
fun del_thm thm (sels, dels) = |
|
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22319
diff
changeset
|
136 |
(Susp.value (remove Thm.eq_thm thm (Susp.force sels)), thm :: dels); |
20600 | 137 |
|
138 |
fun pretty_sdthms ctxt (sels, _) = pretty_lthms ctxt sels; |
|
139 |
||
21119 | 140 |
fun merge_sdthms ((sels1, dels1), (sels2, dels2)) = |
20600 | 141 |
let |
142 |
val (dels_t, dels) = merge_thms (dels1, dels2); |
|
143 |
in if dels_t |
|
144 |
then let |
|
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22319
diff
changeset
|
145 |
val (_, sels) = merge_thms (Susp.force sels1, subtract Thm.eq_thm dels1 (Susp.force sels2)) |
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22319
diff
changeset
|
146 |
val (_, dels) = merge_thms (dels1, subtract Thm.eq_thm (Susp.force sels1) dels2) |
20600 | 147 |
in (true, ((lazy o K) sels, dels)) end |
148 |
else let |
|
149 |
val (sels_t, sels) = merge_lthms (sels1, sels2) |
|
150 |
in (sels_t, (sels, dels)) end |
|
151 |
end; |
|
152 |
||
153 |
||
154 |
(** data structures **) |
|
155 |
||
156 |
structure Consttab = CodegenConsts.Consttab; |
|
157 |
||
158 |
datatype preproc = Preproc of { |
|
159 |
inlines: thm list, |
|
22050 | 160 |
inline_procs: (string * (serial * (theory -> cterm list -> thm list))) list, |
161 |
preprocs: (string * (serial * (theory -> thm list -> thm list))) list |
|
20600 | 162 |
}; |
163 |
||
21119 | 164 |
fun mk_preproc ((inlines, inline_procs), preprocs) = |
165 |
Preproc { inlines = inlines, inline_procs = inline_procs, preprocs = preprocs }; |
|
166 |
fun map_preproc f (Preproc { inlines, inline_procs, preprocs }) = |
|
167 |
mk_preproc (f ((inlines, inline_procs), preprocs)); |
|
168 |
fun merge_preproc (Preproc { inlines = inlines1, inline_procs = inline_procs1, preprocs = preprocs1 }, |
|
169 |
Preproc { inlines = inlines2, inline_procs = inline_procs2, preprocs = preprocs2 }) = |
|
20600 | 170 |
let |
171 |
val (touched1, inlines) = merge_thms (inlines1, inlines2); |
|
22050 | 172 |
val (touched2, inline_procs) = merge_alist (op =) (eq_fst (op =)) (inline_procs1, inline_procs2); |
173 |
val (touched3, preprocs) = merge_alist (op =) (eq_fst (op =)) (preprocs1, preprocs2); |
|
21119 | 174 |
in (touched1 orelse touched2 orelse touched3, |
175 |
mk_preproc ((inlines, inline_procs), preprocs)) end; |
|
20600 | 176 |
|
177 |
fun join_func_thms (tabs as (tab1, tab2)) = |
|
178 |
let |
|
179 |
val cs1 = Consttab.keys tab1; |
|
180 |
val cs2 = Consttab.keys tab2; |
|
181 |
val cs' = filter (member CodegenConsts.eq_const cs2) cs1; |
|
182 |
val cs'' = subtract (op =) cs' cs1 @ subtract (op =) cs' cs2; |
|
183 |
val cs''' = ref [] : CodegenConsts.const list ref; |
|
21119 | 184 |
fun merge c x = let val (touched, thms') = merge_sdthms x in |
20600 | 185 |
(if touched then cs''' := cons c (!cs''') else (); thms') end; |
186 |
in (cs'' @ !cs''', Consttab.join merge tabs) end; |
|
187 |
fun merge_funcs (thms1, thms2) = |
|
188 |
let |
|
189 |
val (consts, thms) = join_func_thms (thms1, thms2); |
|
190 |
in (SOME consts, thms) end; |
|
191 |
||
192 |
val eq_string = op = : string * string -> bool; |
|
22423 | 193 |
fun eq_dtyp ((vs1, cs1), (vs2, cs2)) = |
20600 | 194 |
gen_eq_set (eq_pair eq_string (gen_eq_set eq_string)) (vs1, vs2) |
195 |
andalso gen_eq_set (eq_pair eq_string (eq_list (is_equal o Term.typ_ord))) (cs1, cs2); |
|
196 |
fun merge_dtyps (tabs as (tab1, tab2)) = |
|
197 |
let |
|
198 |
val tycos1 = Symtab.keys tab1; |
|
199 |
val tycos2 = Symtab.keys tab2; |
|
200 |
val tycos' = filter (member eq_string tycos2) tycos1; |
|
20844 | 201 |
val touched = not (gen_eq_set (op =) (tycos1, tycos2) andalso |
202 |
gen_eq_set (eq_pair (op =) (eq_dtyp)) |
|
20600 | 203 |
(AList.make (the o Symtab.lookup tab1) tycos', |
20844 | 204 |
AList.make (the o Symtab.lookup tab2) tycos')); |
20600 | 205 |
in (touched, Symtab.merge (K true) tabs) end; |
206 |
||
207 |
datatype spec = Spec of { |
|
208 |
funcs: sdthms Consttab.table, |
|
209 |
dconstrs: string Consttab.table, |
|
22423 | 210 |
dtyps: ((string * sort) list * (string * typ list) list) Symtab.table |
20600 | 211 |
}; |
212 |
||
213 |
fun mk_spec ((funcs, dconstrs), dtyps) = |
|
214 |
Spec { funcs = funcs, dconstrs = dconstrs, dtyps = dtyps }; |
|
215 |
fun map_spec f (Spec { funcs = funcs, dconstrs = dconstrs, dtyps = dtyps }) = |
|
216 |
mk_spec (f ((funcs, dconstrs), dtyps)); |
|
217 |
fun merge_spec (Spec { funcs = funcs1, dconstrs = dconstrs1, dtyps = dtyps1 }, |
|
218 |
Spec { funcs = funcs2, dconstrs = dconstrs2, dtyps = dtyps2 }) = |
|
219 |
let |
|
220 |
val (touched_cs, funcs) = merge_funcs (funcs1, funcs2); |
|
221 |
val dconstrs = Consttab.merge (K true) (dconstrs1, dconstrs2); |
|
222 |
val (touched', dtyps) = merge_dtyps (dtyps1, dtyps2); |
|
223 |
val touched = if touched' then NONE else touched_cs; |
|
224 |
in (touched, mk_spec ((funcs, dconstrs), dtyps)) end; |
|
225 |
||
226 |
datatype exec = Exec of { |
|
227 |
preproc: preproc, |
|
228 |
spec: spec |
|
229 |
}; |
|
230 |
||
231 |
fun mk_exec (preproc, spec) = |
|
232 |
Exec { preproc = preproc, spec = spec }; |
|
233 |
fun map_exec f (Exec { preproc = preproc, spec = spec }) = |
|
234 |
mk_exec (f (preproc, spec)); |
|
235 |
fun merge_exec (Exec { preproc = preproc1, spec = spec1 }, |
|
236 |
Exec { preproc = preproc2, spec = spec2 }) = |
|
237 |
let |
|
238 |
val (touched', preproc) = merge_preproc (preproc1, preproc2); |
|
239 |
val (touched_cs, spec) = merge_spec (spec1, spec2); |
|
240 |
val touched = if touched' then NONE else touched_cs; |
|
241 |
in (touched, mk_exec (preproc, spec)) end; |
|
21119 | 242 |
val empty_exec = mk_exec (mk_preproc (([], []), []), |
20600 | 243 |
mk_spec ((Consttab.empty, Consttab.empty), Symtab.empty)); |
244 |
||
245 |
fun the_preproc (Exec { preproc = Preproc x, ...}) = x; |
|
246 |
fun the_spec (Exec { spec = Spec x, ...}) = x; |
|
247 |
val the_funcs = #funcs o the_spec; |
|
248 |
val the_dcontrs = #dconstrs o the_spec; |
|
249 |
val the_dtyps = #dtyps o the_spec; |
|
250 |
val map_preproc = map_exec o apfst o map_preproc; |
|
251 |
val map_funcs = map_exec o apsnd o map_spec o apfst o apfst; |
|
252 |
val map_dconstrs = map_exec o apsnd o map_spec o apfst o apsnd; |
|
253 |
val map_dtyps = map_exec o apsnd o map_spec o apsnd; |
|
254 |
||
255 |
||
256 |
(** code data structures **) |
|
257 |
||
258 |
(*private copy avoids potential conflict of table exceptions*) |
|
259 |
structure Datatab = TableFun(type key = int val ord = int_ord); |
|
260 |
||
261 |
||
262 |
(* data slots *) |
|
263 |
||
264 |
local |
|
265 |
||
266 |
type kind = { |
|
267 |
name: string, |
|
268 |
empty: Object.T, |
|
269 |
merge: Pretty.pp -> Object.T * Object.T -> Object.T, |
|
20937 | 270 |
purge: theory option -> CodegenConsts.const list option -> Object.T -> Object.T |
20600 | 271 |
}; |
272 |
||
273 |
val kinds = ref (Datatab.empty: kind Datatab.table); |
|
274 |
||
275 |
fun invoke meth_name meth_fn k = |
|
276 |
(case Datatab.lookup (! kinds) k of |
|
277 |
SOME kind => meth_fn kind |> transform_failure (fn exn => |
|
278 |
EXCEPTION (exn, "Code data method " ^ #name kind ^ "." ^ meth_name ^ " failed")) |
|
279 |
| NONE => sys_error ("Invalid code data identifier " ^ string_of_int k)); |
|
280 |
||
281 |
||
282 |
in |
|
283 |
||
284 |
fun invoke_name k = invoke "name" (K o #name) k (); |
|
285 |
fun invoke_empty k = invoke "empty" (K o #empty) k (); |
|
286 |
fun invoke_merge pp = invoke "merge" (fn kind => #merge kind pp); |
|
20937 | 287 |
fun invoke_purge thy_opt cs = invoke "purge" (fn kind => #purge kind thy_opt cs); |
20600 | 288 |
|
289 |
fun declare name empty merge purge = |
|
290 |
let |
|
291 |
val k = serial (); |
|
292 |
val kind = {name = name, empty = empty, merge = merge, purge = purge}; |
|
21962 | 293 |
val _ = |
294 |
if Datatab.exists (equal name o #name o #2) (! kinds) then |
|
295 |
warning ("Duplicate declaration of code data " ^ quote name) |
|
296 |
else (); |
|
20600 | 297 |
val _ = change kinds (Datatab.update (k, kind)); |
298 |
in k end; |
|
299 |
||
300 |
end; (*local*) |
|
301 |
||
302 |
||
303 |
(* theory store *) |
|
304 |
||
305 |
type data = Object.T Datatab.table; |
|
306 |
||
307 |
structure CodeData = TheoryDataFun |
|
308 |
(struct |
|
309 |
val name = "Pure/codegen_data"; |
|
310 |
type T = exec * data ref; |
|
311 |
val empty = (empty_exec, ref Datatab.empty : data ref); |
|
312 |
fun copy (exec, data) = (exec, ref (! data)); |
|
313 |
val extend = copy; |
|
314 |
fun merge pp ((exec1, data1), (exec2, data2)) = |
|
315 |
let |
|
316 |
val (touched, exec) = merge_exec (exec1, exec2); |
|
20937 | 317 |
val data1' = Datatab.map' (invoke_purge NONE touched) (! data1); |
318 |
val data2' = Datatab.map' (invoke_purge NONE touched) (! data2); |
|
20600 | 319 |
val data = Datatab.join (invoke_merge pp) (data1', data2'); |
320 |
in (exec, ref data) end; |
|
321 |
fun print thy (exec, _) = |
|
322 |
let |
|
323 |
val ctxt = ProofContext.init thy; |
|
324 |
fun pretty_func (s, lthms) = |
|
325 |
(Pretty.block o Pretty.fbreaks) ( |
|
326 |
Pretty.str s :: pretty_sdthms ctxt lthms |
|
327 |
); |
|
22423 | 328 |
fun pretty_dtyp (s, []) = |
329 |
Pretty.str s |
|
330 |
| pretty_dtyp (s, cos) = |
|
331 |
(Pretty.block o Pretty.breaks) ( |
|
332 |
Pretty.str s |
|
333 |
:: Pretty.str "=" |
|
334 |
:: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str c |
|
335 |
| (c, tys) => |
|
336 |
(Pretty.block o Pretty.breaks) |
|
337 |
(Pretty.str c :: Pretty.str "of" :: map (Pretty.quote o Sign.pretty_typ thy) tys)) cos) |
|
338 |
); |
|
20600 | 339 |
val inlines = (#inlines o the_preproc) exec; |
22050 | 340 |
val inline_procs = (map fst o #inline_procs o the_preproc) exec; |
341 |
val preprocs = (map fst o #preprocs o the_preproc) exec; |
|
20600 | 342 |
val funs = the_funcs exec |
343 |
|> Consttab.dest |
|
344 |
|> (map o apfst) (CodegenConsts.string_of_const thy) |
|
345 |
|> sort (string_ord o pairself fst); |
|
346 |
val dtyps = the_dtyps exec |
|
347 |
|> Symtab.dest |
|
22423 | 348 |
|> map (fn (dtco, (vs, cos)) => (Sign.string_of_typ thy (Type (dtco, map TFree vs)), cos)) |
20600 | 349 |
|> sort (string_ord o pairself fst) |
350 |
in |
|
22303 | 351 |
(Pretty.writeln o Pretty.chunks) [ |
352 |
Pretty.block ( |
|
353 |
Pretty.str "defining equations:" |
|
22423 | 354 |
:: Pretty.fbrk |
355 |
:: (Pretty.fbreaks o map pretty_func) funs |
|
22303 | 356 |
), |
20600 | 357 |
Pretty.block ( |
22050 | 358 |
Pretty.str "inlining theorems:" |
20600 | 359 |
:: Pretty.fbrk |
360 |
:: (Pretty.fbreaks o map (ProofContext.pretty_thm ctxt)) inlines |
|
361 |
), |
|
362 |
Pretty.block ( |
|
22050 | 363 |
Pretty.str "inlining procedures:" |
364 |
:: Pretty.fbrk |
|
365 |
:: (Pretty.fbreaks o map Pretty.str) inline_procs |
|
366 |
), |
|
367 |
Pretty.block ( |
|
368 |
Pretty.str "preprocessors:" |
|
369 |
:: Pretty.fbrk |
|
370 |
:: (Pretty.fbreaks o map Pretty.str) preprocs |
|
371 |
), |
|
372 |
Pretty.block ( |
|
20600 | 373 |
Pretty.str "datatypes:" |
374 |
:: Pretty.fbrk |
|
375 |
:: (Pretty.fbreaks o map pretty_dtyp) dtyps |
|
22303 | 376 |
) |
377 |
] |
|
20600 | 378 |
end; |
379 |
end); |
|
380 |
||
381 |
fun init k = CodeData.map |
|
382 |
(fn (exec, data) => (exec, ref (Datatab.update (k, invoke_empty k) (! data)))); |
|
383 |
||
384 |
fun get k dest data = |
|
385 |
(case Datatab.lookup data k of |
|
386 |
SOME x => (dest x handle Match => |
|
387 |
error ("Failed to access code data " ^ quote (invoke_name k))) |
|
388 |
| NONE => error ("Uninitialized code data " ^ quote (invoke_name k))); |
|
389 |
||
390 |
fun put k mk x = Datatab.update (k, mk x); |
|
391 |
||
20937 | 392 |
fun map_exec_purge touched f thy = |
20600 | 393 |
CodeData.map (fn (exec, data) => |
20937 | 394 |
(f exec, ref (Datatab.map' (invoke_purge (SOME thy) touched) (! data)))) thy; |
20600 | 395 |
|
396 |
val get_exec = fst o CodeData.get; |
|
397 |
||
398 |
val _ = Context.add_setup CodeData.init; |
|
399 |
||
400 |
||
401 |
||
402 |
(** theorem transformation and certification **) |
|
403 |
||
22033 | 404 |
fun common_typ_funcs [] = [] |
405 |
| common_typ_funcs [thm] = [thm] |
|
406 |
| common_typ_funcs thms = |
|
20600 | 407 |
let |
22033 | 408 |
val thy = Thm.theory_of_thm (hd thms) |
20600 | 409 |
fun incr_thm thm max = |
410 |
let |
|
411 |
val thm' = incr_indexes max thm; |
|
21119 | 412 |
val max' = Thm.maxidx_of thm' + 1; |
20600 | 413 |
in (thm', max') end; |
414 |
val (thms', maxidx) = fold_map incr_thm thms 0; |
|
22033 | 415 |
val (ty1::tys) = map CodegenFunc.typ_func thms'; |
20600 | 416 |
fun unify ty env = Sign.typ_unify thy (ty1, ty) env |
417 |
handle Type.TUNIFY => |
|
22197 | 418 |
error ("Type unificaton failed, while unifying defining equations\n" |
20600 | 419 |
^ (cat_lines o map Display.string_of_thm) thms |
420 |
^ "\nwith types\n" |
|
22197 | 421 |
^ (cat_lines o map (CodegenConsts.string_of_typ thy)) (ty1 :: tys)); |
20600 | 422 |
val (env, _) = fold unify tys (Vartab.empty, maxidx) |
423 |
val instT = Vartab.fold (fn (x_i, (sort, ty)) => |
|
424 |
cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env []; |
|
21119 | 425 |
in map (Thm.instantiate (instT, [])) thms' end; |
20600 | 426 |
|
20844 | 427 |
fun certify_const thy c c_thms = |
20600 | 428 |
let |
429 |
fun cert (c', thm) = if CodegenConsts.eq_const (c, c') |
|
22197 | 430 |
then thm else error ("Wrong head of defining equation,\nexpected constant " |
22033 | 431 |
^ CodegenConsts.string_of_const thy c ^ "\n" ^ string_of_thm thm) |
20844 | 432 |
in map cert c_thms end; |
20600 | 433 |
|
434 |
||
435 |
||
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
436 |
(** operational sort algebra and class discipline **) |
22050 | 437 |
|
438 |
local |
|
439 |
||
440 |
fun aggr_neutr f y [] = y |
|
441 |
| aggr_neutr f y (x::xs) = aggr_neutr f (f y x) xs; |
|
442 |
||
443 |
fun aggregate f [] = NONE |
|
444 |
| aggregate f (x::xs) = SOME (aggr_neutr f x xs); |
|
445 |
||
446 |
fun inter_sorts thy = |
|
447 |
let |
|
448 |
val algebra = Sign.classes_of thy; |
|
449 |
val inters = curry (Sorts.inter_sort algebra); |
|
450 |
in aggregate (map2 inters) end; |
|
451 |
||
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
452 |
fun specific_constraints thy (class, tyco) = |
22050 | 453 |
let |
454 |
val vs = Name.invents Name.context "" (Sign.arity_number thy tyco); |
|
455 |
val clsops = (these o Option.map snd o try (AxClass.params_of_class thy)) class; |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
456 |
val funcs = clsops |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
457 |
|> map (fn (clsop, _) => CodegenConsts.norm thy (clsop, [Type (tyco, map (TFree o rpair []) vs)])) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
458 |
|> map (Consttab.lookup ((the_funcs o get_exec) thy)) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
459 |
|> (map o Option.map) (Susp.force o fst) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
460 |
|> maps these |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
461 |
|> map (Thm.transfer thy); |
22050 | 462 |
val sorts = map (map (snd o dest_TVar) o snd o dest_Type o the_single |
463 |
o Sign.const_typargs thy o fst o CodegenFunc.dest_func) funcs; |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
464 |
in sorts end; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
465 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
466 |
fun weakest_constraints thy (class, tyco) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
467 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
468 |
val all_superclasses = class :: Graph.all_succs ((#classes o Sorts.rep_algebra o Sign.classes_of) thy) [class]; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
469 |
in case inter_sorts thy (maps (fn class => specific_constraints thy (class, tyco)) all_superclasses) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
470 |
of SOME sorts => sorts |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
471 |
| NONE => Sign.arity_sorts thy tyco [class] |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
472 |
end; |
22050 | 473 |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
474 |
fun strongest_constraints thy (class, tyco) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
475 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
476 |
val algebra = Sign.classes_of thy; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
477 |
val all_subclasses = class :: Graph.all_preds ((#classes o Sorts.rep_algebra) algebra) [class]; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
478 |
val inst_subclasses = filter (can (Sorts.mg_domain algebra tyco) o single) all_subclasses; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
479 |
in case inter_sorts thy (maps (fn class => specific_constraints thy (class, tyco)) inst_subclasses) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
480 |
of SOME sorts => sorts |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
481 |
| NONE => replicate |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
482 |
(Sign.arity_number thy tyco) (Sign.certify_sort thy (Sign.all_classes thy)) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
483 |
end; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
484 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
485 |
fun gen_classop_typ constr thy class (c, tyco) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
486 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
487 |
val (var, cs) = try (AxClass.params_of_class thy) class |> the_default ("'a", []) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
488 |
val ty = (the o AList.lookup (op =) cs) c; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
489 |
val sort_args = Name.names (Name.declare var Name.context) "'a" |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
490 |
(constr thy (class, tyco)); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
491 |
val ty_inst = Type (tyco, map TFree sort_args); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
492 |
in Logic.varifyT (map_type_tfree (K ty_inst) ty) end; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
493 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
494 |
(*FIXME: make distinct step: building algebra from code theorems*) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
495 |
fun retrieve_algebra thy operational = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
496 |
Sorts.subalgebra (Sign.pp thy) operational |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
497 |
(weakest_constraints thy) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
498 |
(Sign.classes_of thy); |
22050 | 499 |
|
500 |
in |
|
501 |
||
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
502 |
fun coregular_algebra thy = retrieve_algebra thy (K true) |> snd; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
503 |
fun operational_algebra thy = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
504 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
505 |
fun add_iff_operational class classes = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
506 |
if (not o null o these o Option.map #params o try (AxClass.get_definition thy)) class |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
507 |
orelse (length o gen_inter (op =)) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
508 |
((Sign.certify_sort thy o Sign.super_classes thy) class, classes) >= 2 |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
509 |
then class :: classes |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
510 |
else classes; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
511 |
val operational_classes = fold add_iff_operational (Sign.all_classes thy) [] |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
512 |
in retrieve_algebra thy (member (op =) operational_classes) end; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
513 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
514 |
val classop_weakest_typ = gen_classop_typ weakest_constraints; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
515 |
val classop_strongest_typ = gen_classop_typ strongest_constraints; |
22050 | 516 |
|
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
517 |
fun gen_mk_func_typ strict thm = |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
518 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
519 |
val thy = Thm.theory_of_thm thm; |
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
520 |
val raw_funcs = CodegenFunc.mk_func strict thm; |
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
521 |
val error_warning = if strict then error else warning #> K NONE; |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
522 |
fun check_typ_classop class (const as (c, [Type (tyco, _)]), thm) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
523 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
524 |
val ((_, ty), _) = CodegenFunc.dest_func thm; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
525 |
val ty_decl = classop_weakest_typ thy class (c, tyco); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
526 |
val ty_strongest = classop_strongest_typ thy class (c, tyco); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
527 |
fun constrain thm = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
528 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
529 |
val max = Thm.maxidx_of thm + 1; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
530 |
val ty_decl' = Logic.incr_tvar max ty_decl; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
531 |
val ((_, ty'), _) = CodegenFunc.dest_func thm; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
532 |
val (env, _) = Sign.typ_unify thy (ty_decl', ty') (Vartab.empty, max); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
533 |
val instT = Vartab.fold (fn (x_i, (sort, ty)) => |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
534 |
cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env []; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
535 |
in Thm.instantiate (instT, []) thm end; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
536 |
in if Sign.typ_instance thy (ty_strongest, ty) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
537 |
then if Sign.typ_instance thy (ty, ty_decl) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
538 |
then SOME (const, thm) |
22197 | 539 |
else (warning ("Constraining type\n" ^ CodegenConsts.string_of_typ thy ty |
540 |
^ "\nof defining equation\n" |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
541 |
^ string_of_thm thm |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
542 |
^ "\nto permitted most general type\n" |
22197 | 543 |
^ CodegenConsts.string_of_typ thy ty_decl); |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
544 |
SOME (const, constrain thm)) |
22197 | 545 |
else error_warning ("Type\n" ^ CodegenConsts.string_of_typ thy ty |
546 |
^ "\nof defining equation\n" |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
547 |
^ string_of_thm thm |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
548 |
^ "\nis incompatible with permitted least general type\n" |
22197 | 549 |
^ CodegenConsts.string_of_typ thy ty_strongest) |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
550 |
end |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
551 |
| check_typ_classop class ((c, [_]), thm) = |
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
552 |
error_warning ("Illegal type for class operation " ^ quote c |
22197 | 553 |
^ "\nin defining equation\n" |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
554 |
^ string_of_thm thm); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
555 |
fun check_typ_fun (const as (c, _), thm) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
556 |
let |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
557 |
val ((_, ty), _) = CodegenFunc.dest_func thm; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
558 |
val ty_decl = Sign.the_const_type thy c; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
559 |
in if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
560 |
then SOME (const, thm) |
22197 | 561 |
else error_warning ("Type\n" ^ CodegenConsts.string_of_typ thy ty |
562 |
^ "\nof defining equation\n" |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
563 |
^ string_of_thm thm |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
564 |
^ "\nis incompatible declared function type\n" |
22197 | 565 |
^ CodegenConsts.string_of_typ thy ty_decl) |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
566 |
end; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
567 |
fun check_typ (const as (c, tys), thm) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
568 |
case AxClass.class_of_param thy c |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
569 |
of SOME class => check_typ_classop class (const, thm) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
570 |
| NONE => check_typ_fun (const, thm); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
571 |
val funcs = map_filter check_typ raw_funcs; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
572 |
in funcs end; |
22050 | 573 |
|
574 |
end; |
|
575 |
||
576 |
||
20600 | 577 |
(** interfaces **) |
578 |
||
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
579 |
fun add_func strict thm thy = |
20600 | 580 |
let |
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
581 |
val funcs = gen_mk_func_typ strict thm; |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
582 |
val cs = map fst funcs; |
20600 | 583 |
in |
584 |
map_exec_purge (SOME cs) (map_funcs |
|
585 |
(fold (fn (c, thm) => Consttab.map_default |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
586 |
(c, (Susp.value [], [])) (add_thm thm)) funcs)) thy |
20600 | 587 |
end; |
588 |
||
22319 | 589 |
fun delete_force msg key xs = |
590 |
if AList.defined (op =) xs key then AList.delete (op =) key xs |
|
591 |
else error ("No such " ^ msg ^ ": " ^ quote key); |
|
592 |
||
20600 | 593 |
fun del_func thm thy = |
594 |
let |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
595 |
val funcs = gen_mk_func_typ false thm; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
596 |
val cs = map fst funcs; |
20600 | 597 |
in |
598 |
map_exec_purge (SOME cs) (map_funcs |
|
599 |
(fold (fn (c, thm) => Consttab.map_entry c |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
600 |
(del_thm thm)) funcs)) thy |
20600 | 601 |
end; |
602 |
||
603 |
fun add_funcl (c, lthms) thy = |
|
604 |
let |
|
605 |
val c' = CodegenConsts.norm thy c; |
|
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
606 |
val lthms' = certificate thy (fn thy => certify_const thy c' o maps (CodegenFunc.mk_func true)) lthms; |
20600 | 607 |
in |
608 |
map_exec_purge (SOME [c]) (map_funcs (Consttab.map_default (c', (Susp.value [], [])) |
|
609 |
(add_lthms lthms'))) thy |
|
610 |
end; |
|
611 |
||
22423 | 612 |
local |
613 |
||
614 |
fun consts_of_cos thy tyco vs cos = |
|
615 |
let |
|
616 |
val dty = Type (tyco, map TFree vs); |
|
617 |
fun mk_co (co, tys) = CodegenConsts.norm_of_typ thy (co, tys ---> dty); |
|
618 |
in map mk_co cos end; |
|
619 |
||
620 |
fun co_of_const thy (c, ty) = |
|
20600 | 621 |
let |
22423 | 622 |
fun err () = error |
623 |
("Illegal type for datatype constructor: " ^ Sign.string_of_typ thy ty); |
|
624 |
val (tys, ty') = strip_type ty; |
|
625 |
val (tyco, vs) = ((apsnd o map) dest_TFree o dest_Type) ty' |
|
626 |
handle TYPE _ => err (); |
|
627 |
val sorts = if has_duplicates (eq_fst op =) vs then err () |
|
628 |
else map snd vs; |
|
629 |
val vs_names = Name.invent_list [] "'a" (length vs); |
|
630 |
val vs_map = map fst vs ~~ vs_names; |
|
631 |
val vs' = vs_names ~~ sorts; |
|
632 |
val tys' = (map o map_type_tfree) (fn (v, sort) => |
|
633 |
(TFree ((the o AList.lookup (op =) vs_map) v, sort))) tys |
|
634 |
handle Option => err (); |
|
635 |
in (tyco, (vs', (c, tys'))) end; |
|
20600 | 636 |
|
637 |
fun del_datatype tyco thy = |
|
22423 | 638 |
case Symtab.lookup ((the_dtyps o get_exec) thy) tyco |
639 |
of SOME (vs, cos) => let |
|
640 |
val consts = consts_of_cos thy tyco vs cos; |
|
641 |
val del = |
|
642 |
map_dtyps (Symtab.delete tyco) |
|
643 |
#> map_dconstrs (fold Consttab.delete consts) |
|
644 |
in map_exec_purge (SOME consts) del thy end |
|
645 |
| NONE => thy; |
|
646 |
||
647 |
(*FIXME integrate this auxiliary properly*) |
|
648 |
||
649 |
in |
|
650 |
||
651 |
fun add_datatype (tyco, (vs_cos as (vs, cos))) thy = |
|
20600 | 652 |
let |
22423 | 653 |
val consts = consts_of_cos thy tyco vs cos; |
654 |
val add = |
|
655 |
map_dtyps (Symtab.update_new (tyco, vs_cos)) |
|
656 |
#> map_dconstrs (fold (fn c => Consttab.update (c, tyco)) consts) |
|
657 |
in |
|
658 |
thy |
|
659 |
|> del_datatype tyco |
|
660 |
|> map_exec_purge (SOME consts) add |
|
661 |
end; |
|
662 |
||
663 |
fun add_datatype_consts cs thy = |
|
664 |
let |
|
665 |
val raw_cos = map (co_of_const thy) cs; |
|
666 |
val (tyco, (vs_names, sorts_cos)) = if (length o distinct (eq_fst op =)) raw_cos = 1 |
|
667 |
then ((fst o hd) raw_cos, ((map fst o fst o snd o hd) raw_cos, |
|
668 |
map ((apfst o map) snd o snd) raw_cos)) |
|
669 |
else error ("Term constructors not referring to the same type: " |
|
670 |
^ commas (map (CodegenConsts.string_of_const_typ thy) cs)); |
|
671 |
val sorts = foldr1 ((uncurry o map2 o curry o Sorts.inter_sort) (Sign.classes_of thy)) |
|
672 |
(map fst sorts_cos); |
|
673 |
val cos = map snd sorts_cos; |
|
674 |
val vs = vs_names ~~ sorts; |
|
675 |
in |
|
676 |
thy |
|
677 |
|> add_datatype (tyco, (vs, cos)) |
|
678 |
end; |
|
679 |
||
680 |
fun add_datatype_consts_cmd raw_cs thy = |
|
681 |
let |
|
682 |
val cs = map (apsnd Logic.unvarifyT o CodegenConsts.typ_of_inst thy |
|
683 |
o CodegenConsts.read_const thy) raw_cs |
|
684 |
in |
|
685 |
thy |
|
686 |
|> add_datatype_consts cs |
|
687 |
end; |
|
688 |
||
689 |
end; (*local*) |
|
20600 | 690 |
|
691 |
fun add_inline thm thy = |
|
22423 | 692 |
(map_exec_purge NONE o map_preproc o apfst o apfst) |
693 |
(fold (insert Thm.eq_thm) (CodegenFunc.mk_rew thm)) thy; |
|
694 |
(*fully applied in order to get right context for mk_rew!*) |
|
20600 | 695 |
|
696 |
fun del_inline thm thy = |
|
22423 | 697 |
(map_exec_purge NONE o map_preproc o apfst o apfst) |
698 |
(fold (remove Thm.eq_thm) (CodegenFunc.mk_rew thm)) thy; |
|
699 |
(*fully applied in order to get right context for mk_rew!*) |
|
20844 | 700 |
|
22050 | 701 |
fun add_inline_proc (name, f) = |
22423 | 702 |
(map_exec_purge NONE o map_preproc o apfst o apsnd) |
703 |
(AList.update (op =) (name, (serial (), f))); |
|
22050 | 704 |
|
705 |
fun del_inline_proc name = |
|
22423 | 706 |
(map_exec_purge NONE o map_preproc o apfst o apsnd) |
707 |
(delete_force "inline procedure" name); |
|
20844 | 708 |
|
22050 | 709 |
fun add_preproc (name, f) = |
710 |
(map_exec_purge NONE o map_preproc o apsnd) (AList.update (op =) (name, (serial (), f))); |
|
711 |
||
712 |
fun del_preproc name = |
|
22319 | 713 |
(map_exec_purge NONE o map_preproc o apsnd) (delete_force "preprocessor" name); |
20844 | 714 |
|
715 |
local |
|
716 |
||
717 |
fun gen_apply_inline_proc prep post thy f x = |
|
718 |
let |
|
719 |
val cts = prep x; |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
720 |
val rews = map CodegenFunc.assert_rew (f thy cts); |
20844 | 721 |
in post rews x end; |
722 |
||
723 |
val apply_inline_proc = gen_apply_inline_proc (maps |
|
724 |
((fn [args, rhs] => rhs :: (snd o Drule.strip_comb) args) o snd o Drule.strip_comb o Thm.cprop_of)) |
|
22033 | 725 |
(fn rews => map (CodegenFunc.rewrite_func rews)); |
20844 | 726 |
val apply_inline_proc_cterm = gen_apply_inline_proc single |
21708 | 727 |
(MetaSimplifier.rewrite false); |
20600 | 728 |
|
20844 | 729 |
fun apply_preproc thy f [] = [] |
730 |
| apply_preproc thy f (thms as (thm :: _)) = |
|
731 |
let |
|
732 |
val thms' = f thy thms; |
|
22033 | 733 |
val c = (CodegenConsts.norm_of_typ thy o fst o CodegenFunc.dest_func) thm; |
734 |
in (certify_const thy c o map CodegenFunc.mk_head) thms' end; |
|
20844 | 735 |
|
736 |
fun cmp_thms thy = |
|
22033 | 737 |
make_ord (fn (thm1, thm2) => not (Sign.typ_instance thy (CodegenFunc.typ_func thm1, CodegenFunc.typ_func thm2))); |
20844 | 738 |
|
739 |
fun rhs_conv conv thm = |
|
740 |
let |
|
741 |
val thm' = (conv o snd o Drule.dest_equals o Thm.cprop_of) thm; |
|
742 |
in Thm.transitive thm thm' end |
|
743 |
||
744 |
in |
|
20600 | 745 |
|
746 |
fun preprocess thy thms = |
|
20844 | 747 |
thms |
22050 | 748 |
|> fold (fn (_, (_, f)) => apply_preproc thy f) ((#preprocs o the_preproc o get_exec) thy) |
22033 | 749 |
|> map (CodegenFunc.rewrite_func ((#inlines o the_preproc o get_exec) thy)) |
22050 | 750 |
|> fold (fn (_, (_, f)) => apply_inline_proc thy f) ((#inline_procs o the_preproc o get_exec) thy) |
22197 | 751 |
(*FIXME - must check: rewrite rule, defining equation, proper constant |> map (snd o check_func false thy) *) |
20844 | 752 |
|> sort (cmp_thms thy) |
22033 | 753 |
|> common_typ_funcs; |
20600 | 754 |
|
22033 | 755 |
fun preprocess_cterm ct = |
756 |
let |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
757 |
val thy = Thm.theory_of_cterm ct; |
22033 | 758 |
in |
759 |
ct |
|
22319 | 760 |
|> MetaSimplifier.rewrite false ((#inlines o the_preproc o get_exec) thy) |
22050 | 761 |
|> fold (fn (_, (_, f)) => rhs_conv (apply_inline_proc_cterm thy f)) |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
762 |
((#inline_procs o the_preproc o get_exec) thy) |
22033 | 763 |
end; |
20844 | 764 |
|
765 |
end; (*local*) |
|
20600 | 766 |
|
22423 | 767 |
fun get_datatype thy tyco = |
768 |
case Symtab.lookup ((the_dtyps o get_exec) thy) tyco |
|
769 |
of SOME spec => spec |
|
770 |
| NONE => Sign.arity_number thy tyco |
|
771 |
|> Name.invents Name.context "'a" |
|
772 |
|> map (rpair []) |
|
773 |
|> rpair []; |
|
774 |
||
775 |
fun get_datatype_of_constr thy = |
|
776 |
Consttab.lookup ((the_dcontrs o get_exec) thy); |
|
777 |
||
778 |
fun get_datatype_constr thy const = |
|
779 |
case Consttab.lookup ((the_dcontrs o get_exec) thy) const |
|
780 |
of SOME tyco => let |
|
781 |
val (vs, cs) = get_datatype thy tyco; |
|
782 |
(*FIXME continue here*) |
|
783 |
in NONE end |
|
784 |
| NONE => NONE; |
|
785 |
||
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
786 |
local |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
787 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
788 |
fun get_funcs thy const = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
789 |
Consttab.lookup ((the_funcs o get_exec) thy) const |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
790 |
|> Option.map (Susp.force o fst) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
791 |
|> these |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
792 |
|> map (Thm.transfer thy); |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
793 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
794 |
in |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
795 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
796 |
fun these_funcs thy const = |
20600 | 797 |
let |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
798 |
fun get_prim_def_funcs (const as (c, tys)) = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
799 |
case CodegenConsts.find_def thy const |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
800 |
of SOME (_, thm) => |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
801 |
thm |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
802 |
|> Thm.transfer thy |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
803 |
|> gen_mk_func_typ false |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
804 |
|> map (CodegenFunc.expand_eta ~1 o snd) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
805 |
| NONE => [] |
20600 | 806 |
fun drop_refl thy = filter_out (is_equal o Term.fast_term_ord o Logic.dest_equals |
807 |
o ObjectLogic.drop_judgment thy o Drule.plain_prop_of); |
|
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
808 |
val funcs = case get_funcs thy const |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
809 |
of [] => get_prim_def_funcs const |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
810 |
| funcs => funcs |
20600 | 811 |
in |
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
812 |
funcs |
20600 | 813 |
|> preprocess thy |
814 |
|> drop_refl thy |
|
815 |
end; |
|
816 |
||
22184
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
817 |
fun tap_typ thy const = |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
818 |
case get_funcs thy const |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
819 |
of thm :: _ => SOME (CodegenFunc.typ_func thm) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
820 |
| [] => NONE; |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
821 |
|
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
822 |
end; (*local*) |
a125f38a559a
added explicit maintainance of coregular code theorems for overloaded constants
haftmann
parents:
22050
diff
changeset
|
823 |
|
20600 | 824 |
|
825 |
(** code attributes **) |
|
826 |
||
827 |
local |
|
828 |
fun add_simple_attribute (name, f) = |
|
829 |
(Codegen.add_attribute name o (Scan.succeed o Thm.declaration_attribute)) |
|
20897 | 830 |
(fn th => Context.mapping (f th) I); |
20600 | 831 |
in |
832 |
val _ = map (Context.add_setup o add_simple_attribute) [ |
|
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22423
diff
changeset
|
833 |
("func", add_func true), |
20600 | 834 |
("nofunc", del_func), |
835 |
("unfold", (fn thm => Codegen.add_unfold thm #> add_inline thm)), |
|
836 |
("inline", add_inline), |
|
837 |
("noinline", del_inline) |
|
838 |
] |
|
839 |
end; (*local*) |
|
840 |
||
22303 | 841 |
|
842 |
(** Isar setup **) |
|
843 |
||
844 |
local |
|
845 |
||
846 |
structure P = OuterParse |
|
847 |
and K = OuterKeyword |
|
848 |
||
849 |
val print_codesetupK = "print_codesetup"; |
|
22423 | 850 |
val code_datatypeK = "code_datatype"; |
22303 | 851 |
|
852 |
in |
|
853 |
||
854 |
val print_codesetupP = |
|
22423 | 855 |
OuterSyntax.improper_command print_codesetupK "print code generator setup of this theory" K.diag |
22303 | 856 |
(Scan.succeed |
857 |
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep (CodeData.print o Toplevel.theory_of))); |
|
858 |
||
22423 | 859 |
val code_datatypeP = |
860 |
OuterSyntax.command code_datatypeK "define set of code datatype constructors" K.thy_decl ( |
|
861 |
Scan.repeat1 P.term |
|
862 |
>> (Toplevel.theory o add_datatype_consts_cmd) |
|
863 |
); |
|
864 |
||
865 |
||
866 |
val _ = OuterSyntax.add_parsers [print_codesetupP, code_datatypeP]; |
|
22303 | 867 |
|
868 |
end; (*local*) |
|
869 |
||
20600 | 870 |
end; (*struct*) |
871 |
||
872 |
||
873 |
||
874 |
(** type-safe interfaces for data depedent on executable content **) |
|
875 |
||
876 |
signature CODE_DATA_ARGS = |
|
877 |
sig |
|
878 |
val name: string |
|
879 |
type T |
|
880 |
val empty: T |
|
881 |
val merge: Pretty.pp -> T * T -> T |
|
20937 | 882 |
val purge: theory option -> CodegenConsts.const list option -> T -> T |
20600 | 883 |
end; |
884 |
||
885 |
signature CODE_DATA = |
|
886 |
sig |
|
887 |
type T |
|
888 |
val init: theory -> theory |
|
889 |
val get: theory -> T |
|
890 |
val change: theory -> (T -> T) -> T |
|
891 |
val change_yield: theory -> (T -> 'a * T) -> 'a * T |
|
892 |
end; |
|
893 |
||
894 |
functor CodeDataFun(Data: CODE_DATA_ARGS): CODE_DATA = |
|
895 |
struct |
|
896 |
||
897 |
type T = Data.T; |
|
898 |
exception Data of T; |
|
899 |
fun dest (Data x) = x |
|
900 |
||
901 |
val kind = CodegenData.declare Data.name (Data Data.empty) |
|
902 |
(fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2))) |
|
20937 | 903 |
(fn thy_opt => fn cs => fn Data x => Data (Data.purge thy_opt cs x)); |
20600 | 904 |
|
905 |
val init = CodegenData.init kind; |
|
22210 | 906 |
|
20600 | 907 |
fun get thy = CodegenData.get kind dest ((! o snd o CodegenData.CodeData.get) thy); |
22210 | 908 |
|
20600 | 909 |
fun change thy f = |
910 |
let |
|
911 |
val data_ref = (snd o CodegenData.CodeData.get) thy; |
|
912 |
val x = (f o CodegenData.get kind dest o !) data_ref; |
|
913 |
val data = CodegenData.put kind Data x (! data_ref); |
|
914 |
in (data_ref := data; x) end; |
|
22210 | 915 |
|
20600 | 916 |
fun change_yield thy f = |
917 |
let |
|
918 |
val data_ref = (snd o CodegenData.CodeData.get) thy; |
|
919 |
val (y, x) = (f o CodegenData.get kind dest o !) data_ref; |
|
920 |
val data = CodegenData.put kind Data x (! data_ref); |
|
921 |
in (data_ref := data; (y, x)) end; |
|
922 |
||
923 |
end; |
|
924 |
||
925 |
structure CodegenData : CODEGEN_DATA = |
|
926 |
struct |
|
927 |
||
928 |
open CodegenData; |
|
929 |
||
930 |
end; |