| author | wenzelm | 
| Tue, 08 Nov 2005 10:43:08 +0100 | |
| changeset 18117 | 61a430a67d7c | 
| parent 17412 | e26cb20ef0cc | 
| child 18220 | 43cf5767f992 | 
| 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 | |
| 16645 | 10 | val add: string option -> theory attribute | 
| 11 | val del: theory attribute | |
| 12447 | 12 | val setup: (theory -> theory) list | 
| 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; | 
| 16645 | 27 | fun merge _ = Symtab.merge_multi' (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: 
15257diff
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 | ||
| 16645 | 39 | fun add optmod (p as (thy, thm)) = | 
| 12447 | 40 | let | 
| 12556 | 41 | val tab = CodegenData.get thy; | 
| 12447 | 42 | val (s, _) = const_of (prop_of thm); | 
| 15321 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 43 | in | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 44 | if Pattern.pattern (lhs_of thm) then | 
| 17412 | 45 | (CodegenData.put (Symtab.update_multi (s, (thm, optmod)) tab) thy, thm) | 
| 15321 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 46 | else (warn thm; p) | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
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); | |
| 17412 | 53 | in case Symtab.lookup tab s of | 
| 15531 | 54 | NONE => p | 
| 17412 | 55 | | SOME thms => (CodegenData.put (Symtab.update (s, | 
| 17261 | 56 | gen_rem (eq_thm o apfst fst) (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: 
15257diff
changeset | 59 | fun del_redundant thy eqs [] = eqs | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 60 | | del_redundant thy eqs (eq :: eqs') = | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 61 | let | 
| 
694f9d3ce90d
Reimplemented some operations on "code lemma" table to avoid that code
 berghofe parents: 
15257diff
changeset | 62 | val matches = curry | 
| 17203 | 63 | (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 | 64 | 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 | 65 | |
| 16645 | 66 | fun get_equations thy defs (s, T) = | 
| 17412 | 67 | (case Symtab.lookup (CodegenData.get thy) s of | 
| 16645 | 68 | NONE => ([], "") | 
| 69 | | SOME thms => | |
| 70 | let val thms' = del_redundant thy [] | |
| 71 | (List.filter (fn (thm, _) => is_instance thy T | |
| 72 | (snd (const_of (prop_of thm)))) thms) | |
| 73 | in if null thms' then ([], "") | |
| 74 | else (preprocess thy (map fst thms'), | |
| 75 | case snd (snd (split_last thms')) of | |
| 76 | NONE => (case get_defn thy defs s T of | |
| 77 | NONE => thyname_of_const s thy | |
| 78 | | SOME ((_, (thyname, _)), _) => thyname) | |
| 79 | | SOME thyname => thyname) | |
| 80 | end); | |
| 12447 | 81 | |
| 16645 | 82 | fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of | 
| 17144 | 83 | SOME (_, SOME i) => " def" ^ string_of_int i | _ => ""); | 
| 12447 | 84 | |
| 85 | exception EQN of string * typ * string; | |
| 86 | ||
| 87 | fun cycle g (xs, x) = | |
| 88 | if x mem xs then xs | |
| 17144 | 89 | else Library.foldl (cycle g) (x :: xs, List.concat (Graph.find_paths (fst g) (x, x))); | 
| 12447 | 90 | |
| 17144 | 91 | fun add_rec_funs thy defs gr dep eqs module = | 
| 12447 | 92 | let | 
| 16645 | 93 | fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t), | 
| 94 | dest_eqn (rename_term t)); | |
| 12447 | 95 | val eqs' = map dest_eq eqs; | 
| 96 | val (dname, _) :: _ = eqs'; | |
| 97 | val (s, T) = const_of (hd eqs); | |
| 98 | ||
| 17144 | 99 | fun mk_fundef module fname prfx gr [] = (gr, []) | 
| 100 | | mk_fundef module fname prfx gr ((fname', (lhs, rhs)) :: xs) = | |
| 12447 | 101 | let | 
| 17144 | 102 | val (gr1, pl) = invoke_codegen thy defs dname module false (gr, lhs); | 
| 103 | val (gr2, pr) = invoke_codegen thy defs dname module false (gr1, rhs); | |
| 104 | val (gr3, rest) = mk_fundef module fname' "and " gr2 xs | |
| 12447 | 105 | in | 
| 106 | (gr3, Pretty.blk (4, [Pretty.str (if fname = fname' then " | " else prfx), | |
| 107 | pl, Pretty.str " =", Pretty.brk 1, pr]) :: rest) | |
| 108 | end; | |
| 109 | ||
| 17144 | 110 | fun put_code module fundef = map_node dname | 
| 111 |       (K (SOME (EQN ("", dummyT, dname)), module, Pretty.string_of (Pretty.blk (0,
 | |
| 12447 | 112 | separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n")); | 
| 113 | ||
| 114 | in | |
| 17144 | 115 | (case try (get_node gr) dname of | 
| 15531 | 116 | NONE => | 
| 12447 | 117 | let | 
| 17144 | 118 | val gr1 = add_edge (dname, dep) | 
| 119 | (new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr); | |
| 120 | val (gr2, fundef) = mk_fundef module "" "fun " gr1 eqs'; | |
| 12447 | 121 | val xs = cycle gr2 ([], dname); | 
| 17144 | 122 | val cs = map (fn x => case get_node gr2 x of | 
| 16645 | 123 | (SOME (EQN (s, T, _)), _, _) => (s, T) | 
| 12447 | 124 |              | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
 | 
| 125 | implode (separate ", " xs))) xs | |
| 126 | in (case xs of | |
| 17144 | 127 | [_] => (put_code module fundef gr2, module) | 
| 12447 | 128 | | _ => | 
| 129 | if not (dep mem xs) then | |
| 130 | let | |
| 16645 | 131 | val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs; | 
| 17144 | 132 | val module' = if_library thyname module; | 
| 16645 | 133 | val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss)); | 
| 17144 | 134 | val (gr3, fundef') = mk_fundef module' "" "fun " | 
| 135 | (add_edge (dname, dep) | |
| 136 | (foldr (uncurry new_node) (del_nodes xs gr2) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 137 | (map (fn k => | 
| 17144 | 138 |                          (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs))) eqs''
 | 
| 139 | in (put_code module' fundef' gr3, module') end | |
| 140 | else (gr2, module)) | |
| 12447 | 141 | end | 
| 17144 | 142 | | SOME (SOME (EQN (_, _, s)), module', _) => | 
| 143 | (if s = "" then | |
| 144 | if dname = dep then gr else add_edge (dname, dep) gr | |
| 145 | else if s = dep then gr else add_edge (s, dep) gr, | |
| 146 | module')) | |
| 12447 | 147 | end; | 
| 148 | ||
| 17144 | 149 | fun recfun_codegen thy defs gr dep module brack t = (case strip_comb t of | 
| 16645 | 150 | (Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy s T) of | 
| 151 | (([], _), _) => NONE | |
| 15531 | 152 | | (_, SOME _) => NONE | 
| 17144 | 153 | | ((eqns, thyname), NONE) => | 
| 16645 | 154 | let | 
| 17144 | 155 | val module' = if_library thyname module; | 
| 16645 | 156 | val (gr', ps) = foldl_map | 
| 17144 | 157 | (invoke_codegen thy defs dep module true) (gr, ts); | 
| 158 | val suffix = mk_suffix thy defs p; | |
| 159 | val (gr'', module'') = | |
| 160 | add_rec_funs thy defs gr' dep (map prop_of eqns) module'; | |
| 161 | val (gr''', fname) = mk_const_id module'' (s ^ suffix) gr'' | |
| 12447 | 162 | in | 
| 17144 | 163 | SOME (gr''', mk_app brack (Pretty.str (mk_qual_id module fname)) ps) | 
| 12447 | 164 | end) | 
| 15531 | 165 | | _ => NONE); | 
| 12447 | 166 | |
| 167 | ||
| 168 | val setup = | |
| 12556 | 169 | [CodegenData.init, | 
| 12447 | 170 | add_codegen "recfun" recfun_codegen, | 
| 16645 | 171 | add_attribute "" | 
| 172 | ( Args.del |-- Scan.succeed del | |
| 173 | || Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add)]; | |
| 12447 | 174 | |
| 175 | end; |