| author | haftmann | 
| Tue, 28 Jun 2005 11:58:56 +0200 | |
| changeset 16576 | 9ce0be075e6a | 
| parent 16424 | 18a07ad8fea8 | 
| child 16645 | a152d6b21c31 | 
| 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  | 
|
10  | 
val add: theory attribute  | 
|
11  | 
val setup: (theory -> theory) list  | 
|
12  | 
end;  | 
|
13  | 
||
14  | 
structure RecfunCodegen : RECFUN_CODEGEN =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
open Codegen;  | 
|
18  | 
||
| 16424 | 19  | 
structure CodegenData = TheoryDataFun  | 
20  | 
(struct  | 
|
| 12447 | 21  | 
val name = "HOL/recfun_codegen";  | 
22  | 
type T = thm list Symtab.table;  | 
|
23  | 
val empty = Symtab.empty;  | 
|
24  | 
val copy = I;  | 
|
| 16424 | 25  | 
val extend = I;  | 
26  | 
fun merge _ = Symtab.merge_multi' Drule.eq_thm_prop;  | 
|
| 12447 | 27  | 
fun print _ _ = ();  | 
| 16424 | 28  | 
end);  | 
| 12447 | 29  | 
|
30  | 
||
31  | 
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
 | 
32  | 
val lhs_of = fst o dest_eqn o prop_of;  | 
| 12447 | 33  | 
val const_of = dest_Const o head_of o fst o dest_eqn;  | 
34  | 
||
35  | 
fun warn thm = warning ("RecfunCodegen: Not a proper equation:\n" ^
 | 
|
36  | 
string_of_thm thm);  | 
|
37  | 
||
38  | 
fun add (p as (thy, thm)) =  | 
|
39  | 
let  | 
|
| 12556 | 40  | 
val tab = CodegenData.get thy;  | 
| 12447 | 41  | 
val (s, _) = const_of (prop_of thm);  | 
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
42  | 
in  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
43  | 
if Pattern.pattern (lhs_of thm) then  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
44  | 
(CodegenData.put (Symtab.update ((s,  | 
| 15570 | 45  | 
thm :: getOpt (Symtab.lookup (tab, s), [])), tab)) thy, thm)  | 
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
46  | 
else (warn thm; p)  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
47  | 
end handle TERM _ => (warn thm; p);  | 
| 12447 | 48  | 
|
| 14196 | 49  | 
fun del (p as (thy, thm)) =  | 
50  | 
let  | 
|
51  | 
val tab = CodegenData.get thy;  | 
|
52  | 
val (s, _) = const_of (prop_of thm);  | 
|
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
53  | 
in case Symtab.lookup (tab, s) of  | 
| 15531 | 54  | 
NONE => p  | 
55  | 
| SOME thms => (CodegenData.put (Symtab.update ((s,  | 
|
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
56  | 
gen_rem eq_thm (thms, thm)), tab)) thy, thm)  | 
| 14196 | 57  | 
end handle TERM _ => (warn thm; p);  | 
58  | 
||
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
59  | 
fun del_redundant thy eqs [] = eqs  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
60  | 
| del_redundant thy eqs (eq :: eqs') =  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
61  | 
let  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
62  | 
val tsig = Sign.tsig_of (sign_of thy);  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
63  | 
val matches = curry  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
64  | 
(Pattern.matches tsig o pairself lhs_of)  | 
| 
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
65  | 
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
 | 
66  | 
|
| 12447 | 67  | 
fun get_equations thy (s, T) =  | 
| 12556 | 68  | 
(case Symtab.lookup (CodegenData.get thy, s) of  | 
| 15531 | 69  | 
NONE => []  | 
70  | 
| SOME thms => preprocess thy (del_redundant thy []  | 
|
| 15570 | 71  | 
(List.filter (fn thm => is_instance thy T  | 
| 
15321
 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 
berghofe 
parents: 
15257 
diff
changeset
 | 
72  | 
(snd (const_of (prop_of thm)))) thms)));  | 
| 12447 | 73  | 
|
74  | 
fun mk_poly_id thy (s, T) = mk_const_id (sign_of thy) s ^  | 
|
75  | 
(case get_defn thy s T of  | 
|
| 15531 | 76  | 
SOME (_, SOME i) => "_def" ^ string_of_int i  | 
| 12447 | 77  | 
| _ => "");  | 
78  | 
||
79  | 
exception EQN of string * typ * string;  | 
|
80  | 
||
81  | 
fun cycle g (xs, x) =  | 
|
82  | 
if x mem xs then xs  | 
|
| 15570 | 83  | 
else Library.foldl (cycle g) (x :: xs, List.concat (Graph.find_paths g (x, x)));  | 
| 12447 | 84  | 
|
85  | 
fun add_rec_funs thy gr dep eqs =  | 
|
86  | 
let  | 
|
87  | 
fun dest_eq t =  | 
|
88  | 
(mk_poly_id thy (const_of t), dest_eqn (rename_term t));  | 
|
89  | 
val eqs' = map dest_eq eqs;  | 
|
90  | 
val (dname, _) :: _ = eqs';  | 
|
91  | 
val (s, T) = const_of (hd eqs);  | 
|
92  | 
||
93  | 
fun mk_fundef fname prfx gr [] = (gr, [])  | 
|
94  | 
| mk_fundef fname prfx gr ((fname', (lhs, rhs)) :: xs) =  | 
|
95  | 
let  | 
|
96  | 
val (gr1, pl) = invoke_codegen thy dname false (gr, lhs);  | 
|
97  | 
val (gr2, pr) = invoke_codegen thy dname false (gr1, rhs);  | 
|
98  | 
val (gr3, rest) = mk_fundef fname' "and " gr2 xs  | 
|
99  | 
in  | 
|
100  | 
(gr3, Pretty.blk (4, [Pretty.str (if fname = fname' then " | " else prfx),  | 
|
101  | 
pl, Pretty.str " =", Pretty.brk 1, pr]) :: rest)  | 
|
102  | 
end;  | 
|
103  | 
||
104  | 
fun put_code fundef = Graph.map_node dname  | 
|
| 15531 | 105  | 
      (K (SOME (EQN ("", dummyT, dname)), Pretty.string_of (Pretty.blk (0,
 | 
| 12447 | 106  | 
separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n"));  | 
107  | 
||
108  | 
in  | 
|
109  | 
(case try (Graph.get_node gr) dname of  | 
|
| 15531 | 110  | 
NONE =>  | 
| 12447 | 111  | 
let  | 
112  | 
val gr1 = Graph.add_edge (dname, dep)  | 
|
| 15531 | 113  | 
(Graph.new_node (dname, (SOME (EQN (s, T, "")), "")) gr);  | 
| 12447 | 114  | 
val (gr2, fundef) = mk_fundef "" "fun " gr1 eqs';  | 
115  | 
val xs = cycle gr2 ([], dname);  | 
|
116  | 
val cs = map (fn x => case Graph.get_node gr2 x of  | 
|
| 15531 | 117  | 
(SOME (EQN (s, T, _)), _) => (s, T)  | 
| 12447 | 118  | 
             | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
 | 
119  | 
implode (separate ", " xs))) xs  | 
|
120  | 
in (case xs of  | 
|
121  | 
[_] => put_code fundef gr2  | 
|
122  | 
| _ =>  | 
|
123  | 
if not (dep mem xs) then  | 
|
124  | 
let  | 
|
125  | 
val eqs'' = map (dest_eq o prop_of)  | 
|
| 15570 | 126  | 
(List.concat (map (get_equations thy) cs));  | 
| 12447 | 127  | 
val (gr3, fundef') = mk_fundef "" "fun "  | 
128  | 
(Graph.add_edge (dname, dep)  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
129  | 
(foldr (uncurry Graph.new_node) (Graph.del_nodes xs gr2)  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
130  | 
(map (fn k =>  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
131  | 
                         (k, (SOME (EQN ("", dummyT, dname)), ""))) xs))) eqs''
 | 
| 12447 | 132  | 
in put_code fundef' gr3 end  | 
133  | 
else gr2)  | 
|
134  | 
end  | 
|
| 15531 | 135  | 
| SOME (SOME (EQN (_, _, s)), _) =>  | 
| 12447 | 136  | 
if s = "" then  | 
137  | 
if dname = dep then gr else Graph.add_edge (dname, dep) gr  | 
|
138  | 
else if s = dep then gr else Graph.add_edge (s, dep) gr)  | 
|
139  | 
end;  | 
|
140  | 
||
141  | 
fun recfun_codegen thy gr dep brack t = (case strip_comb t of  | 
|
| 15257 | 142  | 
(Const (p as (s, T)), ts) => (case (get_equations thy p, get_assoc_code thy s T) of  | 
| 15531 | 143  | 
([], _) => NONE  | 
144  | 
| (_, SOME _) => NONE  | 
|
145  | 
| (eqns, NONE) =>  | 
|
| 12447 | 146  | 
let val (gr', ps) = foldl_map (invoke_codegen thy dep true) (gr, ts)  | 
147  | 
in  | 
|
| 15531 | 148  | 
SOME (add_rec_funs thy gr' dep (map prop_of eqns),  | 
| 12447 | 149  | 
mk_app brack (Pretty.str (mk_poly_id thy p)) ps)  | 
150  | 
end)  | 
|
| 15531 | 151  | 
| _ => NONE);  | 
| 12447 | 152  | 
|
153  | 
||
154  | 
val setup =  | 
|
| 12556 | 155  | 
[CodegenData.init,  | 
| 12447 | 156  | 
add_codegen "recfun" recfun_codegen,  | 
| 14196 | 157  | 
add_attribute "" (Args.del |-- Scan.succeed del || Scan.succeed add)];  | 
| 12447 | 158  | 
|
159  | 
end;  |