| author | bulwahn | 
| Wed, 23 Mar 2011 08:50:31 +0100 | |
| changeset 42088 | 8d00484551fe | 
| parent 41448 | 72ba43b47c7f | 
| child 42411 | ff997038e8eb | 
| permissions | -rw-r--r-- | 
| 24584 | 1  | 
(* Title: HOL/Tools/recfun_codegen.ML  | 
| 12447 | 2  | 
Author: Stefan Berghofer, TU Muenchen  | 
3  | 
||
4  | 
Code generator for recursive functions.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature RECFUN_CODEGEN =  | 
|
8  | 
sig  | 
|
| 18708 | 9  | 
val setup: theory -> theory  | 
| 12447 | 10  | 
end;  | 
11  | 
||
12  | 
structure RecfunCodegen : RECFUN_CODEGEN =  | 
|
13  | 
struct  | 
|
14  | 
||
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
15  | 
val const_of = dest_Const o head_of o fst o Logic.dest_equals;  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
16  | 
|
| 33522 | 17  | 
structure ModuleData = Theory_Data  | 
| 22846 | 18  | 
(  | 
| 28522 | 19  | 
type T = string Symtab.table;  | 
| 12447 | 20  | 
val empty = Symtab.empty;  | 
| 16424 | 21  | 
val extend = I;  | 
| 33522 | 22  | 
fun merge data = Symtab.merge (K true) data;  | 
| 22846 | 23  | 
);  | 
| 12447 | 24  | 
|
| 
31998
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
25  | 
fun add_thm_target module_name thm thy =  | 
| 
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
26  | 
let  | 
| 
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
27  | 
val (thm', _) = Code.mk_eqn thy (thm, true)  | 
| 
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
28  | 
in  | 
| 
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
29  | 
thy  | 
| 
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
30  | 
|> ModuleData.map (Symtab.update (fst (Code.const_typ_eqn thy thm'), module_name))  | 
| 
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
31  | 
end;  | 
| 
24624
 
b8383b1bbae3
distinction between regular and default code theorems
 
haftmann 
parents: 
24584 
diff
changeset
 | 
32  | 
|
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
33  | 
fun avoid_value thy [thm] =  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
34  | 
let val (_, T) = Code.const_typ_eqn thy thm  | 
| 40844 | 35  | 
in  | 
36  | 
if null (Term.add_tvarsT T []) orelse null (binder_types T)  | 
|
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
37  | 
then [thm]  | 
| 34895 | 38  | 
else [Code.expand_eta thy 1 thm]  | 
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
39  | 
end  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
40  | 
| avoid_value thy thms = thms;  | 
| 28522 | 41  | 
|
| 
38864
 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 
haftmann 
parents: 
36692 
diff
changeset
 | 
42  | 
fun get_equations thy defs (raw_c, T) = if raw_c = @{const_name HOL.eq} then ([], "") else
 | 
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
43  | 
let  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
44  | 
val c = AxClass.unoverload_const thy (raw_c, T);  | 
| 
34893
 
ecdc526af73a
function transformer preprocessor applies to both code generators
 
haftmann 
parents: 
34891 
diff
changeset
 | 
45  | 
val raw_thms = Code.get_cert thy (Code_Preproc.preprocess_functrans thy) c  | 
| 35225 | 46  | 
|> Code.bare_thms_of_cert thy  | 
| 
34891
 
99b9a6290446
code certificates as integral part of code generation
 
haftmann 
parents: 
33522 
diff
changeset
 | 
47  | 
|> map (AxClass.overload thy)  | 
| 41448 | 48  | 
|> filter (Codegen.is_instance T o snd o const_of o prop_of);  | 
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
49  | 
val module_name = case Symtab.lookup (ModuleData.get thy) c  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
50  | 
of SOME module_name => module_name  | 
| 41448 | 51  | 
| NONE =>  | 
52  | 
case Codegen.get_defn thy defs c T  | 
|
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
53  | 
of SOME ((_, (thyname, _)), _) => thyname  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
54  | 
| NONE => Codegen.thyname_of_const thy c;  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
55  | 
in if null raw_thms then ([], "") else  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
56  | 
raw_thms  | 
| 41448 | 57  | 
|> Codegen.preprocess thy  | 
| 
32358
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
58  | 
|> avoid_value thy  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
59  | 
|> rpair module_name  | 
| 
 
98c00ee9e786
proper eta expansion in recfun_codegen.ML; no eta expansion at all in code_thingol.ML
 
haftmann 
parents: 
32353 
diff
changeset
 | 
60  | 
end;  | 
| 12447 | 61  | 
|
| 41448 | 62  | 
fun mk_suffix thy defs (s, T) =  | 
63  | 
(case Codegen.get_defn thy defs s T of  | 
|
64  | 
SOME (_, SOME i) => " def" ^ string_of_int i  | 
|
65  | 
| _ => "");  | 
|
| 12447 | 66  | 
|
67  | 
exception EQN of string * typ * string;  | 
|
68  | 
||
| 33244 | 69  | 
fun cycle g x xs =  | 
| 22887 | 70  | 
if member (op =) xs x then xs  | 
| 33244 | 71  | 
else fold (cycle g) (flat (Graph.all_paths (fst g) (x, x))) (x :: xs);  | 
| 12447 | 72  | 
|
| 28535 | 73  | 
fun add_rec_funs thy defs dep module eqs gr =  | 
| 12447 | 74  | 
let  | 
| 16645 | 75  | 
fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t),  | 
| 41448 | 76  | 
Logic.dest_equals (Codegen.rename_term t));  | 
| 12447 | 77  | 
val eqs' = map dest_eq eqs;  | 
78  | 
val (dname, _) :: _ = eqs';  | 
|
79  | 
val (s, T) = const_of (hd eqs);  | 
|
80  | 
||
| 28535 | 81  | 
fun mk_fundef module fname first [] gr = ([], gr)  | 
82  | 
| mk_fundef module fname first ((fname' : string, (lhs, rhs)) :: xs) gr =  | 
|
| 12447 | 83  | 
let  | 
| 41448 | 84  | 
val (pl, gr1) = Codegen.invoke_codegen thy defs dname module false lhs gr;  | 
85  | 
val (pr, gr2) = Codegen.invoke_codegen thy defs dname module false rhs gr1;  | 
|
| 28535 | 86  | 
val (rest, gr3) = mk_fundef module fname' false xs gr2 ;  | 
| 41448 | 87  | 
val (ty, gr4) = Codegen.invoke_tycodegen thy defs dname module false T gr3;  | 
| 28535 | 88  | 
val num_args = (length o snd o strip_comb) lhs;  | 
89  | 
val prfx = if fname = fname' then " |"  | 
|
90  | 
else if not first then "and"  | 
|
91  | 
else if num_args = 0 then "val"  | 
|
92  | 
else "fun";  | 
|
| 41448 | 93  | 
val pl' = Pretty.breaks (Codegen.str prfx  | 
94  | 
:: (if num_args = 0 then [pl, Codegen.str ":", ty] else [pl]));  | 
|
| 12447 | 95  | 
in  | 
| 28535 | 96  | 
(Pretty.blk (4, pl'  | 
| 41448 | 97  | 
@ [Codegen.str " =", Pretty.brk 1, pr]) :: rest, gr4)  | 
| 12447 | 98  | 
end;  | 
99  | 
||
| 41448 | 100  | 
fun put_code module fundef = Codegen.map_node dname  | 
101  | 
      (K (SOME (EQN ("", dummyT, dname)), module, Codegen.string_of (Pretty.blk (0,
 | 
|
102  | 
separate Pretty.fbrk fundef @ [Codegen.str ";"])) ^ "\n\n"));  | 
|
| 12447 | 103  | 
|
104  | 
in  | 
|
| 41448 | 105  | 
(case try (Codegen.get_node gr) dname of  | 
| 15531 | 106  | 
NONE =>  | 
| 12447 | 107  | 
let  | 
| 41448 | 108  | 
val gr1 = Codegen.add_edge (dname, dep)  | 
109  | 
(Codegen.new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr);  | 
|
| 28535 | 110  | 
val (fundef, gr2) = mk_fundef module "" true eqs' gr1 ;  | 
| 33244 | 111  | 
val xs = cycle gr2 dname [];  | 
| 41448 | 112  | 
val cs = map (fn x =>  | 
113  | 
case Codegen.get_node gr2 x of  | 
|
| 16645 | 114  | 
(SOME (EQN (s, T, _)), _, _) => (s, T)  | 
| 12447 | 115  | 
             | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
 | 
116  | 
implode (separate ", " xs))) xs  | 
|
| 41448 | 117  | 
in  | 
118  | 
(case xs of  | 
|
| 28535 | 119  | 
[_] => (module, put_code module fundef gr2)  | 
| 12447 | 120  | 
| _ =>  | 
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
35225 
diff
changeset
 | 
121  | 
if not (member (op =) xs dep) then  | 
| 12447 | 122  | 
let  | 
| 16645 | 123  | 
val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs;  | 
| 41448 | 124  | 
val module' = Codegen.if_library thyname module;  | 
| 32952 | 125  | 
val eqs'' = map (dest_eq o prop_of) (maps fst thmss);  | 
| 28535 | 126  | 
val (fundef', gr3) = mk_fundef module' "" true eqs''  | 
| 41448 | 127  | 
(Codegen.add_edge (dname, dep)  | 
128  | 
(List.foldr (uncurry Codegen.new_node) (Codegen.del_nodes xs gr2)  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
129  | 
(map (fn k =>  | 
| 28535 | 130  | 
                         (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs)))
 | 
131  | 
in (module', put_code module' fundef' gr3) end  | 
|
132  | 
else (module, gr2))  | 
|
| 12447 | 133  | 
end  | 
| 17144 | 134  | 
| SOME (SOME (EQN (_, _, s)), module', _) =>  | 
| 28535 | 135  | 
(module', if s = "" then  | 
| 41448 | 136  | 
if dname = dep then gr else Codegen.add_edge (dname, dep) gr  | 
137  | 
else if s = dep then gr else Codegen.add_edge (s, dep) gr))  | 
|
| 12447 | 138  | 
end;  | 
139  | 
||
| 41448 | 140  | 
fun recfun_codegen thy defs dep module brack t gr =  | 
141  | 
(case strip_comb t of  | 
|
142  | 
(Const (p as (s, T)), ts) =>  | 
|
143  | 
(case (get_equations thy defs p, Codegen.get_assoc_code thy (s, T)) of  | 
|
| 16645 | 144  | 
(([], _), _) => NONE  | 
| 15531 | 145  | 
| (_, SOME _) => NONE  | 
| 17144 | 146  | 
| ((eqns, thyname), NONE) =>  | 
| 16645 | 147  | 
let  | 
| 41448 | 148  | 
val module' = Codegen.if_library thyname module;  | 
| 28535 | 149  | 
val (ps, gr') = fold_map  | 
| 41448 | 150  | 
(Codegen.invoke_codegen thy defs dep module true) ts gr;  | 
| 17144 | 151  | 
val suffix = mk_suffix thy defs p;  | 
| 28535 | 152  | 
val (module'', gr'') =  | 
153  | 
add_rec_funs thy defs dep module' (map prop_of eqns) gr';  | 
|
| 41448 | 154  | 
val (fname, gr''') = Codegen.mk_const_id module'' (s ^ suffix) gr''  | 
| 12447 | 155  | 
in  | 
| 41448 | 156  | 
SOME (Codegen.mk_app brack (Codegen.str (Codegen.mk_qual_id module fname)) ps, gr''')  | 
| 12447 | 157  | 
end)  | 
| 15531 | 158  | 
| _ => NONE);  | 
| 12447 | 159  | 
|
| 
31998
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
160  | 
val setup =  | 
| 41448 | 161  | 
Codegen.add_codegen "recfun" recfun_codegen  | 
| 
31998
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31962 
diff
changeset
 | 
162  | 
#> Code.set_code_target_attr add_thm_target;  | 
| 12447 | 163  | 
|
164  | 
end;  |