berghofe@11520: (* Title: Pure/codegen.ML berghofe@11520: ID: $Id$ berghofe@11520: Author: Stefan Berghofer berghofe@11520: Copyright 2000 TU Muenchen berghofe@11520: berghofe@11520: Generic code generator berghofe@11520: *) berghofe@11520: berghofe@11520: signature CODEGEN = berghofe@11520: sig berghofe@11520: val quiet_mode : bool ref berghofe@11520: val message : string -> unit berghofe@11520: berghofe@11520: datatype mixfix = berghofe@11520: Arg berghofe@11520: | Ignore berghofe@11520: | Pretty of Pretty.T berghofe@11520: | Term of term; berghofe@11520: berghofe@11520: val add_codegen: string -> berghofe@11520: (theory -> (exn option * string) Graph.T -> string -> bool -> term -> berghofe@11520: ((exn option * string) Graph.T * Pretty.T) option) -> theory -> theory berghofe@11520: val print_codegens: theory -> unit berghofe@11520: val generate_code: theory -> (string * string) list -> string berghofe@11520: val generate_code_i: theory -> (string * term) list -> string berghofe@11520: val assoc_consts: (xstring * string option * mixfix list) list -> theory -> theory berghofe@11520: val assoc_consts_i: (xstring * typ option * mixfix list) list -> theory -> theory berghofe@11520: val assoc_types: (xstring * string) list -> theory -> theory berghofe@11520: val get_assoc_code: theory -> string -> typ -> mixfix list option berghofe@11520: val get_assoc_types: theory -> (string * string) list berghofe@11520: val invoke_codegen: theory -> (exn option * string) Graph.T -> berghofe@11520: string -> bool -> term -> (exn option * string) Graph.T * Pretty.T berghofe@11520: val mk_const_id: Sign.sg -> string -> string berghofe@11520: val mk_type_id: Sign.sg -> string -> string berghofe@11520: val rename_term: term -> term berghofe@11520: val get_defn: theory -> string -> typ -> ((term list * term) * int option) option berghofe@11520: val is_instance: theory -> typ -> typ -> bool berghofe@11520: val parens: Pretty.T -> Pretty.T berghofe@11520: val mk_app: bool -> Pretty.T -> Pretty.T list -> Pretty.T berghofe@11520: val eta_expand: term -> term list -> int -> term berghofe@11520: val parsers: OuterSyntax.parser list berghofe@11520: val setup: (theory -> theory) list berghofe@11520: end; berghofe@11520: berghofe@11520: structure Codegen : CODEGEN = berghofe@11520: struct berghofe@11520: berghofe@11520: val quiet_mode = ref true; berghofe@11520: fun message s = if !quiet_mode then () else writeln s; berghofe@11520: berghofe@11520: (**** Mixfix syntax ****) berghofe@11520: berghofe@11520: datatype mixfix = berghofe@11520: Arg berghofe@11520: | Ignore berghofe@11520: | Pretty of Pretty.T berghofe@11520: | Term of term; berghofe@11520: berghofe@11520: fun is_arg Arg = true berghofe@11520: | is_arg Ignore = true berghofe@11520: | is_arg _ = false; berghofe@11520: berghofe@11520: fun terms_of [] = [] berghofe@11520: | terms_of (Term t :: ms) = t :: terms_of ms berghofe@11520: | terms_of (_ :: ms) = terms_of ms; berghofe@11520: berghofe@11520: val num_args = length o filter is_arg; berghofe@11520: berghofe@11520: berghofe@11520: (**** theory data ****) berghofe@11520: berghofe@11520: (* data kind 'Pure/codegen' *) berghofe@11520: berghofe@11520: structure CodegenArgs = berghofe@11520: struct berghofe@11520: val name = "Pure/codegen"; berghofe@11520: type T = berghofe@11520: {codegens : (string * (theory -> (exn option * string) Graph.T -> string -> berghofe@11520: bool -> term -> ((exn option * string) Graph.T * Pretty.T) option)) list, berghofe@11520: consts : ((string * typ) * mixfix list) list, berghofe@11520: types : (string * string) list}; berghofe@11520: berghofe@11520: val empty = {codegens = [], consts = [], types = []}; berghofe@11520: val copy = I; berghofe@11520: val prep_ext = I; berghofe@11520: berghofe@11520: fun merge ({codegens = codegens1, consts = consts1, types = types1}, berghofe@11520: {codegens = codegens2, consts = consts2, types = types2}) = berghofe@11520: {codegens = rev (merge_alists (rev codegens1) (rev codegens2)), berghofe@11520: consts = merge_alists consts1 consts2, berghofe@11520: types = merge_alists types1 types2}; berghofe@11520: berghofe@11520: fun print sg (cs:T) = Pretty.writeln berghofe@11520: (Pretty.strs ("code generators:" :: map fst (#codegens cs))); berghofe@11520: end; berghofe@11520: berghofe@11520: structure CodegenData = TheoryDataFun(CodegenArgs); berghofe@11520: val print_codegens = CodegenData.print; berghofe@11520: berghofe@11520: berghofe@11520: (**** add new code generator to theory ****) berghofe@11520: berghofe@11520: fun add_codegen name f thy = berghofe@11520: let val {codegens, consts, types} = CodegenData.get thy berghofe@11520: in (case assoc (codegens, name) of berghofe@11520: None => CodegenData.put {codegens = (name, f)::codegens, berghofe@11520: consts = consts, types = types} thy berghofe@11520: | Some _ => error ("Code generator " ^ name ^ " already declared")) berghofe@11520: end; berghofe@11520: berghofe@11520: berghofe@11520: (**** associate constants with target language code ****) berghofe@11520: berghofe@11520: fun gen_assoc_consts prep_type xs thy = foldl (fn (thy, (s, tyopt, syn)) => berghofe@11520: let berghofe@11520: val sg = sign_of thy; berghofe@11520: val {codegens, consts, types} = CodegenData.get thy; berghofe@11520: val cname = Sign.intern_const sg s; berghofe@11520: in berghofe@11520: (case Sign.const_type sg cname of berghofe@11520: Some T => berghofe@11520: let val T' = (case tyopt of berghofe@11520: None => T berghofe@11520: | Some ty => berghofe@11520: let val U = prep_type sg ty berghofe@11520: in if Type.typ_instance (Sign.tsig_of sg, U, T) then U berghofe@11520: else error ("Illegal type constraint for constant " ^ cname) berghofe@11520: end) berghofe@11520: in (case assoc (consts, (cname, T')) of berghofe@11520: None => CodegenData.put {codegens = codegens, berghofe@11520: consts = ((cname, T'), syn) :: consts, types = types} thy berghofe@11520: | Some _ => error ("Constant " ^ cname ^ " already associated with code")) berghofe@11520: end berghofe@11520: | _ => error ("Not a constant: " ^ s)) berghofe@11520: end) (thy, xs); berghofe@11520: berghofe@11520: val assoc_consts_i = gen_assoc_consts (K I); berghofe@11520: val assoc_consts = gen_assoc_consts (fn sg => typ_of o read_ctyp sg); berghofe@11520: berghofe@11520: (**** associate types with target language types ****) berghofe@11520: berghofe@11520: fun assoc_types xs thy = foldl (fn (thy, (s, syn)) => berghofe@11520: let berghofe@11520: val {codegens, consts, types} = CodegenData.get thy; berghofe@11520: val tc = Sign.intern_tycon (sign_of thy) s berghofe@11520: in berghofe@11520: (case assoc (types, tc) of berghofe@11520: None => CodegenData.put {codegens = codegens, consts = consts, berghofe@11520: types = (tc, syn) :: types} thy berghofe@11520: | Some _ => error ("Type " ^ tc ^ " already associated with code")) berghofe@11520: end) (thy, xs); berghofe@11520: berghofe@11520: fun get_assoc_types thy = #types (CodegenData.get thy); berghofe@11520: berghofe@11520: berghofe@11520: (**** retrieve definition of constant ****) berghofe@11520: berghofe@11520: fun is_instance thy T1 T2 = berghofe@11520: Type.typ_instance (Sign.tsig_of (sign_of thy), T1, Type.varifyT T2); berghofe@11520: berghofe@11520: fun get_assoc_code thy s T = apsome snd (find_first (fn ((s', T'), _) => berghofe@11520: s = s' andalso is_instance thy T T') (#consts (CodegenData.get thy))); berghofe@11520: berghofe@11520: fun get_defn thy s T = berghofe@11520: let berghofe@11520: val axms = flat (map (Symtab.dest o #axioms o Theory.rep_theory) berghofe@11520: (thy :: Theory.ancestors_of thy)); berghofe@11520: val defs = mapfilter (fn (_, t) => berghofe@11520: (let berghofe@11520: val (lhs, rhs) = Logic.dest_equals t; berghofe@11520: val (c, args) = strip_comb lhs; berghofe@11520: val (s', T') = dest_Const c berghofe@11520: in if s=s' then Some (T', (args, rhs)) else None end) handle TERM _ => berghofe@11520: None) axms; berghofe@11520: val i = find_index (is_instance thy T o fst) defs berghofe@11520: in berghofe@11520: if i>=0 then Some (snd (nth_elem (i, defs)), berghofe@11520: if length defs = 1 then None else Some i) berghofe@11520: else None berghofe@11520: end; berghofe@11520: berghofe@11520: berghofe@11520: (**** make valid ML identifiers ****) berghofe@11520: berghofe@11520: fun gen_mk_id kind rename sg s = berghofe@11520: let berghofe@11520: val (xs as x::_) = explode (rename (space_implode "_" berghofe@11520: (NameSpace.unpack (Sign.cond_extern sg kind s)))); berghofe@11520: fun check_str [] = "" berghofe@11520: | check_str (x::xs) = berghofe@11520: (if Symbol.is_letter x orelse Symbol.is_digit x orelse x="_" then x berghofe@11520: else "_" ^ string_of_int (ord x)) ^ check_str xs berghofe@11520: in berghofe@11520: (if not (Symbol.is_letter x) then "id" else "") ^ check_str xs berghofe@11520: end; berghofe@11520: berghofe@11520: val mk_const_id = gen_mk_id Sign.constK I; berghofe@11520: val mk_type_id = gen_mk_id Sign.typeK berghofe@11520: (fn s => if s mem ThmDatabase.ml_reserved then s ^ "_type" else s); berghofe@11520: berghofe@11520: fun rename_term t = berghofe@11520: let berghofe@11520: val names = add_term_names (t, map (fst o fst o dest_Var) (term_vars t)); berghofe@11520: val clash = names inter ThmDatabase.ml_reserved; berghofe@11520: val ps = clash ~~ variantlist (clash, names); berghofe@11520: berghofe@11520: fun rename (Var ((a, i), T)) = Var ((if_none (assoc (ps, a)) a, i), T) berghofe@11520: | rename (Free (a, T)) = Free (if_none (assoc (ps, a)) a, T) berghofe@11520: | rename (Abs (s, T, t)) = Abs (s, T, rename t) berghofe@11520: | rename (t $ u) = rename t $ rename u berghofe@11520: | rename t = t; berghofe@11520: in berghofe@11520: rename t berghofe@11520: end; berghofe@11520: berghofe@11520: berghofe@11520: (**** invoke suitable code generator for term t ****) berghofe@11520: berghofe@11520: fun invoke_codegen thy gr dep brack t = (case get_first berghofe@11520: (fn (_, f) => f thy gr dep brack t) (#codegens (CodegenData.get thy)) of berghofe@11520: None => error ("Unable to generate code for term:\n" ^ berghofe@11520: Sign.string_of_term (sign_of thy) t ^ "\nrequired by:\n" ^ berghofe@11520: commas (Graph.all_succs gr [dep])) berghofe@11520: | Some x => x) berghofe@11520: berghofe@11520: berghofe@11520: (**** code generator for mixfix expressions ****) berghofe@11520: berghofe@11520: fun parens p = Pretty.block [Pretty.str "(", p, Pretty.str ")"]; berghofe@11520: berghofe@11520: fun pretty_fn [] p = [p] berghofe@11520: | pretty_fn (x::xs) p = Pretty.str ("fn " ^ x ^ " =>") :: berghofe@11520: Pretty.brk 1 :: pretty_fn xs p; berghofe@11520: berghofe@11520: fun pretty_mixfix [] [] _ = [] berghofe@11520: | pretty_mixfix (Arg :: ms) (p :: ps) qs = p :: pretty_mixfix ms ps qs berghofe@11520: | pretty_mixfix (Ignore :: ms) (p :: ps) qs = pretty_mixfix ms ps qs berghofe@11520: | pretty_mixfix (Pretty p :: ms) ps qs = p :: pretty_mixfix ms ps qs berghofe@11520: | pretty_mixfix (Term _ :: ms) ps (q :: qs) = q :: pretty_mixfix ms ps qs; berghofe@11520: berghofe@11520: berghofe@11520: (**** default code generator ****) berghofe@11520: berghofe@11520: fun eta_expand t ts i = berghofe@11520: let berghofe@11520: val (Ts, _) = strip_type (fastype_of t); berghofe@11520: val j = i - length ts berghofe@11520: in berghofe@11520: foldr (fn (T, t) => Abs ("x", T, t)) berghofe@11520: (take (j, Ts), list_comb (t, ts @ map Bound (j-1 downto 0))) berghofe@11520: end; berghofe@11520: berghofe@11520: fun mk_app _ p [] = p berghofe@11520: | mk_app brack p ps = if brack then berghofe@11520: Pretty.block (Pretty.str "(" :: berghofe@11520: separate (Pretty.brk 1) (p :: ps) @ [Pretty.str ")"]) berghofe@11520: else Pretty.block (separate (Pretty.brk 1) (p :: ps)); berghofe@11520: berghofe@11520: fun new_names t xs = variantlist (xs, berghofe@11520: map (fst o fst o dest_Var) (term_vars t) union berghofe@11520: add_term_names (t, ThmDatabase.ml_reserved)); berghofe@11520: berghofe@11520: fun new_name t x = hd (new_names t [x]); berghofe@11520: berghofe@11520: fun default_codegen thy gr dep brack t = berghofe@11520: let berghofe@11520: val (u, ts) = strip_comb t; berghofe@11520: fun mapcode brack' gr ts = foldl_map berghofe@11520: (fn (gr, t) => invoke_codegen thy gr dep brack' t) (gr, ts) berghofe@11520: berghofe@11520: in (case u of berghofe@11520: Var ((s, i), _) => berghofe@11520: let val (gr', ps) = mapcode true gr ts berghofe@11520: in Some (gr', mk_app brack (Pretty.str (s ^ berghofe@11520: (if i=0 then "" else string_of_int i))) ps) berghofe@11520: end berghofe@11520: berghofe@11520: | Free (s, _) => berghofe@11520: let val (gr', ps) = mapcode true gr ts berghofe@11520: in Some (gr', mk_app brack (Pretty.str s) ps) end berghofe@11520: berghofe@11520: | Const (s, T) => berghofe@11520: (case get_assoc_code thy s T of berghofe@11520: Some ms => berghofe@11520: let val i = num_args ms berghofe@11520: in if length ts < i then berghofe@11520: default_codegen thy gr dep brack (eta_expand u ts i) berghofe@11520: else berghofe@11520: let berghofe@11520: val ts1 = take (i, ts); berghofe@11520: val ts2 = drop (i, ts); berghofe@11520: val (gr1, ps1) = mapcode false gr ts1; berghofe@11520: val (gr2, ps2) = mapcode true gr1 ts2; berghofe@11520: val (gr3, ps3) = mapcode false gr2 (terms_of ms); berghofe@11520: in berghofe@11520: Some (gr3, mk_app brack (Pretty.block (pretty_mixfix ms ps1 ps3)) ps2) berghofe@11520: end berghofe@11520: end berghofe@11520: | None => (case get_defn thy s T of berghofe@11520: None => None berghofe@11520: | Some ((args, rhs), k) => berghofe@11520: let berghofe@11520: val id = mk_const_id (sign_of thy) s ^ (case k of berghofe@11520: None => "" | Some i => "_def" ^ string_of_int i); berghofe@11520: val (gr', ps) = mapcode true gr ts; berghofe@11520: in berghofe@11520: Some (Graph.add_edge (id, dep) gr' handle Graph.UNDEF _ => berghofe@11520: let berghofe@11520: val _ = message ("expanding definition of " ^ s); berghofe@11520: val (Ts, _) = strip_type T; berghofe@11520: val (args', rhs') = berghofe@11520: if not (null args) orelse null Ts then (args, rhs) else berghofe@11520: let val v = Free (new_name rhs "x", hd Ts) berghofe@11520: in ([v], betapply (rhs, v)) end; berghofe@11520: val (gr1, p) = invoke_codegen thy (Graph.add_edge (id, dep) berghofe@11520: (Graph.new_node (id, (None, "")) gr')) id false rhs'; berghofe@11520: val (gr2, xs) = mapcode false gr1 args'; berghofe@11520: in Graph.map_node id (K (None, Pretty.string_of (Pretty.block berghofe@11520: (Pretty.str (if null args' then "val " else "fun ") :: berghofe@11520: separate (Pretty.brk 1) (Pretty.str id :: xs) @ berghofe@11520: [Pretty.str " =", Pretty.brk 1, p, Pretty.str ";"])) ^ "\n\n")) gr2 berghofe@11520: end, mk_app brack (Pretty.str id) ps) berghofe@11520: end)) berghofe@11520: berghofe@11520: | Abs _ => berghofe@11520: let berghofe@11520: val (bs, Ts) = ListPair.unzip (strip_abs_vars u); berghofe@11520: val t = strip_abs_body u berghofe@11520: val bs' = new_names t bs; berghofe@11520: val (gr1, ps) = mapcode true gr ts; berghofe@11520: val (gr2, p) = invoke_codegen thy gr1 dep false berghofe@11520: (subst_bounds (map Free (rev (bs' ~~ Ts)), t)); berghofe@11520: in berghofe@11520: Some (gr2, mk_app brack (Pretty.block (Pretty.str "(" :: pretty_fn bs' p @ berghofe@11520: [Pretty.str ")"])) ps) berghofe@11520: end berghofe@11520: berghofe@11520: | _ => None) berghofe@11520: end; berghofe@11520: berghofe@11520: berghofe@11520: fun output_code gr xs = implode (map (snd o Graph.get_node gr) berghofe@11520: (rev (Graph.all_preds gr xs))); berghofe@11520: berghofe@11520: fun gen_generate_code prep_term thy = Pretty.setmp_margin 80 (fn xs => berghofe@11520: let berghofe@11520: val sg = sign_of thy; berghofe@11520: val gr = Graph.new_node ("", (None, "")) Graph.empty; berghofe@11520: val (gr', ps) = foldl_map (fn (gr, (s, t)) => apsnd (pair s) berghofe@11520: (invoke_codegen thy gr "" false t)) (gr, map (apsnd (prep_term sg)) xs) berghofe@11520: in berghofe@11520: "structure Generated =\nstruct\n\n" ^ berghofe@11520: output_code gr' [""] ^ berghofe@11520: space_implode "\n\n" (map (fn (s', p) => Pretty.string_of (Pretty.block berghofe@11520: [Pretty.str ("val " ^ s' ^ " ="), Pretty.brk 1, p, Pretty.str ";"])) ps) ^ berghofe@11520: "\n\nend;\n\nopen Generated;\n" berghofe@11520: end); berghofe@11520: berghofe@11520: val generate_code_i = gen_generate_code (K I); berghofe@11520: val generate_code = gen_generate_code berghofe@11520: (fn sg => term_of o read_cterm sg o rpair TypeInfer.logicT); berghofe@11520: berghofe@11520: fun parse_mixfix rd s = berghofe@11520: (case Scan.finite Symbol.stopper (Scan.repeat berghofe@11520: ( $$ "_" >> K Arg berghofe@11520: || $$ "?" >> K Ignore berghofe@11520: || $$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length) berghofe@11520: || $$ "{" |-- $$ "*" |-- Scan.repeat1 berghofe@11520: ( $$ "'" |-- Scan.one Symbol.not_eof berghofe@11520: || Scan.unless ($$ "*" -- $$ "}") (Scan.one Symbol.not_eof)) --| berghofe@11520: $$ "*" --| $$ "}" >> (Term o rd o implode) berghofe@11520: || Scan.repeat1 berghofe@11520: ( $$ "'" |-- Scan.one Symbol.not_eof berghofe@11520: || Scan.unless ($$ "_" || $$ "?" || $$ "/" || $$ "{" |-- $$ "*") berghofe@11520: (Scan.one Symbol.not_eof)) >> (Pretty o Pretty.str o implode))) berghofe@11520: (Symbol.explode s) of berghofe@11520: (p, []) => p berghofe@11520: | _ => error ("Malformed annotation: " ^ quote s)); berghofe@11520: berghofe@11520: structure P = OuterParse; berghofe@11520: berghofe@11520: val assoc_typeP = berghofe@11520: OuterSyntax.command "types_code" berghofe@11520: "associate types with target language types" berghofe@11520: OuterSyntax.Keyword.thy_decl berghofe@11520: (Scan.repeat1 (P.xname --| P.$$$ "(" -- P.string --| P.$$$ ")") >> berghofe@11520: (Toplevel.theory o assoc_types)); berghofe@11520: berghofe@11520: val assoc_constP = berghofe@11520: OuterSyntax.command "consts_code" berghofe@11520: "associate constants with target language code" berghofe@11520: OuterSyntax.Keyword.thy_decl berghofe@11520: (Scan.repeat1 berghofe@11520: (P.xname -- (Scan.option (P.$$$ "::" |-- P.string)) --| berghofe@11520: P.$$$ "(" -- P.string --| P.$$$ ")") >> berghofe@11520: (fn xs => Toplevel.theory (fn thy => assoc_consts berghofe@11520: (map (fn ((name, optype), mfx) => (name, optype, parse_mixfix berghofe@11520: (term_of o read_cterm (sign_of thy) o rpair TypeInfer.logicT) mfx)) berghofe@11520: xs) thy))); berghofe@11520: berghofe@11520: val generate_codeP = berghofe@11520: OuterSyntax.command "generate_code" "generates code for terms" berghofe@11520: OuterSyntax.Keyword.thy_decl berghofe@11520: (Scan.option (P.$$$ "(" |-- P.string --| P.$$$ ")") -- berghofe@11520: Scan.repeat1 (P.short_ident --| P.$$$ "=" -- P.string) >> berghofe@11520: (fn (opt_fname, xs) => Toplevel.theory (fn thy => berghofe@11520: ((case opt_fname of berghofe@11520: None => use_text Context.ml_output false berghofe@11520: | Some fname => File.write (Path.unpack fname)) berghofe@11520: (generate_code thy xs); thy)))); berghofe@11520: berghofe@11520: val parsers = [assoc_typeP, assoc_constP, generate_codeP]; berghofe@11520: berghofe@11520: val setup = [CodegenData.init, add_codegen "default" default_codegen]; berghofe@11520: berghofe@11520: end; berghofe@11520: berghofe@11520: OuterSyntax.add_parsers Codegen.parsers;