| author | streckem | 
| Wed, 23 Oct 2002 16:10:42 +0200 | |
| changeset 13674 | f4c64597fb02 | 
| parent 13073 | cc9d7f403a4b | 
| child 13731 | e2d17090052b | 
| permissions | -rw-r--r-- | 
| 11520 | 1 | (* Title: Pure/codegen.ML | 
| 2 | ID: $Id$ | |
| 11539 | 3 | Author: Stefan Berghofer, TU Muenchen | 
| 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | |
| 11520 | 5 | |
| 11539 | 6 | Generic code generator. | 
| 11520 | 7 | *) | 
| 8 | ||
| 9 | signature CODEGEN = | |
| 10 | sig | |
| 11 | val quiet_mode : bool ref | |
| 12 | val message : string -> unit | |
| 13 | ||
| 12452 | 14 | datatype 'a mixfix = | 
| 11520 | 15 | Arg | 
| 16 | | Ignore | |
| 17 | | Pretty of Pretty.T | |
| 12452 | 18 | | Quote of 'a; | 
| 11520 | 19 | |
| 12452 | 20 | type 'a codegen | 
| 21 | ||
| 22 | val add_codegen: string -> term codegen -> theory -> theory | |
| 23 | val add_tycodegen: string -> typ codegen -> theory -> theory | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 24 | val add_attribute: string -> theory attribute -> theory -> theory | 
| 11520 | 25 | val print_codegens: theory -> unit | 
| 26 | val generate_code: theory -> (string * string) list -> string | |
| 27 | val generate_code_i: theory -> (string * term) list -> string | |
| 12452 | 28 | val assoc_consts: (xstring * string option * term mixfix list) list -> theory -> theory | 
| 29 | val assoc_consts_i: (xstring * typ option * term mixfix list) list -> theory -> theory | |
| 30 | val assoc_types: (xstring * typ mixfix list) list -> theory -> theory | |
| 31 | val get_assoc_code: theory -> string -> typ -> term mixfix list option | |
| 32 | val get_assoc_type: theory -> string -> typ mixfix list option | |
| 33 | val invoke_codegen: theory -> string -> bool -> | |
| 34 | (exn option * string) Graph.T * term -> (exn option * string) Graph.T * Pretty.T | |
| 35 | val invoke_tycodegen: theory -> string -> bool -> | |
| 36 | (exn option * string) Graph.T * typ -> (exn option * string) Graph.T * Pretty.T | |
| 11520 | 37 | val mk_const_id: Sign.sg -> string -> string | 
| 38 | val mk_type_id: Sign.sg -> string -> string | |
| 39 | val rename_term: term -> term | |
| 40 | val get_defn: theory -> string -> typ -> ((term list * term) * int option) option | |
| 41 | val is_instance: theory -> typ -> typ -> bool | |
| 42 | val parens: Pretty.T -> Pretty.T | |
| 43 | val mk_app: bool -> Pretty.T -> Pretty.T list -> Pretty.T | |
| 44 | val eta_expand: term -> term list -> int -> term | |
| 12452 | 45 | val parse_mixfix: (string -> 'a) -> string -> 'a mixfix list | 
| 11520 | 46 | val parsers: OuterSyntax.parser list | 
| 47 | val setup: (theory -> theory) list | |
| 48 | end; | |
| 49 | ||
| 50 | structure Codegen : CODEGEN = | |
| 51 | struct | |
| 52 | ||
| 53 | val quiet_mode = ref true; | |
| 54 | fun message s = if !quiet_mode then () else writeln s; | |
| 55 | ||
| 56 | (**** Mixfix syntax ****) | |
| 57 | ||
| 12452 | 58 | datatype 'a mixfix = | 
| 11520 | 59 | Arg | 
| 60 | | Ignore | |
| 61 | | Pretty of Pretty.T | |
| 12452 | 62 | | Quote of 'a; | 
| 11520 | 63 | |
| 64 | fun is_arg Arg = true | |
| 65 | | is_arg Ignore = true | |
| 66 | | is_arg _ = false; | |
| 67 | ||
| 12452 | 68 | fun quotes_of [] = [] | 
| 69 | | quotes_of (Quote q :: ms) = q :: quotes_of ms | |
| 70 | | quotes_of (_ :: ms) = quotes_of ms; | |
| 71 | ||
| 72 | fun args_of [] xs = ([], xs) | |
| 73 | | args_of (Arg :: ms) (x :: xs) = apfst (cons x) (args_of ms xs) | |
| 74 | | args_of (Ignore :: ms) (_ :: xs) = args_of ms xs | |
| 75 | | args_of (_ :: ms) xs = args_of ms xs; | |
| 11520 | 76 | |
| 12490 | 77 | fun num_args x = length (filter is_arg x); | 
| 11520 | 78 | |
| 79 | ||
| 80 | (**** theory data ****) | |
| 81 | ||
| 82 | (* data kind 'Pure/codegen' *) | |
| 83 | ||
| 12452 | 84 | type 'a codegen = theory -> (exn option * string) Graph.T -> | 
| 85 | string -> bool -> 'a -> ((exn option * string) Graph.T * Pretty.T) option; | |
| 86 | ||
| 11520 | 87 | structure CodegenArgs = | 
| 88 | struct | |
| 89 | val name = "Pure/codegen"; | |
| 90 | type T = | |
| 12452 | 91 |     {codegens : (string * term codegen) list,
 | 
| 92 | tycodegens : (string * typ codegen) list, | |
| 93 | consts : ((string * typ) * term mixfix list) list, | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 94 | types : (string * typ mixfix list) list, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 95 | attrs: (string * theory attribute) list}; | 
| 11520 | 96 | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 97 | val empty = | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 98 |     {codegens = [], tycodegens = [], consts = [], types = [], attrs = []};
 | 
| 11520 | 99 | val copy = I; | 
| 100 | val prep_ext = I; | |
| 101 | ||
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 102 | fun merge | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 103 |     ({codegens = codegens1, tycodegens = tycodegens1,
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 104 | consts = consts1, types = types1, attrs = attrs1}, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 105 |      {codegens = codegens2, tycodegens = tycodegens2,
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 106 | consts = consts2, types = types2, attrs = attrs2}) = | 
| 11520 | 107 |     {codegens = rev (merge_alists (rev codegens1) (rev codegens2)),
 | 
| 12452 | 108 | tycodegens = rev (merge_alists (rev tycodegens1) (rev tycodegens2)), | 
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 109 | consts = merge_alists consts1 consts2, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 110 | types = merge_alists types1 types2, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 111 | attrs = merge_alists attrs1 attrs2}; | 
| 11520 | 112 | |
| 12452 | 113 |   fun print sg ({codegens, tycodegens, ...} : T) =
 | 
| 114 | Pretty.writeln (Pretty.chunks | |
| 115 |       [Pretty.strs ("term code generators:" :: map fst codegens),
 | |
| 116 |        Pretty.strs ("type code generators:" :: map fst tycodegens)]);
 | |
| 11520 | 117 | end; | 
| 118 | ||
| 119 | structure CodegenData = TheoryDataFun(CodegenArgs); | |
| 120 | val print_codegens = CodegenData.print; | |
| 121 | ||
| 122 | ||
| 12452 | 123 | (**** add new code generators to theory ****) | 
| 11520 | 124 | |
| 125 | fun add_codegen name f thy = | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 126 |   let val {codegens, tycodegens, consts, types, attrs} = CodegenData.get thy
 | 
| 11520 | 127 | in (case assoc (codegens, name) of | 
| 12452 | 128 |       None => CodegenData.put {codegens = (name, f) :: codegens,
 | 
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 129 | tycodegens = tycodegens, consts = consts, types = types, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 130 | attrs = attrs} thy | 
| 12452 | 131 |     | Some _ => error ("Code generator " ^ name ^ " already declared"))
 | 
| 132 | end; | |
| 133 | ||
| 134 | fun add_tycodegen name f thy = | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 135 |   let val {codegens, tycodegens, consts, types, attrs} = CodegenData.get thy
 | 
| 12452 | 136 | in (case assoc (tycodegens, name) of | 
| 137 |       None => CodegenData.put {tycodegens = (name, f) :: tycodegens,
 | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 138 | codegens = codegens, consts = consts, types = types, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 139 | attrs = attrs} thy | 
| 11520 | 140 |     | Some _ => error ("Code generator " ^ name ^ " already declared"))
 | 
| 141 | end; | |
| 142 | ||
| 143 | ||
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 144 | (**** code attribute ****) | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 145 | |
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 146 | fun add_attribute name att thy = | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 147 |   let val {codegens, tycodegens, consts, types, attrs} = CodegenData.get thy
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 148 | in (case assoc (attrs, name) of | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 149 |       None => CodegenData.put {tycodegens = tycodegens,
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 150 | codegens = codegens, consts = consts, types = types, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 151 | attrs = (name, att) :: attrs} thy | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 152 |     | Some _ => error ("Code attribute " ^ name ^ " already declared"))
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 153 | end; | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 154 | |
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 155 | val code_attr = | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 156 | Attrib.syntax (Scan.depend (fn thy => Scan.optional Args.name "" >> | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 157 | (fn name => (thy, case assoc (#attrs (CodegenData.get thy), name) of | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 158 |           None => error ("Unknown code attribute: " ^ quote name)
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 159 | | Some att => att)))); | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 160 | |
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 161 | |
| 11520 | 162 | (**** associate constants with target language code ****) | 
| 163 | ||
| 164 | fun gen_assoc_consts prep_type xs thy = foldl (fn (thy, (s, tyopt, syn)) => | |
| 165 | let | |
| 166 | val sg = sign_of thy; | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 167 |     val {codegens, tycodegens, consts, types, attrs} = CodegenData.get thy;
 | 
| 11520 | 168 | val cname = Sign.intern_const sg s; | 
| 169 | in | |
| 170 | (case Sign.const_type sg cname of | |
| 171 | Some T => | |
| 172 | let val T' = (case tyopt of | |
| 173 | None => T | |
| 174 | | Some ty => | |
| 175 | let val U = prep_type sg ty | |
| 176 | in if Type.typ_instance (Sign.tsig_of sg, U, T) then U | |
| 177 |                     else error ("Illegal type constraint for constant " ^ cname)
 | |
| 178 | end) | |
| 179 | in (case assoc (consts, (cname, T')) of | |
| 180 |              None => CodegenData.put {codegens = codegens,
 | |
| 12452 | 181 | tycodegens = tycodegens, | 
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 182 | consts = ((cname, T'), syn) :: consts, | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 183 | types = types, attrs = attrs} thy | 
| 11520 | 184 |            | Some _ => error ("Constant " ^ cname ^ " already associated with code"))
 | 
| 185 | end | |
| 186 |      | _ => error ("Not a constant: " ^ s))
 | |
| 187 | end) (thy, xs); | |
| 188 | ||
| 189 | val assoc_consts_i = gen_assoc_consts (K I); | |
| 190 | val assoc_consts = gen_assoc_consts (fn sg => typ_of o read_ctyp sg); | |
| 191 | ||
| 192 | (**** associate types with target language types ****) | |
| 193 | ||
| 194 | fun assoc_types xs thy = foldl (fn (thy, (s, syn)) => | |
| 195 | let | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 196 |     val {codegens, tycodegens, consts, types, attrs} = CodegenData.get thy;
 | 
| 11520 | 197 | val tc = Sign.intern_tycon (sign_of thy) s | 
| 198 | in | |
| 199 | (case assoc (types, tc) of | |
| 12452 | 200 |        None => CodegenData.put {codegens = codegens,
 | 
| 201 | tycodegens = tycodegens, consts = consts, | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 202 | types = (tc, syn) :: types, attrs = attrs} thy | 
| 11520 | 203 |      | Some _ => error ("Type " ^ tc ^ " already associated with code"))
 | 
| 204 | end) (thy, xs); | |
| 205 | ||
| 12452 | 206 | fun get_assoc_type thy s = assoc (#types (CodegenData.get thy), s); | 
| 11546 | 207 | |
| 11520 | 208 | |
| 209 | (**** retrieve definition of constant ****) | |
| 210 | ||
| 211 | fun is_instance thy T1 T2 = | |
| 212 | Type.typ_instance (Sign.tsig_of (sign_of thy), T1, Type.varifyT T2); | |
| 213 | ||
| 214 | fun get_assoc_code thy s T = apsome snd (find_first (fn ((s', T'), _) => | |
| 215 | s = s' andalso is_instance thy T T') (#consts (CodegenData.get thy))); | |
| 216 | ||
| 217 | fun get_defn thy s T = | |
| 218 | let | |
| 219 | val axms = flat (map (Symtab.dest o #axioms o Theory.rep_theory) | |
| 220 | (thy :: Theory.ancestors_of thy)); | |
| 221 | val defs = mapfilter (fn (_, t) => | |
| 222 | (let | |
| 223 | val (lhs, rhs) = Logic.dest_equals t; | |
| 224 | val (c, args) = strip_comb lhs; | |
| 225 | val (s', T') = dest_Const c | |
| 226 | in if s=s' then Some (T', (args, rhs)) else None end) handle TERM _ => | |
| 227 | None) axms; | |
| 228 | val i = find_index (is_instance thy T o fst) defs | |
| 229 | in | |
| 230 | if i>=0 then Some (snd (nth_elem (i, defs)), | |
| 231 | if length defs = 1 then None else Some i) | |
| 232 | else None | |
| 233 | end; | |
| 234 | ||
| 235 | ||
| 236 | (**** make valid ML identifiers ****) | |
| 237 | ||
| 238 | fun gen_mk_id kind rename sg s = | |
| 239 | let | |
| 240 | val (xs as x::_) = explode (rename (space_implode "_" | |
| 241 | (NameSpace.unpack (Sign.cond_extern sg kind s)))); | |
| 242 | fun check_str [] = "" | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 243 |       | check_str (" " :: xs) = "_" ^ check_str xs
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 244 | | check_str (x :: xs) = | 
| 12824 | 245 | (if Symbol.is_letdig x then x | 
| 11520 | 246 | else "_" ^ string_of_int (ord x)) ^ check_str xs | 
| 247 | in | |
| 248 | (if not (Symbol.is_letter x) then "id" else "") ^ check_str xs | |
| 249 | end; | |
| 250 | ||
| 13073 
cc9d7f403a4b
mk_const_id now checks for clashes with reserved ML identifiers.
 berghofe parents: 
13003diff
changeset | 251 | val mk_const_id = gen_mk_id Sign.constK | 
| 
cc9d7f403a4b
mk_const_id now checks for clashes with reserved ML identifiers.
 berghofe parents: 
13003diff
changeset | 252 | (fn s => if s mem ThmDatabase.ml_reserved then s ^ "_const" else s); | 
| 
cc9d7f403a4b
mk_const_id now checks for clashes with reserved ML identifiers.
 berghofe parents: 
13003diff
changeset | 253 | |
| 11520 | 254 | val mk_type_id = gen_mk_id Sign.typeK | 
| 255 | (fn s => if s mem ThmDatabase.ml_reserved then s ^ "_type" else s); | |
| 256 | ||
| 257 | fun rename_term t = | |
| 258 | let | |
| 259 | val names = add_term_names (t, map (fst o fst o dest_Var) (term_vars t)); | |
| 260 | val clash = names inter ThmDatabase.ml_reserved; | |
| 261 | val ps = clash ~~ variantlist (clash, names); | |
| 262 | ||
| 263 | fun rename (Var ((a, i), T)) = Var ((if_none (assoc (ps, a)) a, i), T) | |
| 264 | | rename (Free (a, T)) = Free (if_none (assoc (ps, a)) a, T) | |
| 265 | | rename (Abs (s, T, t)) = Abs (s, T, rename t) | |
| 266 | | rename (t $ u) = rename t $ rename u | |
| 267 | | rename t = t; | |
| 268 | in | |
| 269 | rename t | |
| 270 | end; | |
| 271 | ||
| 272 | ||
| 12452 | 273 | (**** invoke suitable code generator for term / type ****) | 
| 11520 | 274 | |
| 12452 | 275 | fun invoke_codegen thy dep brack (gr, t) = (case get_first | 
| 11520 | 276 | (fn (_, f) => f thy gr dep brack t) (#codegens (CodegenData.get thy)) of | 
| 277 |       None => error ("Unable to generate code for term:\n" ^
 | |
| 278 | Sign.string_of_term (sign_of thy) t ^ "\nrequired by:\n" ^ | |
| 279 | commas (Graph.all_succs gr [dep])) | |
| 12452 | 280 | | Some x => x); | 
| 281 | ||
| 282 | fun invoke_tycodegen thy dep brack (gr, T) = (case get_first | |
| 283 | (fn (_, f) => f thy gr dep brack T) (#tycodegens (CodegenData.get thy)) of | |
| 284 |       None => error ("Unable to generate code for type:\n" ^
 | |
| 285 | Sign.string_of_typ (sign_of thy) T ^ "\nrequired by:\n" ^ | |
| 286 | commas (Graph.all_succs gr [dep])) | |
| 287 | | Some x => x); | |
| 11520 | 288 | |
| 289 | ||
| 290 | (**** code generator for mixfix expressions ****) | |
| 291 | ||
| 292 | fun parens p = Pretty.block [Pretty.str "(", p, Pretty.str ")"];
 | |
| 293 | ||
| 294 | fun pretty_fn [] p = [p] | |
| 295 |   | pretty_fn (x::xs) p = Pretty.str ("fn " ^ x ^ " =>") ::
 | |
| 296 | Pretty.brk 1 :: pretty_fn xs p; | |
| 297 | ||
| 298 | fun pretty_mixfix [] [] _ = [] | |
| 299 | | pretty_mixfix (Arg :: ms) (p :: ps) qs = p :: pretty_mixfix ms ps qs | |
| 12452 | 300 | | pretty_mixfix (Ignore :: ms) ps qs = pretty_mixfix ms ps qs | 
| 11520 | 301 | | pretty_mixfix (Pretty p :: ms) ps qs = p :: pretty_mixfix ms ps qs | 
| 12452 | 302 | | pretty_mixfix (Quote _ :: ms) ps (q :: qs) = q :: pretty_mixfix ms ps qs; | 
| 11520 | 303 | |
| 304 | ||
| 12452 | 305 | (**** default code generators ****) | 
| 11520 | 306 | |
| 307 | fun eta_expand t ts i = | |
| 308 | let | |
| 309 | val (Ts, _) = strip_type (fastype_of t); | |
| 310 | val j = i - length ts | |
| 311 | in | |
| 312 |     foldr (fn (T, t) => Abs ("x", T, t))
 | |
| 313 | (take (j, Ts), list_comb (t, ts @ map Bound (j-1 downto 0))) | |
| 314 | end; | |
| 315 | ||
| 316 | fun mk_app _ p [] = p | |
| 317 | | mk_app brack p ps = if brack then | |
| 318 |        Pretty.block (Pretty.str "(" ::
 | |
| 319 | separate (Pretty.brk 1) (p :: ps) @ [Pretty.str ")"]) | |
| 320 | else Pretty.block (separate (Pretty.brk 1) (p :: ps)); | |
| 321 | ||
| 322 | fun new_names t xs = variantlist (xs, | |
| 323 | map (fst o fst o dest_Var) (term_vars t) union | |
| 324 | add_term_names (t, ThmDatabase.ml_reserved)); | |
| 325 | ||
| 326 | fun new_name t x = hd (new_names t [x]); | |
| 327 | ||
| 328 | fun default_codegen thy gr dep brack t = | |
| 329 | let | |
| 330 | val (u, ts) = strip_comb t; | |
| 12452 | 331 | fun codegens brack = foldl_map (invoke_codegen thy dep brack) | 
| 11520 | 332 | in (case u of | 
| 333 | Var ((s, i), _) => | |
| 12452 | 334 | let val (gr', ps) = codegens true (gr, ts) | 
| 11520 | 335 | in Some (gr', mk_app brack (Pretty.str (s ^ | 
| 336 | (if i=0 then "" else string_of_int i))) ps) | |
| 337 | end | |
| 338 | ||
| 339 | | Free (s, _) => | |
| 12452 | 340 | let val (gr', ps) = codegens true (gr, ts) | 
| 11520 | 341 | in Some (gr', mk_app brack (Pretty.str s) ps) end | 
| 342 | ||
| 343 | | Const (s, T) => | |
| 344 | (case get_assoc_code thy s T of | |
| 345 | Some ms => | |
| 346 | let val i = num_args ms | |
| 347 | in if length ts < i then | |
| 348 | default_codegen thy gr dep brack (eta_expand u ts i) | |
| 349 | else | |
| 350 | let | |
| 12452 | 351 | val (ts1, ts2) = args_of ms ts; | 
| 352 | val (gr1, ps1) = codegens false (gr, ts1); | |
| 353 | val (gr2, ps2) = codegens true (gr1, ts2); | |
| 354 | val (gr3, ps3) = codegens false (gr2, quotes_of ms); | |
| 11520 | 355 | in | 
| 356 | Some (gr3, mk_app brack (Pretty.block (pretty_mixfix ms ps1 ps3)) ps2) | |
| 357 | end | |
| 358 | end | |
| 359 | | None => (case get_defn thy s T of | |
| 360 | None => None | |
| 361 | | Some ((args, rhs), k) => | |
| 362 | let | |
| 363 | val id = mk_const_id (sign_of thy) s ^ (case k of | |
| 364 | None => "" | Some i => "_def" ^ string_of_int i); | |
| 12452 | 365 | val (gr', ps) = codegens true (gr, ts); | 
| 11520 | 366 | in | 
| 367 | Some (Graph.add_edge (id, dep) gr' handle Graph.UNDEF _ => | |
| 368 | let | |
| 369 |                    val _ = message ("expanding definition of " ^ s);
 | |
| 370 | val (Ts, _) = strip_type T; | |
| 371 | val (args', rhs') = | |
| 372 | if not (null args) orelse null Ts then (args, rhs) else | |
| 373 | let val v = Free (new_name rhs "x", hd Ts) | |
| 374 | in ([v], betapply (rhs, v)) end; | |
| 12452 | 375 | val (gr1, p) = invoke_codegen thy id false | 
| 376 | (Graph.add_edge (id, dep) | |
| 377 | (Graph.new_node (id, (None, "")) gr'), rhs'); | |
| 378 | val (gr2, xs) = codegens false (gr1, args'); | |
| 12580 
7fdc00bb2a9e
Code generator now adds type constraints to val declarations (to make
 berghofe parents: 
12555diff
changeset | 379 | val (gr3, ty) = invoke_tycodegen thy id false (gr2, T); | 
| 11520 | 380 | in Graph.map_node id (K (None, Pretty.string_of (Pretty.block | 
| 12580 
7fdc00bb2a9e
Code generator now adds type constraints to val declarations (to make
 berghofe parents: 
12555diff
changeset | 381 | (separate (Pretty.brk 1) (if null args' then | 
| 
7fdc00bb2a9e
Code generator now adds type constraints to val declarations (to make
 berghofe parents: 
12555diff
changeset | 382 |                        [Pretty.str ("val " ^ id ^ " :"), ty]
 | 
| 
7fdc00bb2a9e
Code generator now adds type constraints to val declarations (to make
 berghofe parents: 
12555diff
changeset | 383 |                      else Pretty.str ("fun " ^ id) :: xs) @
 | 
| 
7fdc00bb2a9e
Code generator now adds type constraints to val declarations (to make
 berghofe parents: 
12555diff
changeset | 384 | [Pretty.str " =", Pretty.brk 1, p, Pretty.str ";"])) ^ "\n\n")) gr3 | 
| 11520 | 385 | end, mk_app brack (Pretty.str id) ps) | 
| 386 | end)) | |
| 387 | ||
| 388 | | Abs _ => | |
| 389 | let | |
| 390 | val (bs, Ts) = ListPair.unzip (strip_abs_vars u); | |
| 391 | val t = strip_abs_body u | |
| 392 | val bs' = new_names t bs; | |
| 12452 | 393 | val (gr1, ps) = codegens true (gr, ts); | 
| 394 | val (gr2, p) = invoke_codegen thy dep false | |
| 395 | (gr1, subst_bounds (map Free (rev (bs' ~~ Ts)), t)); | |
| 11520 | 396 | in | 
| 397 |         Some (gr2, mk_app brack (Pretty.block (Pretty.str "(" :: pretty_fn bs' p @
 | |
| 398 | [Pretty.str ")"])) ps) | |
| 399 | end | |
| 400 | ||
| 401 | | _ => None) | |
| 402 | end; | |
| 403 | ||
| 12452 | 404 | fun default_tycodegen thy gr dep brack (TVar ((s, i), _)) = | 
| 405 | Some (gr, Pretty.str (s ^ (if i = 0 then "" else string_of_int i))) | |
| 406 | | default_tycodegen thy gr dep brack (TFree (s, _)) = Some (gr, Pretty.str s) | |
| 407 | | default_tycodegen thy gr dep brack (Type (s, Ts)) = | |
| 408 | (case assoc (#types (CodegenData.get thy), s) of | |
| 409 | None => None | |
| 410 | | Some ms => | |
| 411 | let | |
| 412 | val (gr', ps) = foldl_map | |
| 413 | (invoke_tycodegen thy dep false) (gr, fst (args_of ms Ts)); | |
| 414 | val (gr'', qs) = foldl_map | |
| 415 | (invoke_tycodegen thy dep false) (gr', quotes_of ms) | |
| 416 | in Some (gr'', Pretty.block (pretty_mixfix ms ps qs)) end); | |
| 417 | ||
| 11520 | 418 | |
| 419 | fun output_code gr xs = implode (map (snd o Graph.get_node gr) | |
| 420 | (rev (Graph.all_preds gr xs))); | |
| 421 | ||
| 422 | fun gen_generate_code prep_term thy = Pretty.setmp_margin 80 (fn xs => | |
| 423 | let | |
| 424 | val sg = sign_of thy; | |
| 425 |     val gr = Graph.new_node ("<Top>", (None, "")) Graph.empty;
 | |
| 426 | val (gr', ps) = foldl_map (fn (gr, (s, t)) => apsnd (pair s) | |
| 12452 | 427 | (invoke_codegen thy "<Top>" false (gr, t))) | 
| 428 | (gr, map (apsnd (prep_term sg)) xs) | |
| 11520 | 429 | in | 
| 430 | "structure Generated =\nstruct\n\n" ^ | |
| 431 | output_code gr' ["<Top>"] ^ | |
| 432 | space_implode "\n\n" (map (fn (s', p) => Pretty.string_of (Pretty.block | |
| 433 |       [Pretty.str ("val " ^ s' ^ " ="), Pretty.brk 1, p, Pretty.str ";"])) ps) ^
 | |
| 434 | "\n\nend;\n\nopen Generated;\n" | |
| 435 | end); | |
| 436 | ||
| 437 | val generate_code_i = gen_generate_code (K I); | |
| 438 | val generate_code = gen_generate_code | |
| 439 | (fn sg => term_of o read_cterm sg o rpair TypeInfer.logicT); | |
| 440 | ||
| 12452 | 441 | |
| 442 | (**** Interface ****) | |
| 443 | ||
| 11520 | 444 | fun parse_mixfix rd s = | 
| 445 | (case Scan.finite Symbol.stopper (Scan.repeat | |
| 446 | ( $$ "_" >> K Arg | |
| 447 | || $$ "?" >> K Ignore | |
| 448 | || $$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length) | |
| 449 |       || $$ "{" |-- $$ "*" |-- Scan.repeat1
 | |
| 450 | ( $$ "'" |-- Scan.one Symbol.not_eof | |
| 451 | || Scan.unless ($$ "*" -- $$ "}") (Scan.one Symbol.not_eof)) --| | |
| 12452 | 452 | $$ "*" --| $$ "}" >> (Quote o rd o implode) | 
| 11520 | 453 | || Scan.repeat1 | 
| 454 | ( $$ "'" |-- Scan.one Symbol.not_eof | |
| 455 |             || Scan.unless ($$ "_" || $$ "?" || $$ "/" || $$ "{" |-- $$ "*")
 | |
| 456 | (Scan.one Symbol.not_eof)) >> (Pretty o Pretty.str o implode))) | |
| 457 | (Symbol.explode s) of | |
| 458 | (p, []) => p | |
| 459 |    | _ => error ("Malformed annotation: " ^ quote s));
 | |
| 460 | ||
| 11546 | 461 | structure P = OuterParse and K = OuterSyntax.Keyword; | 
| 11520 | 462 | |
| 463 | val assoc_typeP = | |
| 464 | OuterSyntax.command "types_code" | |
| 11546 | 465 | "associate types with target language types" K.thy_decl | 
| 11520 | 466 |     (Scan.repeat1 (P.xname --| P.$$$ "(" -- P.string --| P.$$$ ")") >>
 | 
| 12452 | 467 | (fn xs => Toplevel.theory (fn thy => assoc_types | 
| 468 | (map (fn (name, mfx) => (name, parse_mixfix | |
| 469 | (typ_of o read_ctyp (sign_of thy)) mfx)) xs) thy))); | |
| 11520 | 470 | |
| 471 | val assoc_constP = | |
| 472 | OuterSyntax.command "consts_code" | |
| 11546 | 473 | "associate constants with target language code" K.thy_decl | 
| 11520 | 474 | (Scan.repeat1 | 
| 13003 | 475 | (P.xname -- (Scan.option (P.$$$ "::" |-- P.typ)) --| | 
| 11520 | 476 |         P.$$$ "(" -- P.string --| P.$$$ ")") >>
 | 
| 477 | (fn xs => Toplevel.theory (fn thy => assoc_consts | |
| 478 | (map (fn ((name, optype), mfx) => (name, optype, parse_mixfix | |
| 479 | (term_of o read_cterm (sign_of thy) o rpair TypeInfer.logicT) mfx)) | |
| 480 | xs) thy))); | |
| 481 | ||
| 482 | val generate_codeP = | |
| 11546 | 483 | OuterSyntax.command "generate_code" "generates code for terms" K.thy_decl | 
| 13003 | 484 |     (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") --
 | 
| 485 | Scan.repeat1 (P.name --| P.$$$ "=" -- P.term) >> | |
| 11520 | 486 | (fn (opt_fname, xs) => Toplevel.theory (fn thy => | 
| 487 | ((case opt_fname of | |
| 488 | None => use_text Context.ml_output false | |
| 489 | | Some fname => File.write (Path.unpack fname)) | |
| 490 | (generate_code thy xs); thy)))); | |
| 491 | ||
| 492 | val parsers = [assoc_typeP, assoc_constP, generate_codeP]; | |
| 493 | ||
| 12452 | 494 | val setup = | 
| 495 | [CodegenData.init, | |
| 496 | add_codegen "default" default_codegen, | |
| 497 | add_tycodegen "default" default_tycodegen, | |
| 12555 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 498 |    assoc_types [("fun", parse_mixfix (K dummyT) "(_ ->/ _)")],
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 499 |    Attrib.add_attributes [("code",
 | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 500 | (code_attr, K Attrib.undef_local_attribute), | 
| 
e6d7f040fdc7
"code" attribute is now managed by basic code generator module.
 berghofe parents: 
12490diff
changeset | 501 | "declare theorems for code generation")]]; | 
| 11520 | 502 | |
| 503 | end; | |
| 504 | ||
| 505 | OuterSyntax.add_parsers Codegen.parsers; |