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