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