| author | krauss | 
| Thu, 14 Sep 2006 15:25:23 +0200 | |
| changeset 20535 | b4b3933ec026 | 
| parent 19788 | be3a84d22a58 | 
| child 20597 | 65fe827aa595 | 
| permissions | -rw-r--r-- | 
| 12447 | 1  | 
(* Title: HOL/recfun_codegen.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Stefan Berghofer, TU Muenchen  | 
|
4  | 
||
5  | 
Code generator for recursive functions.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature RECFUN_CODEGEN =  | 
|
9  | 
sig  | 
|
| 18728 | 10  | 
val add: string option -> attribute  | 
11  | 
val del: attribute  | 
|
| 18708 | 12  | 
val setup: theory -> theory  | 
| 12447 | 13  | 
end;  | 
14  | 
||
15  | 
structure RecfunCodegen : RECFUN_CODEGEN =  | 
|
16  | 
struct  | 
|
17  | 
||
18  | 
open Codegen;  | 
|
19  | 
||
| 16424 | 20  | 
structure CodegenData = TheoryDataFun  | 
21  | 
(struct  | 
|
| 12447 | 22  | 
val name = "HOL/recfun_codegen";  | 
| 16645 | 23  | 
type T = (thm * string option) list Symtab.table;  | 
| 12447 | 24  | 
val empty = Symtab.empty;  | 
25  | 
val copy = I;  | 
|
| 16424 | 26  | 
val extend = I;  | 
| 18931 | 27  | 
fun merge _ = Symtab.merge_list (Drule.eq_thm_prop o pairself fst);  | 
| 12447 | 28  | 
fun print _ _ = ();  | 
| 16424 | 29  | 
end);  | 
| 12447 | 30  | 
|
31  | 
||
32  | 
val dest_eqn = HOLogic.dest_eq o HOLogic.dest_Trueprop;  | 
|
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
33  | 
val lhs_of = fst o dest_eqn o prop_of;  | 
| 12447 | 34  | 
val const_of = dest_Const o head_of o fst o dest_eqn;  | 
35  | 
||
36  | 
fun warn thm = warning ("RecfunCodegen: Not a proper equation:\n" ^
 | 
|
37  | 
string_of_thm thm);  | 
|
38  | 
||
| 19344 | 39  | 
fun add optmod = Thm.declaration_attribute (fn thm => Context.map_theory  | 
| 12447 | 40  | 
let  | 
| 19344 | 41  | 
fun add thm =  | 
42  | 
if Pattern.pattern (lhs_of thm) then  | 
|
43  | 
CodegenData.map  | 
|
44  | 
(Symtab.update_list ((fst o const_of o prop_of) thm, (thm, optmod)))  | 
|
45  | 
else tap (fn _ => warn thm)  | 
|
46  | 
handle TERM _ => tap (fn _ => warn thm);  | 
|
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
47  | 
in  | 
| 19344 | 48  | 
add thm  | 
49  | 
#> CodegenTheorems.add_fun thm  | 
|
50  | 
end);  | 
|
| 12447 | 51  | 
|
| 19344 | 52  | 
fun del_thm thm thy =  | 
| 14196 | 53  | 
let  | 
54  | 
val tab = CodegenData.get thy;  | 
|
55  | 
val (s, _) = const_of (prop_of thm);  | 
|
| 17412 | 56  | 
in case Symtab.lookup tab s of  | 
| 18728 | 57  | 
NONE => thy  | 
58  | 
| SOME thms =>  | 
|
59  | 
CodegenData.put (Symtab.update (s, gen_rem (eq_thm o apfst fst) (thms, thm)) tab) thy  | 
|
| 19344 | 60  | 
end handle TERM _ => (warn thm; thy);  | 
61  | 
||
62  | 
val del = Thm.declaration_attribute  | 
|
| 19600 | 63  | 
(fn thm => Context.map_theory (del_thm thm #> CodegenTheorems.del_fun thm))  | 
| 14196 | 64  | 
|
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
65  | 
fun del_redundant thy eqs [] = eqs  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
66  | 
| del_redundant thy eqs (eq :: eqs') =  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
67  | 
let  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
68  | 
val matches = curry  | 
| 17203 | 69  | 
(Pattern.matches thy o pairself (lhs_of o fst))  | 
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
70  | 
in del_redundant thy (eq :: eqs) (filter_out (matches eq) eqs') end;  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
71  | 
|
| 16645 | 72  | 
fun get_equations thy defs (s, T) =  | 
| 17412 | 73  | 
(case Symtab.lookup (CodegenData.get thy) s of  | 
| 16645 | 74  | 
NONE => ([], "")  | 
75  | 
| SOME thms =>  | 
|
76  | 
let val thms' = del_redundant thy []  | 
|
77  | 
(List.filter (fn (thm, _) => is_instance thy T  | 
|
78  | 
(snd (const_of (prop_of thm)))) thms)  | 
|
79  | 
in if null thms' then ([], "")  | 
|
80  | 
else (preprocess thy (map fst thms'),  | 
|
81  | 
case snd (snd (split_last thms')) of  | 
|
82  | 
NONE => (case get_defn thy defs s T of  | 
|
83  | 
NONE => thyname_of_const s thy  | 
|
84  | 
| SOME ((_, (thyname, _)), _) => thyname)  | 
|
85  | 
| SOME thyname => thyname)  | 
|
86  | 
end);  | 
|
| 12447 | 87  | 
|
| 16645 | 88  | 
fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of  | 
| 17144 | 89  | 
SOME (_, SOME i) => " def" ^ string_of_int i | _ => "");  | 
| 12447 | 90  | 
|
91  | 
exception EQN of string * typ * string;  | 
|
92  | 
||
93  | 
fun cycle g (xs, x) =  | 
|
94  | 
if x mem xs then xs  | 
|
| 
19575
 
2d9940cd52d3
replaced Graph.find_paths by Graph.irreducible_paths;
 
wenzelm 
parents: 
19344 
diff
changeset
 | 
95  | 
else Library.foldl (cycle g) (x :: xs, List.concat (Graph.irreducible_paths (fst g) (x, x)));  | 
| 12447 | 96  | 
|
| 17144 | 97  | 
fun add_rec_funs thy defs gr dep eqs module =  | 
| 12447 | 98  | 
let  | 
| 16645 | 99  | 
fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t),  | 
100  | 
dest_eqn (rename_term t));  | 
|
| 12447 | 101  | 
val eqs' = map dest_eq eqs;  | 
102  | 
val (dname, _) :: _ = eqs';  | 
|
103  | 
val (s, T) = const_of (hd eqs);  | 
|
104  | 
||
| 17144 | 105  | 
fun mk_fundef module fname prfx gr [] = (gr, [])  | 
106  | 
| mk_fundef module fname prfx gr ((fname', (lhs, rhs)) :: xs) =  | 
|
| 12447 | 107  | 
let  | 
| 17144 | 108  | 
val (gr1, pl) = invoke_codegen thy defs dname module false (gr, lhs);  | 
109  | 
val (gr2, pr) = invoke_codegen thy defs dname module false (gr1, rhs);  | 
|
110  | 
val (gr3, rest) = mk_fundef module fname' "and " gr2 xs  | 
|
| 12447 | 111  | 
in  | 
112  | 
(gr3, Pretty.blk (4, [Pretty.str (if fname = fname' then " | " else prfx),  | 
|
113  | 
pl, Pretty.str " =", Pretty.brk 1, pr]) :: rest)  | 
|
114  | 
end;  | 
|
115  | 
||
| 17144 | 116  | 
fun put_code module fundef = map_node dname  | 
117  | 
      (K (SOME (EQN ("", dummyT, dname)), module, Pretty.string_of (Pretty.blk (0,
 | 
|
| 12447 | 118  | 
separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n"));  | 
119  | 
||
120  | 
in  | 
|
| 17144 | 121  | 
(case try (get_node gr) dname of  | 
| 15531 | 122  | 
NONE =>  | 
| 12447 | 123  | 
let  | 
| 17144 | 124  | 
val gr1 = add_edge (dname, dep)  | 
125  | 
(new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr);  | 
|
126  | 
val (gr2, fundef) = mk_fundef module "" "fun " gr1 eqs';  | 
|
| 12447 | 127  | 
val xs = cycle gr2 ([], dname);  | 
| 17144 | 128  | 
val cs = map (fn x => case get_node gr2 x of  | 
| 16645 | 129  | 
(SOME (EQN (s, T, _)), _, _) => (s, T)  | 
| 12447 | 130  | 
             | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
 | 
131  | 
implode (separate ", " xs))) xs  | 
|
132  | 
in (case xs of  | 
|
| 17144 | 133  | 
[_] => (put_code module fundef gr2, module)  | 
| 12447 | 134  | 
| _ =>  | 
135  | 
if not (dep mem xs) then  | 
|
136  | 
let  | 
|
| 16645 | 137  | 
val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs;  | 
| 17144 | 138  | 
val module' = if_library thyname module;  | 
| 16645 | 139  | 
val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss));  | 
| 17144 | 140  | 
val (gr3, fundef') = mk_fundef module' "" "fun "  | 
141  | 
(add_edge (dname, dep)  | 
|
142  | 
(foldr (uncurry new_node) (del_nodes xs gr2)  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
143  | 
(map (fn k =>  | 
| 17144 | 144  | 
                         (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs))) eqs''
 | 
145  | 
in (put_code module' fundef' gr3, module') end  | 
|
146  | 
else (gr2, module))  | 
|
| 12447 | 147  | 
end  | 
| 17144 | 148  | 
| SOME (SOME (EQN (_, _, s)), module', _) =>  | 
149  | 
(if s = "" then  | 
|
150  | 
if dname = dep then gr else add_edge (dname, dep) gr  | 
|
151  | 
else if s = dep then gr else add_edge (s, dep) gr,  | 
|
152  | 
module'))  | 
|
| 12447 | 153  | 
end;  | 
154  | 
||
| 17144 | 155  | 
fun recfun_codegen thy defs gr dep module brack t = (case strip_comb t of  | 
| 16645 | 156  | 
(Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy s T) of  | 
157  | 
(([], _), _) => NONE  | 
|
| 15531 | 158  | 
| (_, SOME _) => NONE  | 
| 17144 | 159  | 
| ((eqns, thyname), NONE) =>  | 
| 16645 | 160  | 
let  | 
| 17144 | 161  | 
val module' = if_library thyname module;  | 
| 16645 | 162  | 
val (gr', ps) = foldl_map  | 
| 17144 | 163  | 
(invoke_codegen thy defs dep module true) (gr, ts);  | 
164  | 
val suffix = mk_suffix thy defs p;  | 
|
165  | 
val (gr'', module'') =  | 
|
166  | 
add_rec_funs thy defs gr' dep (map prop_of eqns) module';  | 
|
167  | 
val (gr''', fname) = mk_const_id module'' (s ^ suffix) gr''  | 
|
| 12447 | 168  | 
in  | 
| 17144 | 169  | 
SOME (gr''', mk_app brack (Pretty.str (mk_qual_id module fname)) ps)  | 
| 12447 | 170  | 
end)  | 
| 15531 | 171  | 
| _ => NONE);  | 
| 12447 | 172  | 
|
173  | 
||
| 18708 | 174  | 
val setup =  | 
175  | 
CodegenData.init #>  | 
|
176  | 
add_codegen "recfun" recfun_codegen #>  | 
|
| 18220 | 177  | 
add_attribute ""  | 
| 19344 | 178  | 
(Args.del |-- Scan.succeed del  | 
179  | 
|| Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add);  | 
|
| 12447 | 180  | 
|
181  | 
end;  |