| author | bulwahn |
| Tue, 04 Aug 2009 08:34:56 +0200 | |
| changeset 32318 | bca7fd849829 |
| parent 31998 | 2c7a24f74db9 |
| child 32345 | 4da4fa060bb6 |
| 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 |
||
15 |
open Codegen; |
|
16 |
||
| 28522 | 17 |
structure ModuleData = TheoryDataFun |
| 22846 | 18 |
( |
| 28522 | 19 |
type T = string Symtab.table; |
| 12447 | 20 |
val empty = Symtab.empty; |
21 |
val copy = I; |
|
| 16424 | 22 |
val extend = I; |
| 28522 | 23 |
fun merge _ = Symtab.merge (K true); |
| 22846 | 24 |
); |
| 12447 | 25 |
|
|
31998
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
26 |
fun add_thm_target module_name thm thy = |
|
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
27 |
let |
|
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
28 |
val (thm', _) = Code.mk_eqn thy (thm, true) |
|
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
29 |
in |
|
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
30 |
thy |
|
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
31 |
|> 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
|
32 |
end; |
|
24624
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24584
diff
changeset
|
33 |
|
| 28522 | 34 |
fun expand_eta thy [] = [] |
35 |
| expand_eta thy (thms as thm :: _) = |
|
36 |
let |
|
| 31957 | 37 |
val (_, ty) = Code.const_typ_eqn thy thm; |
| 29272 | 38 |
in if null (Term.add_tvarsT ty []) orelse (null o fst o strip_type) ty |
| 28522 | 39 |
then thms |
| 31156 | 40 |
else map (Code.expand_eta thy 1) thms |
| 28522 | 41 |
end; |
| 14196 | 42 |
|
| 28522 | 43 |
fun retrieve_equations thy (c, T) = if c = @{const_name "op ="} then NONE else
|
44 |
let |
|
45 |
val c' = AxClass.unoverload_const thy (c, T); |
|
46 |
val opt_name = Symtab.lookup (ModuleData.get thy) c'; |
|
|
31125
80218ee73167
transferred code generator preprocessor into separate module
haftmann
parents:
31090
diff
changeset
|
47 |
val thms = Code.these_eqns thy c' |
| 28522 | 48 |
|> map_filter (fn (thm, linear) => if linear then SOME thm else NONE) |
49 |
|> expand_eta thy |
|
| 31156 | 50 |
|> Code.norm_varnames thy |
| 28522 | 51 |
|> map (rpair opt_name) |
52 |
in if null thms then NONE else SOME thms end; |
|
53 |
||
|
31954
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
haftmann
parents:
31156
diff
changeset
|
54 |
val dest_eqn = Logic.dest_equals; |
| 28522 | 55 |
val const_of = dest_Const o head_of o fst o dest_eqn; |
|
15321
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
berghofe
parents:
15257
diff
changeset
|
56 |
|
| 16645 | 57 |
fun get_equations thy defs (s, T) = |
| 28522 | 58 |
(case retrieve_equations thy (s, T) of |
| 16645 | 59 |
NONE => ([], "") |
60 |
| SOME thms => |
|
| 28522 | 61 |
let val thms' = filter (fn (thm, _) => is_instance T |
62 |
(snd (const_of (prop_of thm)))) thms |
|
| 16645 | 63 |
in if null thms' then ([], "") |
64 |
else (preprocess thy (map fst thms'), |
|
65 |
case snd (snd (split_last thms')) of |
|
66 |
NONE => (case get_defn thy defs s T of |
|
|
27398
768da1da59d6
simplified retrieval of theory names of consts and types
haftmann
parents:
26975
diff
changeset
|
67 |
NONE => Codegen.thyname_of_const thy s |
| 16645 | 68 |
| SOME ((_, (thyname, _)), _) => thyname) |
69 |
| SOME thyname => thyname) |
|
70 |
end); |
|
| 12447 | 71 |
|
| 16645 | 72 |
fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of |
| 17144 | 73 |
SOME (_, SOME i) => " def" ^ string_of_int i | _ => ""); |
| 12447 | 74 |
|
75 |
exception EQN of string * typ * string; |
|
76 |
||
| 22887 | 77 |
fun cycle g (xs, x : string) = |
78 |
if member (op =) xs x then xs |
|
79 |
else Library.foldl (cycle g) (x :: xs, flat (Graph.all_paths (fst g) (x, x))); |
|
| 12447 | 80 |
|
| 28535 | 81 |
fun add_rec_funs thy defs dep module eqs gr = |
| 12447 | 82 |
let |
| 16645 | 83 |
fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t), |
84 |
dest_eqn (rename_term t)); |
|
| 12447 | 85 |
val eqs' = map dest_eq eqs; |
86 |
val (dname, _) :: _ = eqs'; |
|
87 |
val (s, T) = const_of (hd eqs); |
|
88 |
||
| 28535 | 89 |
fun mk_fundef module fname first [] gr = ([], gr) |
90 |
| mk_fundef module fname first ((fname' : string, (lhs, rhs)) :: xs) gr = |
|
| 12447 | 91 |
let |
| 28535 | 92 |
val (pl, gr1) = invoke_codegen thy defs dname module false lhs gr; |
93 |
val (pr, gr2) = invoke_codegen thy defs dname module false rhs gr1; |
|
94 |
val (rest, gr3) = mk_fundef module fname' false xs gr2 ; |
|
95 |
val (ty, gr4) = invoke_tycodegen thy defs dname module false T gr3; |
|
96 |
val num_args = (length o snd o strip_comb) lhs; |
|
97 |
val prfx = if fname = fname' then " |" |
|
98 |
else if not first then "and" |
|
99 |
else if num_args = 0 then "val" |
|
100 |
else "fun"; |
|
101 |
val pl' = Pretty.breaks (str prfx |
|
102 |
:: (if num_args = 0 then [pl, str ":", ty] else [pl])); |
|
| 12447 | 103 |
in |
| 28535 | 104 |
(Pretty.blk (4, pl' |
105 |
@ [str " =", Pretty.brk 1, pr]) :: rest, gr4) |
|
| 12447 | 106 |
end; |
107 |
||
| 17144 | 108 |
fun put_code module fundef = map_node dname |
|
26975
103dca19ef2e
Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents:
26928
diff
changeset
|
109 |
(K (SOME (EQN ("", dummyT, dname)), module, string_of (Pretty.blk (0,
|
|
103dca19ef2e
Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents:
26928
diff
changeset
|
110 |
separate Pretty.fbrk fundef @ [str ";"])) ^ "\n\n")); |
| 12447 | 111 |
|
112 |
in |
|
| 17144 | 113 |
(case try (get_node gr) dname of |
| 15531 | 114 |
NONE => |
| 12447 | 115 |
let |
| 17144 | 116 |
val gr1 = add_edge (dname, dep) |
117 |
(new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr); |
|
| 28535 | 118 |
val (fundef, gr2) = mk_fundef module "" true eqs' gr1 ; |
| 12447 | 119 |
val xs = cycle gr2 ([], dname); |
| 17144 | 120 |
val cs = map (fn x => case get_node gr2 x of |
| 16645 | 121 |
(SOME (EQN (s, T, _)), _, _) => (s, T) |
| 12447 | 122 |
| _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
|
123 |
implode (separate ", " xs))) xs |
|
124 |
in (case xs of |
|
| 28535 | 125 |
[_] => (module, put_code module fundef gr2) |
| 12447 | 126 |
| _ => |
127 |
if not (dep mem xs) then |
|
128 |
let |
|
| 16645 | 129 |
val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs; |
| 17144 | 130 |
val module' = if_library thyname module; |
| 16645 | 131 |
val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss)); |
| 28535 | 132 |
val (fundef', gr3) = mk_fundef module' "" true eqs'' |
| 17144 | 133 |
(add_edge (dname, dep) |
| 30190 | 134 |
(List.foldr (uncurry new_node) (del_nodes xs gr2) |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
135 |
(map (fn k => |
| 28535 | 136 |
(k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs)))
|
137 |
in (module', put_code module' fundef' gr3) end |
|
138 |
else (module, gr2)) |
|
| 12447 | 139 |
end |
| 17144 | 140 |
| SOME (SOME (EQN (_, _, s)), module', _) => |
| 28535 | 141 |
(module', if s = "" then |
| 17144 | 142 |
if dname = dep then gr else add_edge (dname, dep) gr |
| 28535 | 143 |
else if s = dep then gr else add_edge (s, dep) gr)) |
| 12447 | 144 |
end; |
145 |
||
| 28535 | 146 |
fun recfun_codegen thy defs dep module brack t gr = (case strip_comb t of |
|
22921
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22887
diff
changeset
|
147 |
(Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy (s, T)) of |
| 16645 | 148 |
(([], _), _) => NONE |
| 15531 | 149 |
| (_, SOME _) => NONE |
| 17144 | 150 |
| ((eqns, thyname), NONE) => |
| 16645 | 151 |
let |
| 17144 | 152 |
val module' = if_library thyname module; |
| 28535 | 153 |
val (ps, gr') = fold_map |
154 |
(invoke_codegen thy defs dep module true) ts gr; |
|
| 17144 | 155 |
val suffix = mk_suffix thy defs p; |
| 28535 | 156 |
val (module'', gr'') = |
157 |
add_rec_funs thy defs dep module' (map prop_of eqns) gr'; |
|
158 |
val (fname, gr''') = mk_const_id module'' (s ^ suffix) gr'' |
|
| 12447 | 159 |
in |
| 28535 | 160 |
SOME (mk_app brack (str (mk_qual_id module fname)) ps, gr''') |
| 12447 | 161 |
end) |
| 15531 | 162 |
| _ => NONE); |
| 12447 | 163 |
|
|
31998
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
164 |
val setup = |
| 24219 | 165 |
add_codegen "recfun" recfun_codegen |
|
31998
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31962
diff
changeset
|
166 |
#> Code.set_code_target_attr add_thm_target; |
| 12447 | 167 |
|
168 |
end; |