| author | wenzelm | 
| Wed, 04 Jul 2007 16:49:34 +0200 | |
| changeset 23566 | b65692d4adcd | 
| parent 22921 | 475ff421a6a3 | 
| child 24219 | e558fe311376 | 
| 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 | ||
| 20597 | 20 | structure RecCodegenData = TheoryDataFun | 
| 22846 | 21 | ( | 
| 16645 | 22 | type T = (thm * string option) list Symtab.table; | 
| 12447 | 23 | val empty = Symtab.empty; | 
| 24 | val copy = I; | |
| 16424 | 25 | val extend = I; | 
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
21410diff
changeset | 26 | fun merge _ = Symtab.merge_list (Thm.eq_thm_prop o pairself fst); | 
| 22846 | 27 | ); | 
| 12447 | 28 | |
| 29 | ||
| 30 | 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: 
15257diff
changeset | 31 | val lhs_of = fst o dest_eqn o prop_of; | 
| 12447 | 32 | val const_of = dest_Const o head_of o fst o dest_eqn; | 
| 33 | ||
| 34 | fun warn thm = warning ("RecfunCodegen: Not a proper equation:\n" ^
 | |
| 35 | string_of_thm thm); | |
| 36 | ||
| 21410 | 37 | fun add optmod = | 
| 12447 | 38 | let | 
| 19344 | 39 | fun add thm = | 
| 40 | if Pattern.pattern (lhs_of thm) then | |
| 20597 | 41 | RecCodegenData.map | 
| 19344 | 42 | (Symtab.update_list ((fst o const_of o prop_of) thm, (thm, optmod))) | 
| 43 | else tap (fn _ => warn thm) | |
| 44 | handle TERM _ => tap (fn _ => warn thm); | |
| 15321 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 45 | in | 
| 21410 | 46 | Thm.declaration_attribute (fn thm => Context.mapping | 
| 22484 
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
 haftmann parents: 
22360diff
changeset | 47 | (add thm #> CodegenData.add_func false thm) I) | 
| 21410 | 48 | end; | 
| 12447 | 49 | |
| 19344 | 50 | fun del_thm thm thy = | 
| 14196 | 51 | let | 
| 20597 | 52 | val tab = RecCodegenData.get thy; | 
| 14196 | 53 | val (s, _) = const_of (prop_of thm); | 
| 17412 | 54 | in case Symtab.lookup tab s of | 
| 18728 | 55 | NONE => thy | 
| 56 | | SOME thms => | |
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
21410diff
changeset | 57 | RecCodegenData.put (Symtab.update (s, remove (Thm.eq_thm o apsnd fst) thm thms) tab) thy | 
| 19344 | 58 | end handle TERM _ => (warn thm; thy); | 
| 59 | ||
| 60 | val del = Thm.declaration_attribute | |
| 20897 | 61 | (fn thm => Context.mapping (del_thm thm #> CodegenData.del_func thm) I) | 
| 14196 | 62 | |
| 15321 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 63 | fun del_redundant thy eqs [] = eqs | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 64 | | del_redundant thy eqs (eq :: eqs') = | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 65 | let | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 66 | val matches = curry | 
| 17203 | 67 | (Pattern.matches thy o pairself (lhs_of o fst)) | 
| 15321 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 68 | 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: 
15257diff
changeset | 69 | |
| 16645 | 70 | fun get_equations thy defs (s, T) = | 
| 20597 | 71 | (case Symtab.lookup (RecCodegenData.get thy) s of | 
| 16645 | 72 | NONE => ([], "") | 
| 73 | | SOME thms => | |
| 74 | let val thms' = del_redundant thy [] | |
| 22887 | 75 | (filter (fn (thm, _) => is_instance thy T | 
| 16645 | 76 | (snd (const_of (prop_of thm)))) thms) | 
| 77 | in if null thms' then ([], "") | |
| 78 | else (preprocess thy (map fst thms'), | |
| 79 | case snd (snd (split_last thms')) of | |
| 80 | NONE => (case get_defn thy defs s T of | |
| 81 | NONE => thyname_of_const s thy | |
| 82 | | SOME ((_, (thyname, _)), _) => thyname) | |
| 83 | | SOME thyname => thyname) | |
| 84 | end); | |
| 12447 | 85 | |
| 16645 | 86 | fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of | 
| 17144 | 87 | SOME (_, SOME i) => " def" ^ string_of_int i | _ => ""); | 
| 12447 | 88 | |
| 89 | exception EQN of string * typ * string; | |
| 90 | ||
| 22887 | 91 | fun cycle g (xs, x : string) = | 
| 92 | if member (op =) xs x then xs | |
| 93 | else Library.foldl (cycle g) (x :: xs, flat (Graph.all_paths (fst g) (x, x))); | |
| 12447 | 94 | |
| 17144 | 95 | fun add_rec_funs thy defs gr dep eqs module = | 
| 12447 | 96 | let | 
| 16645 | 97 | fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t), | 
| 98 | dest_eqn (rename_term t)); | |
| 12447 | 99 | val eqs' = map dest_eq eqs; | 
| 100 | val (dname, _) :: _ = eqs'; | |
| 101 | val (s, T) = const_of (hd eqs); | |
| 102 | ||
| 17144 | 103 | fun mk_fundef module fname prfx gr [] = (gr, []) | 
| 21410 | 104 | | mk_fundef module fname prfx gr ((fname' : string, (lhs, rhs)) :: xs) = | 
| 12447 | 105 | let | 
| 17144 | 106 | val (gr1, pl) = invoke_codegen thy defs dname module false (gr, lhs); | 
| 107 | val (gr2, pr) = invoke_codegen thy defs dname module false (gr1, rhs); | |
| 108 | val (gr3, rest) = mk_fundef module fname' "and " gr2 xs | |
| 12447 | 109 | in | 
| 110 | (gr3, Pretty.blk (4, [Pretty.str (if fname = fname' then " | " else prfx), | |
| 111 | pl, Pretty.str " =", Pretty.brk 1, pr]) :: rest) | |
| 112 | end; | |
| 113 | ||
| 17144 | 114 | fun put_code module fundef = map_node dname | 
| 115 |       (K (SOME (EQN ("", dummyT, dname)), module, Pretty.string_of (Pretty.blk (0,
 | |
| 12447 | 116 | separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n")); | 
| 117 | ||
| 118 | in | |
| 17144 | 119 | (case try (get_node gr) dname of | 
| 15531 | 120 | NONE => | 
| 12447 | 121 | let | 
| 17144 | 122 | val gr1 = add_edge (dname, dep) | 
| 123 | (new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr); | |
| 124 | val (gr2, fundef) = mk_fundef module "" "fun " gr1 eqs'; | |
| 12447 | 125 | val xs = cycle gr2 ([], dname); | 
| 17144 | 126 | val cs = map (fn x => case get_node gr2 x of | 
| 16645 | 127 | (SOME (EQN (s, T, _)), _, _) => (s, T) | 
| 12447 | 128 |              | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
 | 
| 129 | implode (separate ", " xs))) xs | |
| 130 | in (case xs of | |
| 17144 | 131 | [_] => (put_code module fundef gr2, module) | 
| 12447 | 132 | | _ => | 
| 133 | if not (dep mem xs) then | |
| 134 | let | |
| 16645 | 135 | val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs; | 
| 17144 | 136 | val module' = if_library thyname module; | 
| 16645 | 137 | val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss)); | 
| 17144 | 138 | val (gr3, fundef') = mk_fundef module' "" "fun " | 
| 139 | (add_edge (dname, dep) | |
| 140 | (foldr (uncurry new_node) (del_nodes xs gr2) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 141 | (map (fn k => | 
| 17144 | 142 |                          (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs))) eqs''
 | 
| 143 | in (put_code module' fundef' gr3, module') end | |
| 144 | else (gr2, module)) | |
| 12447 | 145 | end | 
| 17144 | 146 | | SOME (SOME (EQN (_, _, s)), module', _) => | 
| 147 | (if s = "" then | |
| 148 | if dname = dep then gr else add_edge (dname, dep) gr | |
| 149 | else if s = dep then gr else add_edge (s, dep) gr, | |
| 150 | module')) | |
| 12447 | 151 | end; | 
| 152 | ||
| 17144 | 153 | fun recfun_codegen thy defs gr dep module brack t = (case strip_comb t of | 
| 22921 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22887diff
changeset | 154 | (Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy (s, T)) of | 
| 16645 | 155 | (([], _), _) => NONE | 
| 15531 | 156 | | (_, SOME _) => NONE | 
| 17144 | 157 | | ((eqns, thyname), NONE) => | 
| 16645 | 158 | let | 
| 17144 | 159 | val module' = if_library thyname module; | 
| 16645 | 160 | val (gr', ps) = foldl_map | 
| 17144 | 161 | (invoke_codegen thy defs dep module true) (gr, ts); | 
| 162 | val suffix = mk_suffix thy defs p; | |
| 163 | val (gr'', module'') = | |
| 164 | add_rec_funs thy defs gr' dep (map prop_of eqns) module'; | |
| 165 | val (gr''', fname) = mk_const_id module'' (s ^ suffix) gr'' | |
| 12447 | 166 | in | 
| 17144 | 167 | SOME (gr''', mk_app brack (Pretty.str (mk_qual_id module fname)) ps) | 
| 12447 | 168 | end) | 
| 15531 | 169 | | _ => NONE); | 
| 12447 | 170 | |
| 171 | ||
| 18708 | 172 | val setup = | 
| 173 | add_codegen "recfun" recfun_codegen #> | |
| 18220 | 174 | add_attribute "" | 
| 19344 | 175 | (Args.del |-- Scan.succeed del | 
| 21128 | 176 | || Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add); | 
| 12447 | 177 | |
| 178 | end; |