author | berghofe |
Tue, 01 Jun 2004 14:59:54 +0200 | |
changeset 14859 | b4be6bdcbb94 |
parent 14560 | 529464cffbfe |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
12453 | 1 |
(* Title: HOL/inductive_codegen.ML |
11537 | 2 |
ID: $Id$ |
11539 | 3 |
Author: Stefan Berghofer, TU Muenchen |
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
11537 | 5 |
|
11539 | 6 |
Code generator for inductive predicates. |
11537 | 7 |
*) |
8 |
||
9 |
signature INDUCTIVE_CODEGEN = |
|
10 |
sig |
|
12557 | 11 |
val add : theory attribute |
11537 | 12 |
val setup : (theory -> theory) list |
13 |
end; |
|
14 |
||
15 |
structure InductiveCodegen : INDUCTIVE_CODEGEN = |
|
16 |
struct |
|
17 |
||
18 |
open Codegen; |
|
19 |
||
12557 | 20 |
(**** theory data ****) |
21 |
||
22 |
structure CodegenArgs = |
|
23 |
struct |
|
24 |
val name = "HOL/inductive_codegen"; |
|
14162
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
25 |
type T = thm list Symtab.table * unit Graph.T; |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
26 |
val empty = (Symtab.empty, Graph.empty); |
12557 | 27 |
val copy = I; |
28 |
val prep_ext = I; |
|
14162
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
29 |
fun merge ((tab1, graph1), (tab2, graph2)) = |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
30 |
(Symtab.merge_multi Drule.eq_thm_prop (tab1, tab2), |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
31 |
Graph.merge (K true) (graph1, graph2)); |
12557 | 32 |
fun print _ _ = (); |
33 |
end; |
|
34 |
||
35 |
structure CodegenData = TheoryDataFun(CodegenArgs); |
|
36 |
||
37 |
fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^ |
|
38 |
string_of_thm thm); |
|
11537 | 39 |
|
14162
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
40 |
fun add_node (g, x) = Graph.new_node (x, ()) g handle Graph.DUP _ => g; |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
41 |
|
12557 | 42 |
fun add (p as (thy, thm)) = |
14162
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
43 |
let val (tab, graph) = CodegenData.get thy; |
12557 | 44 |
in (case concl_of thm of |
45 |
_ $ (Const ("op :", _) $ _ $ t) => (case head_of t of |
|
14162
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
46 |
Const (s, _) => |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
47 |
let val cs = foldr add_term_consts (prems_of thm, []) |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
48 |
in (CodegenData.put |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
49 |
(Symtab.update ((s, |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
50 |
if_none (Symtab.lookup (tab, s)) [] @ [thm]), tab), |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
51 |
foldr (uncurry (Graph.add_edge o pair s)) |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
52 |
(cs, foldl add_node (graph, s :: cs))) thy, thm) |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
53 |
end |
12557 | 54 |
| _ => (warn thm; p)) |
55 |
| _ => (warn thm; p)) |
|
12562 | 56 |
end; |
12557 | 57 |
|
58 |
fun get_clauses thy s = |
|
14162
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
59 |
let val (tab, graph) = CodegenData.get thy |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
60 |
in case Symtab.lookup (tab, s) of |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
61 |
None => (case InductivePackage.get_inductive thy s of |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
62 |
None => None |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
63 |
| Some ({names, ...}, {intrs, ...}) => Some (names, intrs)) |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
64 |
| Some _ => |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
65 |
let val Some names = find_first |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
66 |
(fn xs => s mem xs) (Graph.strong_conn graph) |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
67 |
in Some (names, |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
68 |
flat (map (fn s => the (Symtab.lookup (tab, s))) names)) |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
69 |
end |
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents:
13105
diff
changeset
|
70 |
end; |
12557 | 71 |
|
72 |
||
73 |
(**** improper tuples ****) |
|
11537 | 74 |
|
75 |
fun prod_factors p (Const ("Pair", _) $ t $ u) = |
|
76 |
p :: prod_factors (1::p) t @ prod_factors (2::p) u |
|
77 |
| prod_factors p _ = []; |
|
78 |
||
79 |
fun split_prod p ps t = if p mem ps then (case t of |
|
80 |
Const ("Pair", _) $ t $ u => |
|
81 |
split_prod (1::p) ps t @ split_prod (2::p) ps u |
|
82 |
| _ => error "Inconsistent use of products") else [t]; |
|
83 |
||
12557 | 84 |
datatype factors = FVar of int list list | FFix of int list list; |
85 |
||
86 |
exception Factors; |
|
87 |
||
88 |
fun mg_factor (FVar f) (FVar f') = FVar (f inter f') |
|
89 |
| mg_factor (FVar f) (FFix f') = |
|
90 |
if f' subset f then FFix f' else raise Factors |
|
91 |
| mg_factor (FFix f) (FVar f') = |
|
92 |
if f subset f' then FFix f else raise Factors |
|
93 |
| mg_factor (FFix f) (FFix f') = |
|
94 |
if f subset f' andalso f' subset f then FFix f else raise Factors; |
|
95 |
||
96 |
fun dest_factors (FVar f) = f |
|
97 |
| dest_factors (FFix f) = f; |
|
98 |
||
99 |
fun infer_factors sg extra_fs (fs, (optf, t)) = |
|
100 |
let fun err s = error (s ^ "\n" ^ Sign.string_of_term sg t) |
|
101 |
in (case (optf, strip_comb t) of |
|
102 |
(Some f, (Const (name, _), args)) => |
|
103 |
(case assoc (extra_fs, name) of |
|
104 |
None => overwrite (fs, (name, if_none |
|
105 |
(apsome (mg_factor f) (assoc (fs, name))) f)) |
|
106 |
| Some (fs', f') => (mg_factor f (FFix f'); |
|
107 |
foldl (infer_factors sg extra_fs) |
|
108 |
(fs, map (apsome FFix) fs' ~~ args))) |
|
109 |
| (Some f, (Var ((name, _), _), [])) => |
|
110 |
overwrite (fs, (name, if_none |
|
111 |
(apsome (mg_factor f) (assoc (fs, name))) f)) |
|
112 |
| (None, _) => fs |
|
113 |
| _ => err "Illegal term") |
|
114 |
handle Factors => err "Product factor mismatch in" |
|
115 |
end; |
|
116 |
||
11537 | 117 |
fun string_of_factors p ps = if p mem ps then |
118 |
"(" ^ string_of_factors (1::p) ps ^ ", " ^ string_of_factors (2::p) ps ^ ")" |
|
119 |
else "_"; |
|
120 |
||
12557 | 121 |
|
11537 | 122 |
(**** check if a term contains only constructor functions ****) |
123 |
||
124 |
fun is_constrt thy = |
|
125 |
let |
|
126 |
val cnstrs = flat (flat (map |
|
127 |
(map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd) |
|
128 |
(Symtab.dest (DatatypePackage.get_datatypes thy)))); |
|
129 |
fun check t = (case strip_comb t of |
|
130 |
(Var _, []) => true |
|
131 |
| (Const (s, _), ts) => (case assoc (cnstrs, s) of |
|
132 |
None => false |
|
133 |
| Some i => length ts = i andalso forall check ts) |
|
134 |
| _ => false) |
|
135 |
in check end; |
|
136 |
||
137 |
(**** check if a type is an equality type (i.e. doesn't contain fun) ****) |
|
138 |
||
139 |
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts |
|
140 |
| is_eqT _ = true; |
|
141 |
||
142 |
(**** mode inference ****) |
|
143 |
||
144 |
val term_vs = map (fst o fst o dest_Var) o term_vars; |
|
145 |
val terms_vs = distinct o flat o (map term_vs); |
|
146 |
||
13038 | 147 |
fun assoc' s tab key = (case assoc (tab, key) of |
148 |
None => error ("Unable to determine " ^ s ^ " of " ^ quote key) |
|
149 |
| Some x => x); |
|
150 |
||
11537 | 151 |
(** collect all Vars in a term (with duplicates!) **) |
152 |
fun term_vTs t = map (apfst fst o dest_Var) |
|
153 |
(filter is_Var (foldl_aterms (op :: o Library.swap) ([], t))); |
|
154 |
||
155 |
fun known_args _ _ [] = [] |
|
156 |
| known_args vs i (t::ts) = if term_vs t subset vs then i::known_args vs (i+1) ts |
|
157 |
else known_args vs (i+1) ts; |
|
158 |
||
159 |
fun get_args _ _ [] = ([], []) |
|
160 |
| get_args is i (x::xs) = (if i mem is then apfst else apsnd) (cons x) |
|
161 |
(get_args is (i+1) xs); |
|
162 |
||
163 |
fun merge xs [] = xs |
|
164 |
| merge [] ys = ys |
|
165 |
| merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys) |
|
166 |
else y::merge (x::xs) ys; |
|
167 |
||
168 |
fun subsets i j = if i <= j then |
|
169 |
let val is = subsets (i+1) j |
|
170 |
in merge (map (fn ks => i::ks) is) is end |
|
171 |
else [[]]; |
|
172 |
||
12557 | 173 |
fun cprod ([], ys) = [] |
174 |
| cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys); |
|
175 |
||
176 |
fun cprods xss = foldr (map op :: o cprod) (xss, [[]]); |
|
177 |
||
178 |
datatype mode = Mode of (int list option list * int list) * mode option list; |
|
179 |
||
180 |
fun modes_of modes t = |
|
181 |
let |
|
182 |
fun mk_modes name args = flat |
|
183 |
(map (fn (m as (iss, is)) => map (Mode o pair m) (cprods (map |
|
184 |
(fn (None, _) => [None] |
|
185 |
| (Some js, arg) => map Some |
|
186 |
(filter (fn Mode ((_, js'), _) => js=js') (modes_of modes arg))) |
|
13038 | 187 |
(iss ~~ args)))) (assoc' "modes" modes name)) |
12557 | 188 |
|
189 |
in (case strip_comb t of |
|
14163
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
berghofe
parents:
14162
diff
changeset
|
190 |
(Const ("op =", Type (_, [T, _])), _) => |
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
berghofe
parents:
14162
diff
changeset
|
191 |
[Mode (([], [1]), []), Mode (([], [2]), [])] @ |
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
berghofe
parents:
14162
diff
changeset
|
192 |
(if is_eqT T then [Mode (([], [1, 2]), [])] else []) |
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
berghofe
parents:
14162
diff
changeset
|
193 |
| (Const (name, _), args) => mk_modes name args |
13038 | 194 |
| (Var ((name, _), _), args) => mk_modes name args |
195 |
| (Free (name, _), args) => mk_modes name args) |
|
12557 | 196 |
end; |
197 |
||
198 |
datatype indprem = Prem of term list * term | Sidecond of term; |
|
199 |
||
11537 | 200 |
fun select_mode_prem thy modes vs ps = |
201 |
find_first (is_some o snd) (ps ~~ map |
|
12557 | 202 |
(fn Prem (us, t) => find_first (fn Mode ((_, is), _) => |
11537 | 203 |
let |
204 |
val (_, out_ts) = get_args is 1 us; |
|
205 |
val vTs = flat (map term_vTs out_ts); |
|
206 |
val dupTs = map snd (duplicates vTs) @ |
|
207 |
mapfilter (curry assoc vTs) vs; |
|
208 |
in |
|
209 |
is subset known_args vs 1 us andalso |
|
210 |
forall (is_constrt thy) (snd (get_args is 1 us)) andalso |
|
12557 | 211 |
term_vs t subset vs andalso |
11537 | 212 |
forall is_eqT dupTs |
213 |
end) |
|
12557 | 214 |
(modes_of modes t) |
215 |
| Sidecond t => if term_vs t subset vs then Some (Mode (([], []), [])) |
|
216 |
else None) ps); |
|
11537 | 217 |
|
12557 | 218 |
fun check_mode_clause thy arg_vs modes (iss, is) (ts, ps) = |
11537 | 219 |
let |
12557 | 220 |
val modes' = modes @ mapfilter |
221 |
(fn (_, None) => None | (v, Some js) => Some (v, [([], js)])) |
|
222 |
(arg_vs ~~ iss); |
|
11537 | 223 |
fun check_mode_prems vs [] = Some vs |
12557 | 224 |
| check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of |
11537 | 225 |
None => None |
226 |
| Some (x, _) => check_mode_prems |
|
12557 | 227 |
(case x of Prem (us, _) => vs union terms_vs us | _ => vs) |
11537 | 228 |
(filter_out (equal x) ps)); |
12557 | 229 |
val (in_ts', _) = get_args is 1 ts; |
11537 | 230 |
val in_ts = filter (is_constrt thy) in_ts'; |
231 |
val in_vs = terms_vs in_ts; |
|
232 |
val concl_vs = terms_vs ts |
|
233 |
in |
|
14560 | 234 |
forall is_eqT (map snd (duplicates (flat (map term_vTs in_ts)))) andalso |
11537 | 235 |
(case check_mode_prems (arg_vs union in_vs) ps of |
236 |
None => false |
|
237 |
| Some vs => concl_vs subset vs) |
|
238 |
end; |
|
239 |
||
240 |
fun check_modes_pred thy arg_vs preds modes (p, ms) = |
|
241 |
let val Some rs = assoc (preds, p) |
|
242 |
in (p, filter (fn m => forall (check_mode_clause thy arg_vs modes m) rs) ms) end |
|
243 |
||
244 |
fun fixp f x = |
|
245 |
let val y = f x |
|
246 |
in if x = y then x else fixp f y end; |
|
247 |
||
12557 | 248 |
fun infer_modes thy extra_modes factors arg_vs preds = fixp (fn modes => |
11537 | 249 |
map (check_modes_pred thy arg_vs preds (modes @ extra_modes)) modes) |
12557 | 250 |
(map (fn (s, (fs, f)) => (s, cprod (cprods (map |
251 |
(fn None => [None] |
|
252 |
| Some f' => map Some (subsets 1 (length f' + 1))) fs), |
|
253 |
subsets 1 (length f + 1)))) factors); |
|
11537 | 254 |
|
255 |
(**** code generation ****) |
|
256 |
||
257 |
fun mk_eq (x::xs) = |
|
258 |
let fun mk_eqs _ [] = [] |
|
259 |
| mk_eqs a (b::cs) = Pretty.str (a ^ " = " ^ b) :: mk_eqs b cs |
|
260 |
in mk_eqs x xs end; |
|
261 |
||
262 |
fun mk_tuple xs = Pretty.block (Pretty.str "(" :: |
|
263 |
flat (separate [Pretty.str ",", Pretty.brk 1] (map single xs)) @ |
|
264 |
[Pretty.str ")"]); |
|
265 |
||
266 |
fun mk_v ((names, vs), s) = (case assoc (vs, s) of |
|
267 |
None => ((names, (s, [s])::vs), s) |
|
268 |
| Some xs => let val s' = variant names s in |
|
269 |
((s'::names, overwrite (vs, (s, s'::xs))), s') end); |
|
270 |
||
271 |
fun distinct_v (nvs, Var ((s, 0), T)) = |
|
272 |
apsnd (Var o rpair T o rpair 0) (mk_v (nvs, s)) |
|
273 |
| distinct_v (nvs, t $ u) = |
|
274 |
let |
|
275 |
val (nvs', t') = distinct_v (nvs, t); |
|
276 |
val (nvs'', u') = distinct_v (nvs', u); |
|
277 |
in (nvs'', t' $ u') end |
|
278 |
| distinct_v x = x; |
|
279 |
||
280 |
fun compile_match nvs eq_ps out_ps success_p fail_p = |
|
281 |
let val eqs = flat (separate [Pretty.str " andalso", Pretty.brk 1] |
|
282 |
(map single (flat (map (mk_eq o snd) nvs) @ eq_ps))); |
|
283 |
in |
|
284 |
Pretty.block |
|
285 |
([Pretty.str "(fn ", mk_tuple out_ps, Pretty.str " =>", Pretty.brk 1] @ |
|
286 |
(Pretty.block ((if eqs=[] then [] else Pretty.str "if " :: |
|
287 |
[Pretty.block eqs, Pretty.brk 1, Pretty.str "then "]) @ |
|
288 |
(success_p :: |
|
289 |
(if eqs=[] then [] else [Pretty.brk 1, Pretty.str "else ", fail_p]))) :: |
|
290 |
[Pretty.brk 1, Pretty.str "| _ => ", fail_p, Pretty.str ")"])) |
|
291 |
end; |
|
292 |
||
12557 | 293 |
fun modename thy s (iss, is) = space_implode "__" |
294 |
(mk_const_id (sign_of thy) s :: |
|
295 |
map (space_implode "_" o map string_of_int) (mapfilter I iss @ [is])); |
|
11537 | 296 |
|
12557 | 297 |
fun compile_expr thy dep brack (gr, (None, t)) = |
298 |
apsnd single (invoke_codegen thy dep brack (gr, t)) |
|
299 |
| compile_expr _ _ _ (gr, (Some _, Var ((name, _), _))) = |
|
300 |
(gr, [Pretty.str name]) |
|
301 |
| compile_expr thy dep brack (gr, (Some (Mode (mode, ms)), t)) = |
|
302 |
let |
|
303 |
val (Const (name, _), args) = strip_comb t; |
|
304 |
val (gr', ps) = foldl_map |
|
305 |
(compile_expr thy dep true) (gr, ms ~~ args); |
|
306 |
in (gr', (if brack andalso not (null ps) then |
|
307 |
single o parens o Pretty.block else I) |
|
308 |
(flat (separate [Pretty.brk 1] |
|
309 |
([Pretty.str (modename thy name mode)] :: ps)))) |
|
310 |
end; |
|
311 |
||
312 |
fun compile_clause thy gr dep all_vs arg_vs modes (iss, is) (ts, ps) = |
|
11537 | 313 |
let |
12557 | 314 |
val modes' = modes @ mapfilter |
315 |
(fn (_, None) => None | (v, Some js) => Some (v, [([], js)])) |
|
316 |
(arg_vs ~~ iss); |
|
317 |
||
11537 | 318 |
fun check_constrt ((names, eqs), t) = |
319 |
if is_constrt thy t then ((names, eqs), t) else |
|
320 |
let val s = variant names "x"; |
|
321 |
in ((s::names, (s, t)::eqs), Var ((s, 0), fastype_of t)) end; |
|
322 |
||
12557 | 323 |
val (in_ts, out_ts) = get_args is 1 ts; |
11537 | 324 |
val ((all_vs', eqs), in_ts') = |
325 |
foldl_map check_constrt ((all_vs, []), in_ts); |
|
326 |
||
327 |
fun compile_prems out_ts' vs names gr [] = |
|
328 |
let |
|
12453 | 329 |
val (gr2, out_ps) = foldl_map |
330 |
(invoke_codegen thy dep false) (gr, out_ts); |
|
11537 | 331 |
val (gr3, eq_ps) = foldl_map (fn (gr, (s, t)) => |
332 |
apsnd (Pretty.block o cons (Pretty.str (s ^ " = ")) o single) |
|
12453 | 333 |
(invoke_codegen thy dep false (gr, t))) (gr2, eqs); |
11537 | 334 |
val (nvs, out_ts'') = foldl_map distinct_v |
335 |
((names, map (fn x => (x, [x])) vs), out_ts'); |
|
12453 | 336 |
val (gr4, out_ps') = foldl_map |
337 |
(invoke_codegen thy dep false) (gr3, out_ts''); |
|
11537 | 338 |
in |
339 |
(gr4, compile_match (snd nvs) eq_ps out_ps' |
|
340 |
(Pretty.block [Pretty.str "Seq.single", Pretty.brk 1, mk_tuple out_ps]) |
|
341 |
(Pretty.str "Seq.empty")) |
|
342 |
end |
|
343 |
| compile_prems out_ts vs names gr ps = |
|
344 |
let |
|
345 |
val vs' = distinct (flat (vs :: map term_vs out_ts)); |
|
12557 | 346 |
val Some (p, mode as Some (Mode ((_, js), _))) = |
347 |
select_mode_prem thy modes' (arg_vs union vs') ps; |
|
11537 | 348 |
val ps' = filter_out (equal p) ps; |
349 |
in |
|
350 |
(case p of |
|
12557 | 351 |
Prem (us, t) => |
11537 | 352 |
let |
12557 | 353 |
val (in_ts, out_ts') = get_args js 1 us; |
12453 | 354 |
val (gr1, in_ps) = foldl_map |
355 |
(invoke_codegen thy dep false) (gr, in_ts); |
|
11537 | 356 |
val (nvs, out_ts'') = foldl_map distinct_v |
357 |
((names, map (fn x => (x, [x])) vs), out_ts); |
|
12557 | 358 |
val (gr2, out_ps) = foldl_map |
359 |
(invoke_codegen thy dep false) (gr1, out_ts''); |
|
360 |
val (gr3, ps) = compile_expr thy dep false (gr2, (mode, t)); |
|
11537 | 361 |
val (gr4, rest) = compile_prems out_ts' vs' (fst nvs) gr3 ps'; |
362 |
in |
|
363 |
(gr4, compile_match (snd nvs) [] out_ps |
|
12557 | 364 |
(Pretty.block (ps @ |
11537 | 365 |
[Pretty.brk 1, mk_tuple in_ps, |
366 |
Pretty.str " :->", Pretty.brk 1, rest])) |
|
367 |
(Pretty.str "Seq.empty")) |
|
368 |
end |
|
369 |
| Sidecond t => |
|
370 |
let |
|
12453 | 371 |
val (gr1, side_p) = invoke_codegen thy dep true (gr, t); |
11537 | 372 |
val (nvs, out_ts') = foldl_map distinct_v |
373 |
((names, map (fn x => (x, [x])) vs), out_ts); |
|
12453 | 374 |
val (gr2, out_ps) = foldl_map |
375 |
(invoke_codegen thy dep false) (gr1, out_ts') |
|
11537 | 376 |
val (gr3, rest) = compile_prems [] vs' (fst nvs) gr2 ps'; |
377 |
in |
|
378 |
(gr3, compile_match (snd nvs) [] out_ps |
|
379 |
(Pretty.block [Pretty.str "?? ", side_p, |
|
380 |
Pretty.str " :->", Pretty.brk 1, rest]) |
|
381 |
(Pretty.str "Seq.empty")) |
|
382 |
end) |
|
383 |
end; |
|
384 |
||
385 |
val (gr', prem_p) = compile_prems in_ts' [] all_vs' gr ps; |
|
386 |
in |
|
387 |
(gr', Pretty.block [Pretty.str "Seq.single inp :->", Pretty.brk 1, prem_p]) |
|
388 |
end; |
|
389 |
||
390 |
fun compile_pred thy gr dep prfx all_vs arg_vs modes s cls mode = |
|
391 |
let val (gr', cl_ps) = foldl_map (fn (gr, cl) => |
|
392 |
compile_clause thy gr dep all_vs arg_vs modes mode cl) (gr, cls) |
|
393 |
in |
|
394 |
((gr', "and "), Pretty.block |
|
395 |
([Pretty.block (separate (Pretty.brk 1) |
|
396 |
(Pretty.str (prfx ^ modename thy s mode) :: map Pretty.str arg_vs) @ |
|
397 |
[Pretty.str " inp ="]), |
|
398 |
Pretty.brk 1] @ |
|
399 |
flat (separate [Pretty.str " ++", Pretty.brk 1] (map single cl_ps)))) |
|
400 |
end; |
|
401 |
||
402 |
fun compile_preds thy gr dep all_vs arg_vs modes preds = |
|
403 |
let val ((gr', _), prs) = foldl_map (fn ((gr, prfx), (s, cls)) => |
|
404 |
foldl_map (fn ((gr', prfx'), mode) => |
|
405 |
compile_pred thy gr' dep prfx' all_vs arg_vs modes s cls mode) |
|
406 |
((gr, prfx), the (assoc (modes, s)))) ((gr, "fun "), preds) |
|
407 |
in |
|
408 |
(gr', space_implode "\n\n" (map Pretty.string_of (flat prs)) ^ ";\n\n") |
|
409 |
end; |
|
410 |
||
411 |
(**** processing of introduction rules ****) |
|
412 |
||
12557 | 413 |
exception Modes of |
414 |
(string * (int list option list * int list) list) list * |
|
415 |
(string * (int list list option list * int list list)) list; |
|
416 |
||
417 |
fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip |
|
418 |
(map ((fn (Some (Modes x), _) => x | _ => ([], [])) o Graph.get_node gr) |
|
419 |
(Graph.all_preds gr [dep])))); |
|
420 |
||
421 |
fun string_of_mode (iss, is) = space_implode " -> " (map |
|
422 |
(fn None => "X" |
|
423 |
| Some js => enclose "[" "]" (commas (map string_of_int js))) |
|
424 |
(iss @ [Some is])); |
|
11537 | 425 |
|
426 |
fun print_modes modes = message ("Inferred modes:\n" ^ |
|
427 |
space_implode "\n" (map (fn (s, ms) => s ^ ": " ^ commas (map |
|
428 |
string_of_mode ms)) modes)); |
|
429 |
||
430 |
fun print_factors factors = message ("Factors:\n" ^ |
|
12557 | 431 |
space_implode "\n" (map (fn (s, (fs, f)) => s ^ ": " ^ |
432 |
space_implode " -> " (map |
|
433 |
(fn None => "X" | Some f' => string_of_factors [] f') |
|
434 |
(fs @ [Some f]))) factors)); |
|
11537 | 435 |
|
12557 | 436 |
fun mk_extra_defs thy gr dep names ts = |
437 |
foldl (fn (gr, name) => |
|
438 |
if name mem names then gr |
|
439 |
else (case get_clauses thy name of |
|
440 |
None => gr |
|
441 |
| Some (names, intrs) => |
|
442 |
mk_ind_def thy gr dep names intrs)) |
|
443 |
(gr, foldr add_term_consts (ts, [])) |
|
444 |
||
445 |
and mk_ind_def thy gr dep names intrs = |
|
11537 | 446 |
let val ids = map (mk_const_id (sign_of thy)) names |
447 |
in Graph.add_edge (hd ids, dep) gr handle Graph.UNDEF _ => |
|
448 |
let |
|
14250
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
449 |
fun dest_prem factors (_ $ (p as (Const ("op :", _) $ t $ u))) = |
12557 | 450 |
(case head_of u of |
14250
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
451 |
Const (name, _) => (case assoc (factors, name) of |
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
452 |
None => Sidecond p |
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
453 |
| Some f => Prem (split_prod [] f t, u)) |
12557 | 454 |
| Var ((name, _), _) => Prem (split_prod [] |
455 |
(the (assoc (factors, name))) t, u)) |
|
456 |
| dest_prem factors (_ $ ((eq as Const ("op =", _)) $ t $ u)) = |
|
457 |
Prem ([t, u], eq) |
|
458 |
| dest_prem factors (_ $ t) = Sidecond t; |
|
11537 | 459 |
|
12557 | 460 |
fun add_clause factors (clauses, intr) = |
11537 | 461 |
let |
462 |
val _ $ (_ $ t $ u) = Logic.strip_imp_concl intr; |
|
12557 | 463 |
val Const (name, _) = head_of u; |
464 |
val prems = map (dest_prem factors) (Logic.strip_imp_prems intr); |
|
11537 | 465 |
in |
466 |
(overwrite (clauses, (name, if_none (assoc (clauses, name)) [] @ |
|
12557 | 467 |
[(split_prod [] (the (assoc (factors, name))) t, prems)]))) |
11537 | 468 |
end; |
469 |
||
12557 | 470 |
fun add_prod_factors extra_fs (fs, _ $ (Const ("op :", _) $ t $ u)) = |
14250
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
471 |
(case apsome (get_clauses thy o fst) (try dest_Const (head_of u)) of |
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
472 |
Some None => fs |
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
473 |
| _ => infer_factors (sign_of thy) extra_fs |
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
berghofe
parents:
14195
diff
changeset
|
474 |
(fs, (Some (FVar (prod_factors [] t)), u))) |
12557 | 475 |
| add_prod_factors _ (fs, _) = fs; |
11537 | 476 |
|
477 |
val intrs' = map (rename_term o #prop o rep_thm o standard) intrs; |
|
478 |
val _ $ (_ $ _ $ u) = Logic.strip_imp_concl (hd intrs'); |
|
479 |
val (_, args) = strip_comb u; |
|
480 |
val arg_vs = flat (map term_vs args); |
|
12557 | 481 |
val gr' = mk_extra_defs thy |
482 |
(Graph.add_edge (hd ids, dep) |
|
483 |
(Graph.new_node (hd ids, (None, "")) gr)) (hd ids) names intrs'; |
|
14163
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
berghofe
parents:
14162
diff
changeset
|
484 |
val (extra_modes, extra_factors) = lookup_modes gr' (hd ids); |
12557 | 485 |
val fs = map (apsnd dest_factors) |
486 |
(foldl (add_prod_factors extra_factors) ([], flat (map (fn t => |
|
487 |
Logic.strip_imp_concl t :: Logic.strip_imp_prems t) intrs'))); |
|
488 |
val _ = (case map fst fs \\ names \\ arg_vs of |
|
489 |
[] => () |
|
490 |
| xs => error ("Non-inductive sets: " ^ commas_quote xs)); |
|
491 |
val factors = mapfilter (fn (name, f) => |
|
492 |
if name mem arg_vs then None |
|
493 |
else Some (name, (map (curry assoc fs) arg_vs, f))) fs; |
|
494 |
val clauses = |
|
495 |
foldl (add_clause (fs @ map (apsnd snd) extra_factors)) ([], intrs'); |
|
496 |
val modes = infer_modes thy extra_modes factors arg_vs clauses; |
|
497 |
val _ = print_factors factors; |
|
11537 | 498 |
val _ = print_modes modes; |
499 |
val (gr'', s) = compile_preds thy gr' (hd ids) (terms_vs intrs') arg_vs |
|
500 |
(modes @ extra_modes) clauses; |
|
501 |
in |
|
502 |
(Graph.map_node (hd ids) (K (Some (Modes (modes, factors)), s)) gr'') |
|
503 |
end |
|
504 |
end; |
|
505 |
||
12557 | 506 |
fun mk_ind_call thy gr dep t u is_query = (case head_of u of |
13038 | 507 |
Const (s, T) => (case (get_clauses thy s, get_assoc_code thy s T) of |
508 |
(None, _) => None |
|
509 |
| (Some (names, intrs), None) => |
|
11537 | 510 |
let |
12565 | 511 |
fun mk_mode (((ts, mode), i), Const ("dummy_pattern", _)) = |
512 |
((ts, mode), i+1) |
|
11537 | 513 |
| mk_mode (((ts, mode), i), t) = ((ts @ [t], mode @ [i]), i+1); |
514 |
||
12557 | 515 |
val gr1 = mk_extra_defs thy |
516 |
(mk_ind_def thy gr dep names intrs) dep names [u]; |
|
517 |
val (modes, factors) = lookup_modes gr1 dep; |
|
518 |
val ts = split_prod [] (snd (the (assoc (factors, s)))) t; |
|
519 |
val (ts', is) = if is_query then |
|
11537 | 520 |
fst (foldl mk_mode ((([], []), 1), ts)) |
521 |
else (ts, 1 upto length ts); |
|
12557 | 522 |
val mode = (case find_first (fn Mode ((_, js), _) => is=js) |
523 |
(modes_of modes u) of |
|
524 |
None => error ("No such mode for " ^ s ^ ": " ^ |
|
525 |
string_of_mode ([], is)) |
|
526 |
| mode => mode); |
|
12453 | 527 |
val (gr2, in_ps) = foldl_map |
528 |
(invoke_codegen thy dep false) (gr1, ts'); |
|
12557 | 529 |
val (gr3, ps) = compile_expr thy dep false (gr2, (mode, u)) |
11537 | 530 |
in |
12557 | 531 |
Some (gr3, Pretty.block |
532 |
(ps @ [Pretty.brk 1, mk_tuple in_ps])) |
|
13038 | 533 |
end |
534 |
| _ => None) |
|
11537 | 535 |
| _ => None); |
536 |
||
537 |
fun inductive_codegen thy gr dep brack (Const ("op :", _) $ t $ u) = |
|
12565 | 538 |
((case mk_ind_call thy gr dep (Term.no_dummy_patterns t) u false of |
11537 | 539 |
None => None |
540 |
| Some (gr', call_p) => Some (gr', (if brack then parens else I) |
|
12453 | 541 |
(Pretty.block [Pretty.str "?! (", call_p, Pretty.str ")"]))) |
12565 | 542 |
handle TERM _ => mk_ind_call thy gr dep t u true) |
11537 | 543 |
| inductive_codegen thy gr dep brack _ = None; |
544 |
||
12557 | 545 |
val setup = |
546 |
[add_codegen "inductive" inductive_codegen, |
|
547 |
CodegenData.init, |
|
14195 | 548 |
add_attribute "ind" (Scan.succeed add)]; |
11537 | 549 |
|
550 |
end; |
|
12453 | 551 |
|
552 |
||
553 |
(**** combinators for code generated from inductive predicates ****) |
|
554 |
||
555 |
infix 5 :->; |
|
556 |
infix 3 ++; |
|
557 |
||
558 |
fun s :-> f = Seq.flat (Seq.map f s); |
|
559 |
||
560 |
fun s1 ++ s2 = Seq.append (s1, s2); |
|
561 |
||
562 |
fun ?? b = if b then Seq.single () else Seq.empty; |
|
563 |
||
564 |
fun ?! s = is_some (Seq.pull s); |
|
565 |
||
14859 | 566 |
fun op_61__1 x = Seq.single x; |
12453 | 567 |
|
14859 | 568 |
val op_61__2 = op_61__1; |
12557 | 569 |
|
14859 | 570 |
fun op_61__1_2 (x, y) = ?? (x = y); |