haftmann@24219: (* Title: Tools/code/code_target.ML haftmann@24219: ID: $Id$ haftmann@24219: Author: Florian Haftmann, TU Muenchen haftmann@24219: haftmann@24219: Serializer from intermediate language ("Thin-gol") haftmann@24219: to target languages (like SML or Haskell). haftmann@24219: *) haftmann@24219: haftmann@24219: signature CODE_TARGET = haftmann@24219: sig haftmann@24219: include BASIC_CODE_THINGOL; haftmann@24219: haftmann@24219: val add_syntax_class: string -> class haftmann@24423: -> (string * (string * string) list) option -> theory -> theory; haftmann@24219: val add_syntax_inst: string -> string * class -> bool -> theory -> theory; haftmann@24219: val add_syntax_tycoP: string -> string -> OuterParse.token list haftmann@24219: -> (theory -> theory) * OuterParse.token list; haftmann@24219: val add_syntax_constP: string -> string -> OuterParse.token list haftmann@24219: -> (theory -> theory) * OuterParse.token list; haftmann@24219: haftmann@24219: val add_undefined: string -> string -> string -> theory -> theory; haftmann@24219: val add_pretty_list: string -> string -> string -> theory -> theory; haftmann@24219: val add_pretty_list_string: string -> string -> string haftmann@24219: -> string -> string list -> theory -> theory; haftmann@24219: val add_pretty_char: string -> string -> string list -> theory -> theory huffman@26086: val add_pretty_numeral: string -> bool -> bool -> string -> string -> string haftmann@24219: -> string -> string -> theory -> theory; haftmann@24992: val add_pretty_message: string -> string -> string list -> string haftmann@24219: -> string -> string -> theory -> theory; haftmann@24219: haftmann@27103: val allow_abort: string -> theory -> theory; haftmann@24841: haftmann@27000: type serialization; haftmann@24219: type serializer; haftmann@27103: val add_target: string * serializer -> theory -> theory; haftmann@27550: val extend_target: string * (string * (CodeThingol.program -> CodeThingol.program)) haftmann@27550: -> theory -> theory; haftmann@27103: val assert_target: theory -> string -> string; haftmann@27103: val serialize: theory -> string -> string option -> Args.T list haftmann@27103: -> CodeThingol.program -> string list -> serialization; haftmann@27000: val compile: serialization -> unit; haftmann@27304: val export: serialization -> unit; haftmann@27000: val file: Path.T -> serialization -> unit; haftmann@27304: val string: string list -> serialization -> string; haftmann@27103: haftmann@27304: val code_of: theory -> string -> string -> string list -> string list -> string; haftmann@27103: val eval_conv: string * (unit -> thm) option ref haftmann@27103: -> theory -> cterm -> string list -> thm; haftmann@27103: val eval_term: string * (unit -> 'a) option ref haftmann@27103: -> theory -> term -> string list -> 'a; haftmann@27103: val shell_command: string (*theory name*) -> string (*cg expr*) -> unit; haftmann@24219: haftmann@24219: val setup: theory -> theory; haftmann@27103: val code_width: int ref; haftmann@27437: haftmann@27437: val ml_code_of: theory -> CodeThingol.program -> string list -> string * string list; haftmann@24219: end; haftmann@24219: haftmann@24219: structure CodeTarget : CODE_TARGET = haftmann@24219: struct haftmann@24219: haftmann@24219: open BasicCodeThingol; haftmann@24219: haftmann@24219: (** basics **) haftmann@24219: haftmann@24219: infixr 5 @@; haftmann@24219: infixr 5 @|; haftmann@24219: fun x @@ y = [x, y]; haftmann@24219: fun xs @| y = xs @ [y]; wenzelm@24634: val str = PrintMode.setmp [] Pretty.str; haftmann@24219: val concat = Pretty.block o Pretty.breaks; haftmann@24219: val brackets = Pretty.enclose "(" ")" o Pretty.breaks; haftmann@24219: fun semicolon ps = Pretty.block [concat ps, str ";"]; haftmann@25771: fun enum_default default sep opn cls [] = str default haftmann@25771: | enum_default default sep opn cls xs = Pretty.enum sep opn cls xs; haftmann@24219: haftmann@27304: datatype destination = Compile | Export | File of Path.T | String of string list; haftmann@27436: type serialization = destination -> (string * string list) option; haftmann@27014: haftmann@27103: val code_width = ref 80; (*FIXME after Pretty module no longer depends on print mode*) haftmann@27103: fun code_setmp f = PrintMode.setmp [] (Pretty.setmp_margin (!code_width) f); haftmann@27103: fun code_of_pretty p = code_setmp Pretty.string_of p ^ "\n"; haftmann@27103: fun code_writeln p = Pretty.setmp_margin (!code_width) Pretty.writeln p; haftmann@26998: haftmann@27103: (*FIXME why another code_setmp?*) haftmann@27103: fun compile f = (code_setmp f Compile; ()); haftmann@27304: fun export f = (code_setmp f Export; ()); haftmann@27103: fun file p f = (code_setmp f (File p); ()); haftmann@27436: fun string cs f = fst (the (code_setmp f (String cs))); haftmann@27304: haftmann@27304: fun stmt_names_of_destination (String stmts) = stmts haftmann@27304: | stmt_names_of_destination _ = []; haftmann@27014: haftmann@24219: haftmann@27000: (** generic syntax **) haftmann@24219: haftmann@24219: datatype lrx = L | R | X; haftmann@24219: haftmann@24219: datatype fixity = haftmann@24219: BR haftmann@24219: | NOBR haftmann@24219: | INFX of (int * lrx); haftmann@24219: haftmann@24219: val APP = INFX (~1, L); haftmann@24219: haftmann@26010: fun fixity_lrx L L = false haftmann@26010: | fixity_lrx R R = false haftmann@26010: | fixity_lrx _ _ = true; haftmann@24219: haftmann@27024: fun fixity NOBR _ = false haftmann@27024: | fixity _ NOBR = false haftmann@26010: | fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = haftmann@24219: pr < pr_ctxt haftmann@24219: orelse pr = pr_ctxt haftmann@26010: andalso fixity_lrx lr lr_ctxt haftmann@24219: orelse pr_ctxt = ~1 haftmann@27024: | fixity BR (INFX _) = false haftmann@26010: | fixity _ _ = true; haftmann@24219: haftmann@24219: fun gen_brackify _ [p] = p haftmann@24219: | gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps haftmann@24219: | gen_brackify false (ps as _::_) = Pretty.block ps; haftmann@24219: haftmann@27024: fun brackify fxy_ctxt = haftmann@27024: gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks; haftmann@24219: haftmann@27024: fun brackify_infix infx fxy_ctxt = haftmann@27024: gen_brackify (fixity (INFX infx) fxy_ctxt) o Pretty.breaks; haftmann@24219: haftmann@24219: type class_syntax = string * (string -> string option); haftmann@24219: type typ_syntax = int * ((fixity -> itype -> Pretty.T) haftmann@24219: -> fixity -> itype list -> Pretty.T); haftmann@24219: type term_syntax = int * ((CodeName.var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@27304: -> thm -> bool -> CodeName.var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T); haftmann@24219: haftmann@27550: datatype name_syntax_table = NameSyntaxTable of { haftmann@27550: class: class_syntax Symtab.table, haftmann@27550: inst: unit Symtab.table, haftmann@27550: tyco: typ_syntax Symtab.table, haftmann@27550: const: term_syntax Symtab.table haftmann@27550: }; haftmann@27550: haftmann@24219: haftmann@27000: (** theory data **) haftmann@27000: haftmann@27000: val target_SML = "SML"; haftmann@27000: val target_OCaml = "OCaml"; haftmann@27000: val target_Haskell = "Haskell"; haftmann@27000: haftmann@27024: fun mk_name_syntax_table ((class, inst), (tyco, const)) = haftmann@27024: NameSyntaxTable { class = class, inst = inst, tyco = tyco, const = const }; haftmann@27024: fun map_name_syntax_table f (NameSyntaxTable { class, inst, tyco, const }) = haftmann@27024: mk_name_syntax_table (f ((class, inst), (tyco, const))); haftmann@27024: fun merge_name_syntax_table (NameSyntaxTable { class = class1, inst = inst1, tyco = tyco1, const = const1 }, haftmann@27024: NameSyntaxTable { class = class2, inst = inst2, tyco = tyco2, const = const2 }) = haftmann@27024: mk_name_syntax_table ( haftmann@27000: (Symtab.join (K snd) (class1, class2), haftmann@27000: Symtab.join (K snd) (inst1, inst2)), haftmann@27000: (Symtab.join (K snd) (tyco1, tyco2), haftmann@27000: Symtab.join (K snd) (const1, const2)) haftmann@27000: ); haftmann@27000: haftmann@27304: type serializer = haftmann@27024: string option (*module name*) haftmann@27024: -> Args.T list (*arguments*) haftmann@27024: -> (string -> string) (*labelled_name*) haftmann@27024: -> string list (*reserved symbols*) haftmann@27024: -> (string * Pretty.T) list (*includes*) haftmann@27024: -> (string -> string option) (*module aliasses*) haftmann@27000: -> (string -> class_syntax option) haftmann@27000: -> (string -> typ_syntax option) haftmann@27000: -> (string -> term_syntax option) haftmann@27103: -> CodeThingol.program haftmann@27024: -> string list (*selected statements*) haftmann@27024: -> serialization; haftmann@24219: haftmann@27550: datatype serializer_entry = Serializer of serializer haftmann@27550: | Extends of string * (CodeThingol.program -> CodeThingol.program); haftmann@27550: haftmann@27000: datatype target = Target of { haftmann@27000: serial: serial, haftmann@27550: serializer: serializer_entry, haftmann@27000: reserved: string list, haftmann@27000: includes: Pretty.T Symtab.table, haftmann@27024: name_syntax_table: name_syntax_table, haftmann@27000: module_alias: string Symtab.table haftmann@27000: }; haftmann@24219: haftmann@27024: fun mk_target ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))) = haftmann@27000: Target { serial = serial, serializer = serializer, reserved = reserved, haftmann@27024: includes = includes, name_syntax_table = name_syntax_table, module_alias = module_alias }; haftmann@27024: fun map_target f ( Target { serial, serializer, reserved, includes, name_syntax_table, module_alias } ) = haftmann@27024: mk_target (f ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias)))); haftmann@27550: fun merge_target strict target (Target { serial = serial1, serializer = serializer, haftmann@27000: reserved = reserved1, includes = includes1, haftmann@27024: name_syntax_table = name_syntax_table1, module_alias = module_alias1 }, haftmann@27000: Target { serial = serial2, serializer = _, haftmann@27000: reserved = reserved2, includes = includes2, haftmann@27024: name_syntax_table = name_syntax_table2, module_alias = module_alias2 }) = haftmann@27550: if serial1 = serial2 orelse not strict then haftmann@27000: mk_target ((serial1, serializer), haftmann@27000: ((merge (op =) (reserved1, reserved2), Symtab.merge (op =) (includes1, includes2)), haftmann@27024: (merge_name_syntax_table (name_syntax_table1, name_syntax_table2), haftmann@27000: Symtab.join (K snd) (module_alias1, module_alias2)) haftmann@27000: )) haftmann@27000: else haftmann@27000: error ("Incompatible serializers: " ^ quote target); haftmann@27000: haftmann@27000: structure CodeTargetData = TheoryDataFun haftmann@27000: ( haftmann@27000: type T = target Symtab.table * string list; haftmann@27000: val empty = (Symtab.empty, []); haftmann@27000: val copy = I; haftmann@27000: val extend = I; haftmann@27000: fun merge _ ((target1, exc1) : T, (target2, exc2)) = haftmann@27550: (Symtab.join (merge_target true) (target1, target2), Library.merge (op =) (exc1, exc2)); haftmann@27000: ); haftmann@27000: haftmann@27000: fun the_serializer (Target { serializer, ... }) = serializer; haftmann@27000: fun the_reserved (Target { reserved, ... }) = reserved; haftmann@27000: fun the_includes (Target { includes, ... }) = includes; haftmann@27024: fun the_name_syntax (Target { name_syntax_table = NameSyntaxTable x, ... }) = x; haftmann@27000: fun the_module_alias (Target { module_alias , ... }) = module_alias; haftmann@27000: haftmann@27103: val abort_allowed = snd o CodeTargetData.get; haftmann@27103: haftmann@27103: fun assert_target thy target = haftmann@27000: case Symtab.lookup (fst (CodeTargetData.get thy)) target haftmann@27000: of SOME data => target haftmann@27000: | NONE => error ("Unknown code target language: " ^ quote target); haftmann@27000: haftmann@27550: fun put_target (target, seri) thy = haftmann@24219: let haftmann@27550: val defined_target = is_some o Symtab.lookup (fst (CodeTargetData.get thy)); haftmann@27550: val _ = case seri haftmann@27550: of Extends (super, _) => if defined_target super then () haftmann@27710: else error ("Unknown code target language: " ^ quote super) haftmann@27550: | _ => (); haftmann@27550: val _ = if defined_target target haftmann@27550: then warning ("Overwriting existing target " ^ quote target) haftmann@27550: else (); haftmann@24219: in haftmann@27000: thy haftmann@27000: |> (CodeTargetData.map o apfst oo Symtab.map_default) haftmann@27000: (target, mk_target ((serial (), seri), (([], Symtab.empty), haftmann@27024: (mk_name_syntax_table ((Symtab.empty, Symtab.empty), (Symtab.empty, Symtab.empty)), haftmann@27000: Symtab.empty)))) haftmann@27000: ((map_target o apfst o apsnd o K) seri) haftmann@24219: end; haftmann@24219: haftmann@27550: fun add_target (target, seri) = put_target (target, Serializer seri); haftmann@27550: fun extend_target (target, (super, modify)) = haftmann@27550: put_target (target, Extends (super, modify)); haftmann@27550: haftmann@27103: fun map_target_data target f thy = haftmann@24219: let haftmann@27103: val _ = assert_target thy target; haftmann@24219: in haftmann@27000: thy haftmann@27000: |> (CodeTargetData.map o apfst o Symtab.map_entry target o map_target) f haftmann@24219: end; haftmann@24219: haftmann@27103: fun map_reserved target = haftmann@27103: map_target_data target o apsnd o apfst o apfst; haftmann@27103: fun map_includes target = haftmann@27103: map_target_data target o apsnd o apfst o apsnd; haftmann@27024: fun map_name_syntax target = haftmann@27103: map_target_data target o apsnd o apsnd o apfst o map_name_syntax_table; haftmann@27000: fun map_module_alias target = haftmann@27103: map_target_data target o apsnd o apsnd o apsnd; haftmann@27000: haftmann@27550: fun invoke_serializer thy modify abortable serializer reserved includes haftmann@27103: module_alias class inst tyco const module args program1 cs1 = haftmann@24219: let haftmann@27550: val program2 = modify program1; haftmann@27103: val hidden = Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const; haftmann@27103: val cs2 = subtract (op =) hidden cs1; haftmann@27550: val program3 = Graph.subgraph (not o member (op =) hidden) program2; haftmann@27103: val all_cs = Graph.all_succs program2 cs2; haftmann@27550: val program4 = Graph.subgraph (member (op =) all_cs) program3; haftmann@27103: val empty_funs = filter_out (member (op =) abortable) haftmann@27103: (CodeThingol.empty_funs program3); haftmann@27103: val _ = if null empty_funs then () else error ("No defining equations for " haftmann@27103: ^ commas (map (CodeName.labelled_name thy) empty_funs)); haftmann@27103: in haftmann@27103: serializer module args (CodeName.labelled_name thy) reserved includes haftmann@27103: (Symtab.lookup module_alias) (Symtab.lookup class) haftmann@27103: (Symtab.lookup tyco) (Symtab.lookup const) haftmann@27550: program4 cs2 haftmann@27103: end; haftmann@27103: haftmann@27103: fun mount_serializer thy alt_serializer target = haftmann@27103: let haftmann@27103: val (targets, abortable) = CodeTargetData.get thy; haftmann@27550: fun collapse_hierarchy target = haftmann@27550: let haftmann@27550: val data = case Symtab.lookup targets target haftmann@27550: of SOME data => data haftmann@27550: | NONE => error ("Unknown code target language: " ^ quote target); haftmann@27550: in case the_serializer data haftmann@27550: of Serializer _ => (I, data) haftmann@27550: | Extends (super, modify) => let haftmann@27550: val (modify', data') = collapse_hierarchy super haftmann@27550: in (modify' #> modify, merge_target false target (data', data)) end haftmann@27550: end; haftmann@27550: val (modify, data) = collapse_hierarchy target; haftmann@27550: val serializer = the_default (case the_serializer data haftmann@27550: of Serializer seri => seri) alt_serializer; haftmann@27000: val reserved = the_reserved data; haftmann@27000: val includes = Symtab.dest (the_includes data); haftmann@27103: val module_alias = the_module_alias data; haftmann@27024: val { class, inst, tyco, const } = the_name_syntax data; haftmann@27103: in haftmann@27550: invoke_serializer thy modify abortable serializer reserved haftmann@27103: includes module_alias class inst tyco const haftmann@24219: end; haftmann@24219: haftmann@27103: fun serialize thy = mount_serializer thy NONE; haftmann@27000: haftmann@24219: fun parse_args f args = haftmann@24219: case Scan.read Args.stopper f args haftmann@24219: of SOME x => x haftmann@24219: | NONE => error "Bad serializer arguments"; haftmann@24219: haftmann@24219: haftmann@27103: (** generic code combinators **) haftmann@24219: haftmann@24219: (* list, char, string, numeral and monad abstract syntax transformations *) haftmann@24219: haftmann@27304: fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm thm); haftmann@27304: haftmann@24219: fun implode_list c_nil c_cons t = haftmann@24219: let haftmann@24219: fun dest_cons (IConst (c, _) `$ t1 `$ t2) = haftmann@24219: if c = c_cons haftmann@24219: then SOME (t1, t2) haftmann@24219: else NONE haftmann@24219: | dest_cons _ = NONE; haftmann@24219: val (ts, t') = CodeThingol.unfoldr dest_cons t; haftmann@24219: in case t' haftmann@24219: of IConst (c, _) => if c = c_nil then SOME ts else NONE haftmann@24219: | _ => NONE haftmann@24219: end; haftmann@24219: haftmann@24219: fun decode_char c_nibbles (IConst (c1, _), IConst (c2, _)) = haftmann@24219: let haftmann@24219: fun idx c = find_index (curry (op =) c) c_nibbles; haftmann@24219: fun decode ~1 _ = NONE haftmann@24219: | decode _ ~1 = NONE haftmann@24219: | decode n m = SOME (chr (n * 16 + m)); haftmann@24219: in decode (idx c1) (idx c2) end haftmann@24219: | decode_char _ _ = NONE; haftmann@24219: haftmann@24219: fun implode_string c_char c_nibbles mk_char mk_string ts = haftmann@24219: let haftmann@24219: fun implode_char (IConst (c, _) `$ t1 `$ t2) = haftmann@24219: if c = c_char then decode_char c_nibbles (t1, t2) else NONE haftmann@24219: | implode_char _ = NONE; haftmann@24219: val ts' = map implode_char ts; haftmann@24219: in if forall is_some ts' haftmann@24219: then (SOME o str o mk_string o implode o map_filter I) ts' haftmann@24219: else NONE haftmann@24219: end; haftmann@24219: haftmann@27304: fun implode_numeral thm negative c_pls c_min c_bit0 c_bit1 = haftmann@24219: let haftmann@25936: fun dest_bit (IConst (c, _)) = if c = c_bit0 then 0 haftmann@25936: else if c = c_bit1 then 1 haftmann@27304: else nerror thm "Illegal numeral expression: illegal bit" haftmann@27304: | dest_bit _ = nerror thm "Illegal numeral expression: illegal bit"; wenzelm@24630: fun dest_numeral (IConst (c, _)) = if c = c_pls then SOME 0 haftmann@25936: else if c = c_min then haftmann@25936: if negative then SOME ~1 else NONE haftmann@27304: else nerror thm "Illegal numeral expression: illegal leading digit" huffman@26086: | dest_numeral (t1 `$ t2) = huffman@26086: let val (n, b) = (dest_numeral t2, dest_bit t1) huffman@26086: in case n of SOME n => SOME (2 * n + b) | NONE => NONE end haftmann@27304: | dest_numeral _ = nerror thm "Illegal numeral expression: illegal term"; haftmann@25936: in dest_numeral #> the_default 0 end; haftmann@24219: haftmann@24992: fun implode_monad c_bind t = haftmann@24219: let haftmann@24219: fun dest_monad (IConst (c, _) `$ t1 `$ t2) = haftmann@24992: if c = c_bind haftmann@24219: then case CodeThingol.split_abs t2 haftmann@24992: of SOME (((v, pat), ty), t') => haftmann@24992: SOME ((SOME (((SOME v, pat), ty), true), t1), t') haftmann@24219: | NONE => NONE haftmann@24219: else NONE haftmann@24219: | dest_monad t = case CodeThingol.split_let t haftmann@24992: of SOME (((pat, ty), tbind), t') => haftmann@24992: SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t') haftmann@24219: | NONE => NONE; haftmann@24219: in CodeThingol.unfoldr dest_monad t end; haftmann@24219: haftmann@24219: haftmann@27103: (* applications and bindings *) haftmann@27103: haftmann@27304: fun gen_pr_app pr_app pr_term syntax_const is_cons thm pat haftmann@27304: vars fxy (app as ((c, (_, tys)), ts)) = haftmann@27103: case syntax_const c haftmann@27304: of NONE => if pat andalso not (is_cons c) then haftmann@27710: nerror thm "Non-constructor in pattern" haftmann@27304: else brackify fxy (pr_app thm pat vars app) haftmann@27103: | SOME (i, pr) => haftmann@27103: let haftmann@27103: val k = if i < 0 then length tys else i; haftmann@27304: fun pr' fxy ts = pr (pr_term thm pat) thm pat vars fxy (ts ~~ curry Library.take k tys); haftmann@27103: in if k = length ts haftmann@27103: then pr' fxy ts haftmann@27103: else if k < length ts haftmann@27103: then case chop k ts of (ts1, ts2) => haftmann@27304: brackify fxy (pr' APP ts1 :: map (pr_term thm pat vars BR) ts2) haftmann@27710: else pr_term thm pat vars fxy (CodeThingol.eta_expand k app) haftmann@27103: end; haftmann@27103: haftmann@27304: fun gen_pr_bind pr_bind pr_term thm (fxy : fixity) ((v, pat), ty : itype) vars = haftmann@27103: let haftmann@27103: val vs = case pat haftmann@27103: of SOME pat => CodeThingol.fold_varnames (insert (op =)) pat [] haftmann@27103: | NONE => []; haftmann@27103: val vars' = CodeName.intro_vars (the_list v) vars; haftmann@27103: val vars'' = CodeName.intro_vars vs vars'; haftmann@27103: val v' = Option.map (CodeName.lookup_var vars') v; haftmann@27304: val pat' = Option.map (pr_term thm true vars'' fxy) pat; haftmann@27103: in (pr_bind ((v', pat'), ty), vars'') end; haftmann@27103: haftmann@27103: haftmann@27000: (* name auxiliary *) haftmann@24219: haftmann@24219: val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode; haftmann@24219: val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode; haftmann@24219: haftmann@24219: val dest_name = haftmann@24219: apfst NameSpace.implode o split_last o fst o split_last o NameSpace.explode; haftmann@24219: haftmann@27103: fun mk_name_module reserved_names module_prefix module_alias program = haftmann@24219: let haftmann@27103: fun mk_alias name = case module_alias name haftmann@27103: of SOME name' => name' haftmann@27103: | NONE => name haftmann@27103: |> NameSpace.explode haftmann@27103: |> map (fn name => (the_single o fst) (Name.variants [name] reserved_names)) haftmann@27103: |> NameSpace.implode; haftmann@27103: fun mk_prefix name = case module_prefix haftmann@27103: of SOME module_prefix => NameSpace.append module_prefix name haftmann@27103: | NONE => name; haftmann@24219: val tab = haftmann@24219: Symtab.empty haftmann@24219: |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name)) haftmann@24219: o fst o dest_name o fst) haftmann@27103: program haftmann@27103: in the o Symtab.lookup tab end; haftmann@24219: haftmann@24219: haftmann@24219: haftmann@24219: (** SML/OCaml serializer **) haftmann@24219: haftmann@27103: datatype ml_stmt = haftmann@27304: MLFuns of (string * (typscheme * ((iterm list * iterm) * thm) list)) list haftmann@24219: | MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list haftmann@24811: | MLClass of string * (vname * ((class * string) list * (string * itype) list)) haftmann@24219: | MLClassinst of string * ((class * (string * (vname * sort) list)) haftmann@24219: * ((class * (string * (string * dict list list))) list haftmann@27304: * ((string * const) * thm) list)); haftmann@27304: haftmann@27304: fun stmt_names_of (MLFuns fs) = map fst fs haftmann@27304: | stmt_names_of (MLDatas ds) = map fst ds haftmann@27304: | stmt_names_of (MLClass (c, _)) = [c] haftmann@27304: | stmt_names_of (MLClassinst (i, _)) = [i]; haftmann@24219: haftmann@27103: fun pr_sml_stmt syntax_tyco syntax_const labelled_name reserved_names deresolve is_cons = haftmann@24219: let haftmann@24992: val pr_label_classrel = translate_string (fn "." => "__" | c => c) haftmann@24992: o NameSpace.qualifier; haftmann@24841: val pr_label_classparam = NameSpace.base o NameSpace.qualifier; haftmann@24219: fun pr_dicts fxy ds = haftmann@24219: let haftmann@24219: fun pr_dictvar (v, (_, 1)) = first_upper v ^ "_" haftmann@24219: | pr_dictvar (v, (i, _)) = first_upper v ^ string_of_int (i+1) ^ "_"; haftmann@24219: fun pr_proj [] p = haftmann@24219: p haftmann@24219: | pr_proj [p'] p = haftmann@24219: brackets [p', p] haftmann@24219: | pr_proj (ps as _ :: _) p = haftmann@24219: brackets [Pretty.enum " o" "(" ")" ps, p]; haftmann@27103: fun pr_dict fxy (DictConst (inst, dss)) = haftmann@27103: brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss) haftmann@27103: | pr_dict fxy (DictVar (classrels, v)) = haftmann@27103: pr_proj (map (str o deresolve) classrels) ((str o pr_dictvar) v) haftmann@24219: in case ds haftmann@24219: of [] => str "()" haftmann@27103: | [d] => pr_dict fxy d haftmann@27103: | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds haftmann@24219: end; haftmann@27103: fun pr_tyvar_dicts vs = haftmann@24219: vs haftmann@24992: |> map (fn (v, sort) => map_index (fn (i, _) => haftmann@24992: DictVar ([], (v, (i, length sort)))) sort) haftmann@24219: |> map (pr_dicts BR); haftmann@24219: fun pr_tycoexpr fxy (tyco, tys) = haftmann@24219: let haftmann@27103: val tyco' = (str o deresolve) tyco haftmann@24219: in case map (pr_typ BR) tys haftmann@24219: of [] => tyco' haftmann@24219: | [p] => Pretty.block [p, Pretty.brk 1, tyco'] haftmann@24219: | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco'] haftmann@24219: end haftmann@27304: and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco haftmann@27304: of NONE => pr_tycoexpr fxy (tyco, tys) haftmann@27304: | SOME (i, pr) => pr pr_typ fxy tys) haftmann@27304: | pr_typ fxy (ITyVar v) = str ("'" ^ v); haftmann@27304: fun pr_term thm pat vars fxy (IConst c) = haftmann@27304: pr_app thm pat vars fxy (c, []) haftmann@27304: | pr_term thm pat vars fxy (IVar v) = haftmann@24219: str (CodeName.lookup_var vars v) haftmann@27304: | pr_term thm pat vars fxy (t as t1 `$ t2) = haftmann@24219: (case CodeThingol.unfold_const_app t haftmann@27304: of SOME c_ts => pr_app thm pat vars fxy c_ts haftmann@24219: | NONE => haftmann@27304: brackify fxy [pr_term thm pat vars NOBR t1, pr_term thm pat vars BR t2]) haftmann@27304: | pr_term thm pat vars fxy (t as _ `|-> _) = haftmann@24219: let haftmann@24219: val (binds, t') = CodeThingol.unfold_abs t; haftmann@24219: fun pr ((v, pat), ty) = haftmann@27304: pr_bind thm NOBR ((SOME v, pat), ty) haftmann@24219: #>> (fn p => concat [str "fn", p, str "=>"]); haftmann@24219: val (ps, vars') = fold_map pr binds vars; haftmann@27304: in brackets (ps @ [pr_term thm pat vars' NOBR t']) end haftmann@27304: | pr_term thm pat vars fxy (ICase (cases as (_, t0))) = haftmann@24992: (case CodeThingol.unfold_const_app t0 haftmann@27103: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@27304: then pr_case thm vars fxy cases haftmann@27304: else pr_app thm pat vars fxy c_ts haftmann@27304: | NONE => pr_case thm vars fxy cases) haftmann@27304: and pr_app' thm pat vars (app as ((c, (iss, tys)), ts)) = haftmann@24219: if is_cons c then let haftmann@24219: val k = length tys haftmann@24219: in if k < 2 then haftmann@27304: (str o deresolve) c :: map (pr_term thm pat vars BR) ts haftmann@24219: else if k = length ts then haftmann@27304: [(str o deresolve) c, Pretty.enum "," "(" ")" (map (pr_term thm pat vars NOBR) ts)] haftmann@27710: else [pr_term thm pat vars BR (CodeThingol.eta_expand k app)] end else haftmann@27103: (str o deresolve) c haftmann@27304: :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term thm pat vars BR) ts haftmann@27304: and pr_app thm pat vars = gen_pr_app pr_app' pr_term syntax_const is_cons thm pat vars haftmann@24219: and pr_bind' ((NONE, NONE), _) = str "_" haftmann@24219: | pr_bind' ((SOME v, NONE), _) = str v haftmann@24219: | pr_bind' ((NONE, SOME p), _) = p haftmann@24219: | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p] haftmann@27304: and pr_bind thm = gen_pr_bind pr_bind' pr_term thm haftmann@27304: and pr_case thm vars fxy (cases as ((_, [_]), _)) = haftmann@24219: let haftmann@24219: val (binds, t') = CodeThingol.unfold_let (ICase cases); haftmann@24219: fun pr ((pat, ty), t) vars = haftmann@24219: vars haftmann@27304: |> pr_bind thm NOBR ((NONE, SOME pat), ty) haftmann@27304: |>> (fn p => semicolon [str "val", p, str "=", pr_term thm false vars NOBR t]) haftmann@24219: val (ps, vars') = fold_map pr binds vars; haftmann@24219: in haftmann@24219: Pretty.chunks [ haftmann@24219: [str ("let"), Pretty.fbrk, Pretty.chunks ps] |> Pretty.block, haftmann@27304: [str ("in"), Pretty.fbrk, pr_term thm false vars' NOBR t'] |> Pretty.block, haftmann@24219: str ("end") haftmann@24219: ] haftmann@24219: end haftmann@27304: | pr_case thm vars fxy (((td, ty), b::bs), _) = haftmann@24219: let haftmann@24219: fun pr delim (pat, t) = haftmann@24219: let haftmann@27304: val (p, vars') = pr_bind thm NOBR ((NONE, SOME pat), ty) vars; haftmann@24219: in haftmann@27304: concat [str delim, p, str "=>", pr_term thm false vars' NOBR t] haftmann@24219: end; haftmann@24219: in haftmann@24219: (Pretty.enclose "(" ")" o single o brackify fxy) ( haftmann@24219: str "case" haftmann@27304: :: pr_term thm false vars NOBR td haftmann@24219: :: pr "of" b haftmann@24219: :: map (pr "|") bs haftmann@24219: ) haftmann@24219: end haftmann@27304: | pr_case thm vars fxy ((_, []), _) = str "raise Fail \"empty case\""; haftmann@27103: fun pr_stmt (MLFuns (funns as (funn :: funns'))) = haftmann@24219: let haftmann@24219: val definer = haftmann@24219: let haftmann@27304: fun no_args _ (((ts, _), _) :: _) = length ts haftmann@24841: | no_args ty [] = (length o fst o CodeThingol.unfold_fun) ty; haftmann@24841: fun mk 0 [] = "val" haftmann@24992: | mk 0 vs = if (null o filter_out (null o snd)) vs haftmann@24992: then "val" else "fun" haftmann@24841: | mk k _ = "fun"; haftmann@24841: fun chk (_, ((vs, ty), eqs)) NONE = SOME (mk (no_args ty eqs) vs) haftmann@24841: | chk (_, ((vs, ty), eqs)) (SOME defi) = haftmann@24841: if defi = mk (no_args ty eqs) vs then SOME defi haftmann@24219: else error ("Mixing simultaneous vals and funs not implemented: " haftmann@24219: ^ commas (map (labelled_name o fst) funns)); haftmann@24219: in the (fold chk funns NONE) end; haftmann@27103: fun pr_funn definer (name, ((vs, ty), [])) = haftmann@24219: let haftmann@27103: val vs_dict = filter_out (null o snd) vs; haftmann@27103: val n = length vs_dict + (length o fst o CodeThingol.unfold_fun) ty; haftmann@24992: val exc_str = haftmann@24992: (ML_Syntax.print_string o NameSpace.base o NameSpace.qualifier) name; haftmann@24219: in haftmann@24219: concat ( haftmann@24841: str definer haftmann@27103: :: (str o deresolve) name haftmann@24841: :: map str (replicate n "_") haftmann@24841: @ str "=" haftmann@24841: :: str "raise" haftmann@24841: :: str "(Fail" haftmann@27103: @@ str (exc_str ^ ")") haftmann@24219: ) haftmann@24219: end haftmann@27103: | pr_funn definer (name, ((vs, ty), eqs as eq :: eqs')) = haftmann@24841: let haftmann@27103: val vs_dict = filter_out (null o snd) vs; haftmann@24841: val shift = if null eqs' then I else haftmann@24841: map (Pretty.block o single o Pretty.block o single); haftmann@27304: fun pr_eq definer ((ts, t), thm) = haftmann@24841: let haftmann@24841: val consts = map_filter haftmann@27103: (fn c => if (is_some o syntax_const) c haftmann@27103: then NONE else (SOME o NameSpace.base o deresolve) c) haftmann@24841: ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@27103: val vars = reserved_names haftmann@24841: |> CodeName.intro_vars consts haftmann@24841: |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames) haftmann@24841: (insert (op =)) ts []); haftmann@24841: in haftmann@24841: concat ( haftmann@27103: [str definer, (str o deresolve) name] haftmann@27103: @ (if null ts andalso null vs_dict haftmann@24841: then [str ":", pr_typ NOBR ty] haftmann@24841: else haftmann@27103: pr_tyvar_dicts vs_dict haftmann@27304: @ map (pr_term thm true vars BR) ts) haftmann@27304: @ [str "=", pr_term thm false vars NOBR t] haftmann@24841: ) haftmann@24841: end haftmann@24841: in haftmann@24841: (Pretty.block o Pretty.fbreaks o shift) ( haftmann@24841: pr_eq definer eq haftmann@24841: :: map (pr_eq "|") eqs' haftmann@24841: ) haftmann@24841: end; haftmann@24219: val (ps, p) = split_last (pr_funn definer funn :: map (pr_funn "and") funns'); haftmann@24219: in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end haftmann@27103: | pr_stmt (MLDatas (datas as (data :: datas'))) = haftmann@24219: let haftmann@24219: fun pr_co (co, []) = haftmann@27103: str (deresolve co) haftmann@24219: | pr_co (co, tys) = haftmann@24219: concat [ haftmann@27103: str (deresolve co), haftmann@24219: str "of", haftmann@24219: Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys) haftmann@24219: ]; haftmann@24219: fun pr_data definer (tyco, (vs, [])) = haftmann@24219: concat ( haftmann@24219: str definer haftmann@24219: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@24219: :: str "=" haftmann@24219: @@ str "EMPTY__" haftmann@24219: ) haftmann@24219: | pr_data definer (tyco, (vs, cos)) = haftmann@24219: concat ( haftmann@24219: str definer haftmann@24219: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@24219: :: str "=" haftmann@24219: :: separate (str "|") (map pr_co cos) haftmann@24219: ); haftmann@24992: val (ps, p) = split_last haftmann@24992: (pr_data "datatype" data :: map (pr_data "and") datas'); haftmann@24219: in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end haftmann@27103: | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) = haftmann@24219: let haftmann@24219: val w = first_upper v ^ "_"; haftmann@24219: fun pr_superclass_field (class, classrel) = haftmann@24219: (concat o map str) [ haftmann@27103: pr_label_classrel classrel, ":", "'" ^ v, deresolve class haftmann@24219: ]; haftmann@24841: fun pr_classparam_field (classparam, ty) = haftmann@24219: concat [ haftmann@24841: (str o pr_label_classparam) classparam, str ":", pr_typ NOBR ty haftmann@24219: ]; haftmann@24841: fun pr_classparam_proj (classparam, _) = haftmann@24219: semicolon [ haftmann@24219: str "fun", haftmann@27103: (str o deresolve) classparam, haftmann@27103: Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)], haftmann@24219: str "=", haftmann@24841: str ("#" ^ pr_label_classparam classparam), haftmann@24219: str w haftmann@24219: ]; haftmann@24219: fun pr_superclass_proj (_, classrel) = haftmann@24219: semicolon [ haftmann@24219: str "fun", haftmann@27103: (str o deresolve) classrel, haftmann@27103: Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)], haftmann@24219: str "=", haftmann@24219: str ("#" ^ pr_label_classrel classrel), haftmann@24219: str w haftmann@24219: ]; haftmann@24219: in haftmann@24219: Pretty.chunks ( haftmann@24219: concat [ haftmann@24219: str ("type '" ^ v), haftmann@27103: (str o deresolve) class, haftmann@24219: str "=", haftmann@24219: Pretty.enum "," "{" "};" ( haftmann@24841: map pr_superclass_field superclasses @ map pr_classparam_field classparams haftmann@24219: ) haftmann@24219: ] haftmann@24219: :: map pr_superclass_proj superclasses haftmann@24841: @ map pr_classparam_proj classparams haftmann@24219: ) haftmann@24219: end haftmann@27103: | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) = haftmann@24219: let haftmann@24219: fun pr_superclass (_, (classrel, dss)) = haftmann@24219: concat [ haftmann@24219: (str o pr_label_classrel) classrel, haftmann@24219: str "=", haftmann@24219: pr_dicts NOBR [DictConst dss] haftmann@24219: ]; haftmann@27304: fun pr_classparam ((classparam, c_inst), thm) = haftmann@24591: concat [ haftmann@24841: (str o pr_label_classparam) classparam, haftmann@24591: str "=", haftmann@27304: pr_app thm false reserved_names NOBR (c_inst, []) haftmann@24591: ]; haftmann@24219: in haftmann@24219: semicolon ([ haftmann@24219: str (if null arity then "val" else "fun"), haftmann@27103: (str o deresolve) inst ] @ haftmann@27103: pr_tyvar_dicts arity @ [ haftmann@24219: str "=", haftmann@24992: Pretty.enum "," "{" "}" haftmann@24992: (map pr_superclass superarities @ map pr_classparam classparam_insts), haftmann@24219: str ":", haftmann@24219: pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) haftmann@24219: ]) haftmann@24219: end; haftmann@27103: in pr_stmt end; haftmann@24219: haftmann@27103: fun pr_sml_module name content = haftmann@27103: Pretty.chunks ( haftmann@27103: str ("structure " ^ name ^ " = ") haftmann@27103: :: str "struct" haftmann@27103: :: str "" haftmann@27103: :: content haftmann@27103: @ str "" haftmann@27103: @@ str ("end; (*struct " ^ name ^ "*)") haftmann@27103: ); haftmann@24219: haftmann@27103: fun pr_ocaml_stmt syntax_tyco syntax_const labelled_name reserved_names deresolve is_cons = haftmann@24219: let haftmann@24219: fun pr_dicts fxy ds = haftmann@24219: let haftmann@24219: fun pr_dictvar (v, (_, 1)) = "_" ^ first_upper v haftmann@24219: | pr_dictvar (v, (i, _)) = "_" ^ first_upper v ^ string_of_int (i+1); haftmann@24219: fun pr_proj ps p = haftmann@24219: fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p haftmann@27103: fun pr_dict fxy (DictConst (inst, dss)) = haftmann@27103: brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss) haftmann@27103: | pr_dict fxy (DictVar (classrels, v)) = haftmann@27103: pr_proj (map deresolve classrels) ((str o pr_dictvar) v) haftmann@24219: in case ds haftmann@24219: of [] => str "()" haftmann@27103: | [d] => pr_dict fxy d haftmann@27103: | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds haftmann@24219: end; haftmann@27103: fun pr_tyvar_dicts vs = haftmann@24219: vs haftmann@24992: |> map (fn (v, sort) => map_index (fn (i, _) => haftmann@24992: DictVar ([], (v, (i, length sort)))) sort) haftmann@24219: |> map (pr_dicts BR); haftmann@24219: fun pr_tycoexpr fxy (tyco, tys) = haftmann@24219: let haftmann@27103: val tyco' = (str o deresolve) tyco haftmann@24219: in case map (pr_typ BR) tys haftmann@24219: of [] => tyco' haftmann@24219: | [p] => Pretty.block [p, Pretty.brk 1, tyco'] haftmann@24219: | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco'] haftmann@24219: end haftmann@27304: and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco haftmann@27304: of NONE => pr_tycoexpr fxy (tyco, tys) haftmann@27304: | SOME (i, pr) => pr pr_typ fxy tys) haftmann@27304: | pr_typ fxy (ITyVar v) = str ("'" ^ v); haftmann@27304: fun pr_term thm pat vars fxy (IConst c) = haftmann@27304: pr_app thm pat vars fxy (c, []) haftmann@27304: | pr_term thm pat vars fxy (IVar v) = haftmann@24219: str (CodeName.lookup_var vars v) haftmann@27304: | pr_term thm pat vars fxy (t as t1 `$ t2) = haftmann@24219: (case CodeThingol.unfold_const_app t haftmann@27304: of SOME c_ts => pr_app thm pat vars fxy c_ts haftmann@24219: | NONE => haftmann@27304: brackify fxy [pr_term thm pat vars NOBR t1, pr_term thm pat vars BR t2]) haftmann@27304: | pr_term thm pat vars fxy (t as _ `|-> _) = haftmann@24219: let haftmann@24219: val (binds, t') = CodeThingol.unfold_abs t; haftmann@27304: fun pr ((v, pat), ty) = pr_bind thm BR ((SOME v, pat), ty); haftmann@24219: val (ps, vars') = fold_map pr binds vars; haftmann@27304: in brackets (str "fun" :: ps @ str "->" @@ pr_term thm pat vars' NOBR t') end haftmann@27304: | pr_term thm pat vars fxy (ICase (cases as (_, t0))) = (case CodeThingol.unfold_const_app t0 haftmann@27103: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@27304: then pr_case thm vars fxy cases haftmann@27304: else pr_app thm pat vars fxy c_ts haftmann@27304: | NONE => pr_case thm vars fxy cases) haftmann@27304: and pr_app' thm pat vars (app as ((c, (iss, tys)), ts)) = haftmann@24219: if is_cons c then haftmann@24219: if length tys = length ts haftmann@24219: then case ts haftmann@27103: of [] => [(str o deresolve) c] haftmann@27304: | [t] => [(str o deresolve) c, pr_term thm pat vars BR t] haftmann@27103: | _ => [(str o deresolve) c, Pretty.enum "," "(" ")" haftmann@27304: (map (pr_term thm pat vars NOBR) ts)] haftmann@27710: else [pr_term thm pat vars BR (CodeThingol.eta_expand (length tys) app)] haftmann@27103: else (str o deresolve) c haftmann@27304: :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term thm pat vars BR) ts) haftmann@27304: and pr_app thm pat vars = gen_pr_app pr_app' pr_term syntax_const is_cons thm pat vars haftmann@24219: and pr_bind' ((NONE, NONE), _) = str "_" haftmann@24219: | pr_bind' ((SOME v, NONE), _) = str v haftmann@24219: | pr_bind' ((NONE, SOME p), _) = p haftmann@24219: | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v] haftmann@27304: and pr_bind thm = gen_pr_bind pr_bind' pr_term thm haftmann@27304: and pr_case thm vars fxy (cases as ((_, [_]), _)) = haftmann@24219: let haftmann@24219: val (binds, t') = CodeThingol.unfold_let (ICase cases); haftmann@24219: fun pr ((pat, ty), t) vars = haftmann@24219: vars haftmann@27304: |> pr_bind thm NOBR ((NONE, SOME pat), ty) haftmann@24992: |>> (fn p => concat haftmann@27304: [str "let", p, str "=", pr_term thm false vars NOBR t, str "in"]) haftmann@24219: val (ps, vars') = fold_map pr binds vars; haftmann@27304: in Pretty.chunks (ps @| pr_term thm false vars' NOBR t') end haftmann@27304: | pr_case thm vars fxy (((td, ty), b::bs), _) = haftmann@24219: let haftmann@24219: fun pr delim (pat, t) = haftmann@24219: let haftmann@27304: val (p, vars') = pr_bind thm NOBR ((NONE, SOME pat), ty) vars; haftmann@27304: in concat [str delim, p, str "->", pr_term thm false vars' NOBR t] end; haftmann@24219: in haftmann@24219: (Pretty.enclose "(" ")" o single o brackify fxy) ( haftmann@24219: str "match" haftmann@27304: :: pr_term thm false vars NOBR td haftmann@24219: :: pr "with" b haftmann@24219: :: map (pr "|") bs haftmann@24219: ) haftmann@24219: end haftmann@27304: | pr_case thm vars fxy ((_, []), _) = str "failwith \"empty case\""; haftmann@27103: fun fish_params vars eqs = haftmann@27103: let haftmann@27103: fun fish_param _ (w as SOME _) = w haftmann@27103: | fish_param (IVar v) NONE = SOME v haftmann@27103: | fish_param _ NONE = NONE; haftmann@27103: fun fillup_param _ (_, SOME v) = v haftmann@27103: | fillup_param x (i, NONE) = x ^ string_of_int i; haftmann@27103: val fished1 = fold (map2 fish_param) eqs (replicate (length (hd eqs)) NONE); haftmann@27103: val x = Name.variant (map_filter I fished1) "x"; haftmann@27103: val fished2 = map_index (fillup_param x) fished1; haftmann@27103: val (fished3, _) = Name.variants fished2 Name.context; haftmann@27103: val vars' = CodeName.intro_vars fished3 vars; haftmann@27103: in map (CodeName.lookup_var vars') fished3 end; haftmann@27103: fun pr_stmt (MLFuns (funns as funn :: funns')) = haftmann@24219: let haftmann@27304: fun pr_eq ((ts, t), thm) = haftmann@24219: let haftmann@24219: val consts = map_filter haftmann@27103: (fn c => if (is_some o syntax_const) c haftmann@27103: then NONE else (SOME o NameSpace.base o deresolve) c) haftmann@24219: ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@27103: val vars = reserved_names haftmann@24219: |> CodeName.intro_vars consts haftmann@24219: |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames) haftmann@24219: (insert (op =)) ts []); haftmann@24219: in concat [ haftmann@27304: (Pretty.block o Pretty.commas) (map (pr_term thm true vars NOBR) ts), haftmann@24219: str "->", haftmann@27304: pr_term thm false vars NOBR t haftmann@24219: ] end; haftmann@24841: fun pr_eqs name ty [] = haftmann@24841: let haftmann@24841: val n = (length o fst o CodeThingol.unfold_fun) ty; haftmann@24992: val exc_str = haftmann@24992: (ML_Syntax.print_string o NameSpace.base o NameSpace.qualifier) name; haftmann@24841: in haftmann@24841: concat ( haftmann@24841: map str (replicate n "_") haftmann@24841: @ str "=" haftmann@24841: :: str "failwith" haftmann@24992: @@ str exc_str haftmann@24841: ) haftmann@24841: end haftmann@27304: | pr_eqs _ _ [((ts, t), thm)] = haftmann@24219: let haftmann@24219: val consts = map_filter haftmann@27103: (fn c => if (is_some o syntax_const) c haftmann@27103: then NONE else (SOME o NameSpace.base o deresolve) c) haftmann@24219: ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@27103: val vars = reserved_names haftmann@24219: |> CodeName.intro_vars consts haftmann@24219: |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames) haftmann@24219: (insert (op =)) ts []); haftmann@24219: in haftmann@24219: concat ( haftmann@27304: map (pr_term thm true vars BR) ts haftmann@24219: @ str "=" haftmann@27304: @@ pr_term thm false vars NOBR t haftmann@24219: ) haftmann@24219: end haftmann@27304: | pr_eqs _ _ (eqs as (eq as (([_], _), _)) :: eqs') = haftmann@24219: Pretty.block ( haftmann@24219: str "=" haftmann@24219: :: Pretty.brk 1 haftmann@24219: :: str "function" haftmann@24219: :: Pretty.brk 1 haftmann@24219: :: pr_eq eq haftmann@24992: :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] haftmann@24992: o single o pr_eq) eqs' haftmann@24219: ) haftmann@24841: | pr_eqs _ _ (eqs as eq :: eqs') = haftmann@24219: let haftmann@24219: val consts = map_filter haftmann@27103: (fn c => if (is_some o syntax_const) c haftmann@27103: then NONE else (SOME o NameSpace.base o deresolve) c) haftmann@24992: ((fold o CodeThingol.fold_constnames) haftmann@27304: (insert (op =)) (map (snd o fst) eqs) []); haftmann@27103: val vars = reserved_names haftmann@24219: |> CodeName.intro_vars consts; haftmann@27304: val dummy_parms = (map str o fish_params vars o map (fst o fst)) eqs; haftmann@24219: in haftmann@24219: Pretty.block ( haftmann@24219: Pretty.breaks dummy_parms haftmann@24219: @ Pretty.brk 1 haftmann@24219: :: str "=" haftmann@24219: :: Pretty.brk 1 haftmann@24219: :: str "match" haftmann@24219: :: Pretty.brk 1 haftmann@24219: :: (Pretty.block o Pretty.commas) dummy_parms haftmann@24219: :: Pretty.brk 1 haftmann@24219: :: str "with" haftmann@24219: :: Pretty.brk 1 haftmann@24219: :: pr_eq eq haftmann@24992: :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] haftmann@24992: o single o pr_eq) eqs' haftmann@24219: ) haftmann@24219: end; haftmann@24381: fun pr_funn definer (name, ((vs, ty), eqs)) = haftmann@24219: concat ( haftmann@24219: str definer haftmann@27103: :: (str o deresolve) name haftmann@27103: :: pr_tyvar_dicts (filter_out (null o snd) vs) haftmann@24841: @| pr_eqs name ty eqs haftmann@24219: ); haftmann@24992: val (ps, p) = split_last haftmann@24992: (pr_funn "let rec" funn :: map (pr_funn "and") funns'); haftmann@24219: in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end haftmann@27103: | pr_stmt (MLDatas (datas as (data :: datas'))) = haftmann@24219: let haftmann@24219: fun pr_co (co, []) = haftmann@27103: str (deresolve co) haftmann@24219: | pr_co (co, tys) = haftmann@24219: concat [ haftmann@27103: str (deresolve co), haftmann@24219: str "of", haftmann@24219: Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys) haftmann@24219: ]; haftmann@24219: fun pr_data definer (tyco, (vs, [])) = haftmann@24219: concat ( haftmann@24219: str definer haftmann@24219: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@24219: :: str "=" haftmann@24219: @@ str "EMPTY_" haftmann@24219: ) haftmann@24219: | pr_data definer (tyco, (vs, cos)) = haftmann@24219: concat ( haftmann@24219: str definer haftmann@24219: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@24219: :: str "=" haftmann@24219: :: separate (str "|") (map pr_co cos) haftmann@24219: ); haftmann@24992: val (ps, p) = split_last haftmann@24992: (pr_data "type" data :: map (pr_data "and") datas'); haftmann@24219: in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end haftmann@27103: | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) = haftmann@24219: let haftmann@24219: val w = "_" ^ first_upper v; haftmann@24219: fun pr_superclass_field (class, classrel) = haftmann@24219: (concat o map str) [ haftmann@27103: deresolve classrel, ":", "'" ^ v, deresolve class haftmann@24219: ]; haftmann@24841: fun pr_classparam_field (classparam, ty) = haftmann@24219: concat [ haftmann@27103: (str o deresolve) classparam, str ":", pr_typ NOBR ty haftmann@24219: ]; haftmann@24841: fun pr_classparam_proj (classparam, _) = haftmann@24219: concat [ haftmann@24219: str "let", haftmann@27103: (str o deresolve) classparam, haftmann@24219: str w, haftmann@24219: str "=", haftmann@27103: str (w ^ "." ^ deresolve classparam ^ ";;") haftmann@24219: ]; haftmann@24219: in Pretty.chunks ( haftmann@24219: concat [ haftmann@24219: str ("type '" ^ v), haftmann@27103: (str o deresolve) class, haftmann@24219: str "=", haftmann@25771: enum_default "();;" ";" "{" "};;" ( haftmann@24992: map pr_superclass_field superclasses haftmann@24992: @ map pr_classparam_field classparams haftmann@24219: ) haftmann@24219: ] haftmann@24841: :: map pr_classparam_proj classparams haftmann@24219: ) end haftmann@27103: | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) = haftmann@24219: let haftmann@24219: fun pr_superclass (_, (classrel, dss)) = haftmann@24219: concat [ haftmann@27103: (str o deresolve) classrel, haftmann@24219: str "=", haftmann@24219: pr_dicts NOBR [DictConst dss] haftmann@24219: ]; haftmann@27304: fun pr_classparam_inst ((classparam, c_inst), thm) = haftmann@24591: concat [ haftmann@27103: (str o deresolve) classparam, haftmann@24591: str "=", haftmann@27304: pr_app thm false reserved_names NOBR (c_inst, []) haftmann@24591: ]; haftmann@24219: in haftmann@24219: concat ( haftmann@24219: str "let" haftmann@27103: :: (str o deresolve) inst haftmann@27103: :: pr_tyvar_dicts arity haftmann@24219: @ str "=" haftmann@24219: @@ (Pretty.enclose "(" ");;" o Pretty.breaks) [ haftmann@25771: enum_default "()" ";" "{" "}" (map pr_superclass superarities haftmann@24992: @ map pr_classparam_inst classparam_insts), haftmann@24219: str ":", haftmann@24219: pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) haftmann@24219: ] haftmann@24219: ) haftmann@24219: end; haftmann@27103: in pr_stmt end; haftmann@27103: haftmann@27103: fun pr_ocaml_module name content = haftmann@27103: Pretty.chunks ( haftmann@27103: str ("module " ^ name ^ " = ") haftmann@27103: :: str "struct" haftmann@27103: :: str "" haftmann@27103: :: content haftmann@27103: @ str "" haftmann@27103: @@ str ("end;; (*struct " ^ name ^ "*)") haftmann@27103: ); haftmann@24219: haftmann@27103: local haftmann@27103: haftmann@27103: datatype ml_node = haftmann@27103: Dummy of string haftmann@27103: | Stmt of string * ml_stmt haftmann@27103: | Module of string * ((Name.context * Name.context) * ml_node Graph.T); haftmann@24219: haftmann@27103: in haftmann@27103: haftmann@27103: fun ml_node_of_program labelled_name module_name reserved_names raw_module_alias program = haftmann@24219: let haftmann@27103: val module_alias = if is_some module_name then K module_name else raw_module_alias; haftmann@27103: val reserved_names = Name.make_context reserved_names; haftmann@27103: val empty_module = ((reserved_names, reserved_names), Graph.empty); haftmann@24219: fun map_node [] f = f haftmann@24219: | map_node (m::ms) f = haftmann@27103: Graph.default_node (m, Module (m, empty_module)) haftmann@27103: #> Graph.map_node m (fn (Module (module_name, (nsp, nodes))) => haftmann@27103: Module (module_name, (nsp, map_node ms f nodes))); haftmann@24219: fun map_nsp_yield [] f (nsp, nodes) = haftmann@24219: let haftmann@24219: val (x, nsp') = f nsp haftmann@24219: in (x, (nsp', nodes)) end haftmann@24219: | map_nsp_yield (m::ms) f (nsp, nodes) = haftmann@24219: let haftmann@24219: val (x, nodes') = haftmann@24219: nodes haftmann@27103: |> Graph.default_node (m, Module (m, empty_module)) haftmann@27103: |> Graph.map_node_yield m (fn Module (d_module_name, nsp_nodes) => haftmann@24219: let haftmann@24219: val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes haftmann@27103: in (x, Module (d_module_name, nsp_nodes')) end) haftmann@24219: in (x, (nsp, nodes')) end; haftmann@27103: fun map_nsp_fun_yield f (nsp_fun, nsp_typ) = haftmann@27103: let haftmann@27103: val (x, nsp_fun') = f nsp_fun haftmann@27103: in (x, (nsp_fun', nsp_typ)) end; haftmann@27103: fun map_nsp_typ_yield f (nsp_fun, nsp_typ) = haftmann@27103: let haftmann@27103: val (x, nsp_typ') = f nsp_typ haftmann@27103: in (x, (nsp_fun, nsp_typ')) end; haftmann@27103: val mk_name_module = mk_name_module reserved_names NONE module_alias program; haftmann@27103: fun mk_name_stmt upper name nsp = haftmann@24219: let haftmann@24219: val (_, base) = dest_name name; haftmann@24219: val base' = if upper then first_upper base else base; haftmann@24219: val ([base''], nsp') = Name.variants [base'] nsp; haftmann@24219: in (base'', nsp') end; haftmann@27103: fun add_funs stmts = haftmann@24219: fold_map haftmann@27103: (fn (name, CodeThingol.Fun stmt) => haftmann@27103: map_nsp_fun_yield (mk_name_stmt false name) #>> haftmann@27304: rpair (name, stmt) haftmann@27103: | (name, _) => haftmann@27024: error ("Function block containing illegal statement: " ^ labelled_name name) haftmann@27103: ) stmts haftmann@27103: #>> (split_list #> apsnd MLFuns); haftmann@27103: fun add_datatypes stmts = haftmann@24219: fold_map haftmann@27103: (fn (name, CodeThingol.Datatype stmt) => haftmann@27103: map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt)) haftmann@24219: | (name, CodeThingol.Datatypecons _) => haftmann@27103: map_nsp_fun_yield (mk_name_stmt true name) #>> rpair NONE haftmann@27103: | (name, _) => haftmann@27024: error ("Datatype block containing illegal statement: " ^ labelled_name name) haftmann@27103: ) stmts haftmann@27103: #>> (split_list #> apsnd (map_filter I haftmann@27024: #> (fn [] => error ("Datatype block without data statement: " haftmann@27103: ^ (commas o map (labelled_name o fst)) stmts) haftmann@27103: | stmts => MLDatas stmts))); haftmann@27103: fun add_class stmts = haftmann@24219: fold_map haftmann@24219: (fn (name, CodeThingol.Class info) => haftmann@27103: map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, info)) haftmann@24219: | (name, CodeThingol.Classrel _) => haftmann@27103: map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE haftmann@24841: | (name, CodeThingol.Classparam _) => haftmann@27103: map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE haftmann@27103: | (name, _) => haftmann@27024: error ("Class block containing illegal statement: " ^ labelled_name name) haftmann@27103: ) stmts haftmann@27103: #>> (split_list #> apsnd (map_filter I haftmann@27024: #> (fn [] => error ("Class block without class statement: " haftmann@27103: ^ (commas o map (labelled_name o fst)) stmts) haftmann@27103: | [stmt] => MLClass stmt))); haftmann@27103: fun add_inst [(name, CodeThingol.Classinst stmt)] = haftmann@27103: map_nsp_fun_yield (mk_name_stmt false name) haftmann@27304: #>> (fn base => ([base], MLClassinst (name, stmt))); haftmann@27103: fun add_stmts ((stmts as (_, CodeThingol.Fun _)::_)) = haftmann@27103: add_funs stmts haftmann@27103: | add_stmts ((stmts as (_, CodeThingol.Datatypecons _)::_)) = haftmann@27103: add_datatypes stmts haftmann@27103: | add_stmts ((stmts as (_, CodeThingol.Datatype _)::_)) = haftmann@27103: add_datatypes stmts haftmann@27103: | add_stmts ((stmts as (_, CodeThingol.Class _)::_)) = haftmann@27103: add_class stmts haftmann@27103: | add_stmts ((stmts as (_, CodeThingol.Classrel _)::_)) = haftmann@27103: add_class stmts haftmann@27103: | add_stmts ((stmts as (_, CodeThingol.Classparam _)::_)) = haftmann@27103: add_class stmts haftmann@27103: | add_stmts ((stmts as [(_, CodeThingol.Classinst _)])) = haftmann@27103: add_inst stmts haftmann@27103: | add_stmts stmts = error ("Illegal mutual dependencies: " ^ haftmann@27103: (commas o map (labelled_name o fst)) stmts); haftmann@27103: fun add_stmts' stmts nsp_nodes = haftmann@24219: let haftmann@27103: val names as (name :: names') = map fst stmts; haftmann@24219: val deps = haftmann@24219: [] haftmann@27103: |> fold (fold (insert (op =)) o Graph.imm_succs program) names haftmann@24219: |> subtract (op =) names; haftmann@27103: val (module_names, _) = (split_list o map dest_name) names; haftmann@27103: val module_name = (the_single o distinct (op =) o map mk_name_module) module_names haftmann@24219: handle Empty => haftmann@24811: error ("Different namespace prefixes for mutual dependencies:\n" haftmann@24811: ^ commas (map labelled_name names) haftmann@24811: ^ "\n" haftmann@27103: ^ commas module_names); haftmann@27103: val module_name_path = NameSpace.explode module_name; haftmann@27103: fun add_dep name name' = haftmann@24219: let haftmann@27103: val module_name' = (mk_name_module o fst o dest_name) name'; haftmann@27103: in if module_name = module_name' then haftmann@27103: map_node module_name_path (Graph.add_edge (name, name')) haftmann@24219: else let haftmann@24219: val (common, (diff1::_, diff2::_)) = chop_prefix (op =) haftmann@27103: (module_name_path, NameSpace.explode module_name'); haftmann@24219: in haftmann@24219: map_node common haftmann@27103: (fn node => Graph.add_edge_acyclic (diff1, diff2) node haftmann@24219: handle Graph.CYCLES _ => error ("Dependency " haftmann@27103: ^ quote name ^ " -> " ^ quote name' haftmann@24219: ^ " would result in module dependency cycle")) haftmann@24219: end end; haftmann@24219: in haftmann@24219: nsp_nodes haftmann@27103: |> map_nsp_yield module_name_path (add_stmts stmts) haftmann@27103: |-> (fn (base' :: bases', stmt') => haftmann@27103: apsnd (map_node module_name_path (Graph.new_node (name, (Stmt (base', stmt'))) haftmann@24992: #> fold2 (fn name' => fn base' => haftmann@27103: Graph.new_node (name', (Dummy base'))) names' bases'))) haftmann@24219: |> apsnd (fold (fn name => fold (add_dep name) deps) names) haftmann@27103: |> apsnd (fold_product (curry (map_node module_name_path o Graph.add_edge)) names names) haftmann@24219: end; haftmann@27103: val (_, nodes) = empty_module haftmann@27103: |> fold add_stmts' (map (AList.make (Graph.get_node program)) haftmann@27103: (rev (Graph.strong_conn program))); haftmann@24219: fun deresolver prefix name = haftmann@24219: let haftmann@27103: val module_name = (fst o dest_name) name; haftmann@27103: val module_name' = (NameSpace.explode o mk_name_module) module_name; haftmann@27103: val (_, (_, remainder)) = chop_prefix (op =) (prefix, module_name'); haftmann@27103: val stmt_name = haftmann@24219: nodes haftmann@27103: |> fold (fn name => fn node => case Graph.get_node node name haftmann@27103: of Module (_, (_, node)) => node) module_name' haftmann@27103: |> (fn node => case Graph.get_node node name of Stmt (stmt_name, _) => stmt_name haftmann@27103: | Dummy stmt_name => stmt_name); haftmann@24219: in haftmann@27103: NameSpace.implode (remainder @ [stmt_name]) haftmann@24219: end handle Graph.UNDEF _ => haftmann@27024: error ("Unknown statement name: " ^ labelled_name name); haftmann@27103: in (deresolver, nodes) end; haftmann@27103: haftmann@27436: fun serialize_ml compile pr_module pr_stmt raw_module_name labelled_name reserved_names includes raw_module_alias haftmann@27103: _ syntax_tyco syntax_const program cs destination = haftmann@27103: let haftmann@27103: val is_cons = CodeThingol.is_cons program; haftmann@27304: val stmt_names = stmt_names_of_destination destination; haftmann@27304: val module_name = if null stmt_names then raw_module_name else SOME "Code"; haftmann@27103: val (deresolver, nodes) = ml_node_of_program labelled_name module_name haftmann@27103: reserved_names raw_module_alias program; haftmann@27103: val reserved_names = CodeName.make_vars reserved_names; haftmann@27103: fun pr_node prefix (Dummy _) = haftmann@24219: NONE haftmann@27304: | pr_node prefix (Stmt (_, stmt)) = if null stmt_names orelse haftmann@27304: (not o null o filter (member (op =) stmt_names) o stmt_names_of) stmt then SOME haftmann@27304: (pr_stmt syntax_tyco syntax_const labelled_name reserved_names haftmann@27304: (deresolver prefix) is_cons stmt) haftmann@27304: else NONE haftmann@27103: | pr_node prefix (Module (module_name, (_, nodes))) = haftmann@27304: let haftmann@27304: val ps = separate (str "") haftmann@27304: ((map_filter (pr_node (prefix @ [module_name]) o Graph.get_node nodes) haftmann@27304: o rev o flat o Graph.strong_conn) nodes) haftmann@27304: in SOME (case destination of String _ => Pretty.chunks ps haftmann@27304: | _ => pr_module module_name ps) haftmann@27304: end; haftmann@27103: val cs' = map_filter (try (deresolver (if is_some module_name then the_list module_name else []))) haftmann@27103: cs; haftmann@24992: val p = Pretty.chunks (separate (str "") (map snd includes @ (map_filter haftmann@26113: (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes)); haftmann@27103: fun output Compile = K NONE o compile o code_of_pretty haftmann@27304: | output Export = K NONE o code_writeln haftmann@27103: | output (File file) = K NONE o File.write file o code_of_pretty haftmann@27436: | output (String _) = SOME o rpair cs' o code_of_pretty; haftmann@27436: in output destination p end; haftmann@27103: haftmann@27103: end; (*local*) haftmann@24219: haftmann@27436: (* ML (system language) code for evaluation and instrumentalization *) haftmann@27436: haftmann@27436: fun ml_code_of thy program cs = mount_serializer thy haftmann@27436: (SOME (fn _ => fn [] => serialize_ml (K ()) (K Pretty.chunks) pr_sml_stmt (SOME ""))) haftmann@27436: target_SML NONE [] program cs (String []) haftmann@27436: |> the; haftmann@27436: haftmann@27436: (* generic entry points for SML/OCaml *) haftmann@27436: haftmann@27103: fun isar_seri_sml module_name = haftmann@27000: parse_args (Scan.succeed ()) haftmann@27103: #> (fn () => serialize_ml (use_text (1, "generated code") Output.ml_output false) haftmann@27436: pr_sml_module pr_sml_stmt module_name); haftmann@26113: haftmann@27103: fun isar_seri_ocaml module_name = haftmann@27000: parse_args (Scan.succeed ()) haftmann@27103: #> (fn () => serialize_ml (fn _ => error "OCaml: no internal compilation") haftmann@27436: pr_ocaml_module pr_ocaml_stmt module_name); haftmann@24219: haftmann@24219: haftmann@24219: (** Haskell serializer **) haftmann@24219: haftmann@27304: fun pr_haskell_bind pr_term = haftmann@27103: let haftmann@27103: fun pr_bind ((NONE, NONE), _) = str "_" haftmann@27103: | pr_bind ((SOME v, NONE), _) = str v haftmann@27103: | pr_bind ((NONE, SOME p), _) = p haftmann@27103: | pr_bind ((SOME v, SOME p), _) = brackets [str v, str "@", p]; haftmann@27304: in gen_pr_bind pr_bind pr_term end; haftmann@24219: haftmann@27103: fun pr_haskell_stmt syntax_class syntax_tyco syntax_const labelled_name haftmann@27103: init_syms deresolve is_cons contr_classparam_typs deriving_show = haftmann@24219: let haftmann@27103: val deresolve_base = NameSpace.base o deresolve; haftmann@27103: fun class_name class = case syntax_class class haftmann@27103: of NONE => deresolve class haftmann@24219: | SOME (class, _) => class; haftmann@27103: fun classparam_name class classparam = case syntax_class class haftmann@27103: of NONE => deresolve_base classparam haftmann@24841: | SOME (_, classparam_syntax) => case classparam_syntax classparam haftmann@24841: of NONE => (snd o dest_name) classparam haftmann@25621: | SOME classparam => classparam; haftmann@25621: fun pr_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs haftmann@25621: of [] => [] haftmann@25621: | classbinds => Pretty.enum "," "(" ")" ( haftmann@25621: map (fn (v, class) => haftmann@25621: str (class_name class ^ " " ^ CodeName.lookup_var tyvars v)) classbinds) haftmann@25621: @@ str " => "; haftmann@25621: fun pr_typforall tyvars vs = case map fst vs haftmann@25621: of [] => [] haftmann@25621: | vnames => str "forall " :: Pretty.breaks haftmann@25621: (map (str o CodeName.lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1; haftmann@24219: fun pr_tycoexpr tyvars fxy (tyco, tys) = haftmann@24219: brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) haftmann@27304: and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco haftmann@27304: of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys) haftmann@27304: | SOME (i, pr) => pr (pr_typ tyvars) fxy tys) haftmann@27304: | pr_typ tyvars fxy (ITyVar v) = (str o CodeName.lookup_var tyvars) v; haftmann@25621: fun pr_typdecl tyvars (vs, tycoexpr) = haftmann@25621: Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr); haftmann@24219: fun pr_typscheme tyvars (vs, ty) = haftmann@25621: Pretty.block (pr_typforall tyvars vs @ pr_typcontext tyvars vs @| pr_typ tyvars NOBR ty); haftmann@27304: fun pr_term tyvars thm pat vars fxy (IConst c) = haftmann@27304: pr_app tyvars thm pat vars fxy (c, []) haftmann@27304: | pr_term tyvars thm pat vars fxy (t as (t1 `$ t2)) = haftmann@24219: (case CodeThingol.unfold_const_app t haftmann@27304: of SOME app => pr_app tyvars thm pat vars fxy app haftmann@24219: | _ => haftmann@24219: brackify fxy [ haftmann@27304: pr_term tyvars thm pat vars NOBR t1, haftmann@27304: pr_term tyvars thm pat vars BR t2 haftmann@24219: ]) haftmann@27304: | pr_term tyvars thm pat vars fxy (IVar v) = haftmann@24219: (str o CodeName.lookup_var vars) v haftmann@27304: | pr_term tyvars thm pat vars fxy (t as _ `|-> _) = haftmann@24219: let haftmann@24219: val (binds, t') = CodeThingol.unfold_abs t; haftmann@27304: fun pr ((v, pat), ty) = pr_bind tyvars thm BR ((SOME v, pat), ty); haftmann@24219: val (ps, vars') = fold_map pr binds vars; haftmann@27304: in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm pat vars' NOBR t') end haftmann@27304: | pr_term tyvars thm pat vars fxy (ICase (cases as (_, t0))) = haftmann@24992: (case CodeThingol.unfold_const_app t0 haftmann@27103: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@27304: then pr_case tyvars thm vars fxy cases haftmann@27304: else pr_app tyvars thm pat vars fxy c_ts haftmann@27304: | NONE => pr_case tyvars thm vars fxy cases) haftmann@27304: and pr_app' tyvars thm pat vars ((c, (_, tys)), ts) = case contr_classparam_typs c haftmann@27304: of [] => (str o deresolve) c :: map (pr_term tyvars thm pat vars BR) ts haftmann@25621: | fingerprint => let haftmann@25621: val ts_fingerprint = ts ~~ curry Library.take (length ts) fingerprint; haftmann@25621: val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) => haftmann@25621: (not o CodeThingol.locally_monomorphic) t) ts_fingerprint; haftmann@27304: fun pr_term_anno (t, NONE) _ = pr_term tyvars thm pat vars BR t haftmann@25621: | pr_term_anno (t, SOME _) ty = haftmann@27304: brackets [pr_term tyvars thm pat vars NOBR t, str "::", pr_typ tyvars NOBR ty]; haftmann@25621: in haftmann@25621: if needs_annotation then haftmann@27103: (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry Library.take (length ts) tys) haftmann@27304: else (str o deresolve) c :: map (pr_term tyvars thm pat vars BR) ts haftmann@25621: end haftmann@27304: and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const is_cons haftmann@27103: and pr_bind tyvars = pr_haskell_bind (pr_term tyvars) haftmann@27304: and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) = haftmann@24219: let haftmann@24219: val (binds, t) = CodeThingol.unfold_let (ICase cases); haftmann@24219: fun pr ((pat, ty), t) vars = haftmann@24219: vars haftmann@27304: |> pr_bind tyvars thm BR ((NONE, SOME pat), ty) haftmann@27304: |>> (fn p => semicolon [p, str "=", pr_term tyvars thm false vars NOBR t]) haftmann@24219: val (ps, vars') = fold_map pr binds vars; haftmann@24219: in haftmann@24219: Pretty.block_enclose ( haftmann@24219: str "let {", haftmann@27304: concat [str "}", str "in", pr_term tyvars thm false vars' NOBR t] haftmann@24219: ) ps haftmann@24219: end haftmann@27304: | pr_case tyvars thm vars fxy (((td, ty), bs as _ :: _), _) = haftmann@24219: let haftmann@24219: fun pr (pat, t) = haftmann@24219: let haftmann@27304: val (p, vars') = pr_bind tyvars thm NOBR ((NONE, SOME pat), ty) vars; haftmann@27304: in semicolon [p, str "->", pr_term tyvars thm false vars' NOBR t] end; haftmann@24219: in haftmann@24219: Pretty.block_enclose ( haftmann@27304: concat [str "(case", pr_term tyvars thm false vars NOBR td, str "of", str "{"], haftmann@24219: str "})" haftmann@24219: ) (map pr bs) haftmann@24219: end haftmann@27304: | pr_case tyvars thm vars fxy ((_, []), _) = str "error \"empty case\""; haftmann@27103: fun pr_stmt (name, CodeThingol.Fun ((vs, ty), [])) = haftmann@24841: let haftmann@24841: val tyvars = CodeName.intro_vars (map fst vs) init_syms; haftmann@24841: val n = (length o fst o CodeThingol.unfold_fun) ty; haftmann@24841: in haftmann@24841: Pretty.chunks [ haftmann@24841: Pretty.block [ haftmann@27103: (str o suffix " ::" o deresolve_base) name, haftmann@24841: Pretty.brk 1, haftmann@24841: pr_typscheme tyvars (vs, ty), haftmann@24841: str ";" haftmann@24841: ], haftmann@24841: concat ( haftmann@27103: (str o deresolve_base) name haftmann@24841: :: map str (replicate n "_") haftmann@24841: @ str "=" haftmann@24841: :: str "error" haftmann@24992: @@ (str o (fn s => s ^ ";") o ML_Syntax.print_string haftmann@24992: o NameSpace.base o NameSpace.qualifier) name haftmann@24841: ) haftmann@24841: ] haftmann@24841: end haftmann@27103: | pr_stmt (name, CodeThingol.Fun ((vs, ty), eqs)) = haftmann@24219: let haftmann@24219: val tyvars = CodeName.intro_vars (map fst vs) init_syms; haftmann@27304: fun pr_eq ((ts, t), thm) = haftmann@24219: let haftmann@24219: val consts = map_filter haftmann@27103: (fn c => if (is_some o syntax_const) c haftmann@27103: then NONE else (SOME o NameSpace.base o deresolve) c) haftmann@24219: ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@24219: val vars = init_syms haftmann@24219: |> CodeName.intro_vars consts haftmann@24219: |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames) haftmann@24219: (insert (op =)) ts []); haftmann@24219: in haftmann@24219: semicolon ( haftmann@27103: (str o deresolve_base) name haftmann@27304: :: map (pr_term tyvars thm true vars BR) ts haftmann@24219: @ str "=" haftmann@27304: @@ pr_term tyvars thm false vars NOBR t haftmann@24219: ) haftmann@24219: end; haftmann@24219: in haftmann@24219: Pretty.chunks ( haftmann@24219: Pretty.block [ haftmann@27103: (str o suffix " ::" o deresolve_base) name, haftmann@24219: Pretty.brk 1, haftmann@24219: pr_typscheme tyvars (vs, ty), haftmann@24219: str ";" haftmann@24219: ] haftmann@24219: :: map pr_eq eqs haftmann@24219: ) haftmann@24219: end haftmann@27103: | pr_stmt (name, CodeThingol.Datatype (vs, [])) = haftmann@24219: let haftmann@24219: val tyvars = CodeName.intro_vars (map fst vs) init_syms; haftmann@24219: in haftmann@24219: semicolon [ haftmann@24219: str "data", haftmann@27103: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) haftmann@24219: ] haftmann@24219: end haftmann@27103: | pr_stmt (name, CodeThingol.Datatype (vs, [(co, [ty])])) = haftmann@24219: let haftmann@24219: val tyvars = CodeName.intro_vars (map fst vs) init_syms; haftmann@24219: in haftmann@24219: semicolon ( haftmann@24219: str "newtype" haftmann@27103: :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) haftmann@24219: :: str "=" haftmann@27103: :: (str o deresolve_base) co haftmann@24219: :: pr_typ tyvars BR ty haftmann@24219: :: (if deriving_show name then [str "deriving (Read, Show)"] else []) haftmann@24219: ) haftmann@24219: end haftmann@27103: | pr_stmt (name, CodeThingol.Datatype (vs, co :: cos)) = haftmann@24219: let haftmann@24219: val tyvars = CodeName.intro_vars (map fst vs) init_syms; haftmann@24219: fun pr_co (co, tys) = haftmann@24219: concat ( haftmann@27103: (str o deresolve_base) co haftmann@24219: :: map (pr_typ tyvars BR) tys haftmann@24219: ) haftmann@24219: in haftmann@24219: semicolon ( haftmann@24219: str "data" haftmann@27103: :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) haftmann@24219: :: str "=" haftmann@24219: :: pr_co co haftmann@24219: :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos haftmann@24219: @ (if deriving_show name then [str "deriving (Read, Show)"] else []) haftmann@24219: ) haftmann@24219: end haftmann@27103: | pr_stmt (name, CodeThingol.Class (v, (superclasses, classparams))) = haftmann@24219: let haftmann@24219: val tyvars = CodeName.intro_vars [v] init_syms; haftmann@24841: fun pr_classparam (classparam, ty) = haftmann@24219: semicolon [ haftmann@24841: (str o classparam_name name) classparam, haftmann@24219: str "::", haftmann@24219: pr_typ tyvars NOBR ty haftmann@24219: ] haftmann@24219: in haftmann@24219: Pretty.block_enclose ( haftmann@24219: Pretty.block [ haftmann@24219: str "class ", haftmann@25621: Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]), haftmann@27103: str (deresolve_base name ^ " " ^ CodeName.lookup_var tyvars v), haftmann@24219: str " where {" haftmann@24219: ], haftmann@24219: str "};" haftmann@24841: ) (map pr_classparam classparams) haftmann@24219: end haftmann@27103: | pr_stmt (_, CodeThingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) = haftmann@24219: let haftmann@24219: val tyvars = CodeName.intro_vars (map fst vs) init_syms; haftmann@27304: fun pr_instdef ((classparam, c_inst), thm) = haftmann@24591: semicolon [ haftmann@24841: (str o classparam_name class) classparam, haftmann@24591: str "=", haftmann@27304: pr_app tyvars thm false init_syms NOBR (c_inst, []) haftmann@24591: ]; haftmann@24219: in haftmann@24219: Pretty.block_enclose ( haftmann@24219: Pretty.block [ haftmann@24219: str "instance ", haftmann@25621: Pretty.block (pr_typcontext tyvars vs), haftmann@24219: str (class_name class ^ " "), haftmann@24219: pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), haftmann@24219: str " where {" haftmann@24219: ], haftmann@24219: str "};" haftmann@24841: ) (map pr_instdef classparam_insts) haftmann@24219: end; haftmann@27103: in pr_stmt end; haftmann@24219: haftmann@24992: fun pretty_haskell_monad c_bind = haftmann@24219: let haftmann@27304: fun pretty pr thm pat vars fxy [(t, _)] = haftmann@24219: let haftmann@27304: val pr_bind = pr_haskell_bind (K (K pr)) thm; haftmann@24992: fun pr_monad (NONE, t) vars = haftmann@24219: (semicolon [pr vars NOBR t], vars) haftmann@24992: | pr_monad (SOME (bind, true), t) vars = vars haftmann@24219: |> pr_bind NOBR bind haftmann@24219: |>> (fn p => semicolon [p, str "<-", pr vars NOBR t]) haftmann@24992: | pr_monad (SOME (bind, false), t) vars = vars haftmann@24219: |> pr_bind NOBR bind haftmann@24219: |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]); haftmann@24992: val (binds, t) = implode_monad c_bind t; haftmann@24992: val (ps, vars') = fold_map pr_monad binds vars; haftmann@26010: fun brack p = if fixity BR fxy then Pretty.block [str "(", p, str ")"] else p; haftmann@24219: in (brack o Pretty.block_enclose (str "do {", str "}")) (ps @| pr vars' NOBR t) end; haftmann@24219: in (1, pretty) end; haftmann@24219: haftmann@27103: fun haskell_program_of_program labelled_name module_name module_prefix reserved_names raw_module_alias program = haftmann@24219: let haftmann@27103: val module_alias = if is_some module_name then K module_name else raw_module_alias; haftmann@27103: val reserved_names = Name.make_context reserved_names; haftmann@27103: val mk_name_module = mk_name_module reserved_names module_prefix module_alias program; haftmann@27103: fun add_stmt (name, (stmt, deps)) = haftmann@24219: let haftmann@27103: val (module_name, base) = dest_name name; haftmann@27103: val module_name' = mk_name_module module_name; haftmann@27103: val mk_name_stmt = yield_singleton Name.variants; haftmann@24219: fun add_fun upper (nsp_fun, nsp_typ) = haftmann@24219: let haftmann@24992: val (base', nsp_fun') = haftmann@27103: mk_name_stmt (if upper then first_upper base else base) nsp_fun haftmann@24219: in (base', (nsp_fun', nsp_typ)) end; haftmann@24219: fun add_typ (nsp_fun, nsp_typ) = haftmann@24219: let haftmann@27103: val (base', nsp_typ') = mk_name_stmt (first_upper base) nsp_typ haftmann@24219: in (base', (nsp_fun, nsp_typ')) end; haftmann@27103: val add_name = case stmt haftmann@27103: of CodeThingol.Fun _ => add_fun false haftmann@27103: | CodeThingol.Datatype _ => add_typ haftmann@27103: | CodeThingol.Datatypecons _ => add_fun true haftmann@27103: | CodeThingol.Class _ => add_typ haftmann@27103: | CodeThingol.Classrel _ => pair base haftmann@27103: | CodeThingol.Classparam _ => add_fun false haftmann@27103: | CodeThingol.Classinst _ => pair base; haftmann@27103: fun add_stmt' base' = case stmt haftmann@27103: of CodeThingol.Datatypecons _ => haftmann@27103: cons (name, (NameSpace.append module_name' base', NONE)) haftmann@27103: | CodeThingol.Classrel _ => I haftmann@27103: | CodeThingol.Classparam _ => haftmann@27103: cons (name, (NameSpace.append module_name' base', NONE)) haftmann@27103: | _ => cons (name, (NameSpace.append module_name' base', SOME stmt)); haftmann@24219: in haftmann@27103: Symtab.map_default (module_name', ([], ([], (reserved_names, reserved_names)))) haftmann@24219: (apfst (fold (insert (op = : string * string -> bool)) deps)) haftmann@27103: #> `(fn program => add_name ((snd o snd o the o Symtab.lookup program) module_name')) haftmann@24219: #-> (fn (base', names) => haftmann@27103: (Symtab.map_entry module_name' o apsnd) (fn (stmts, _) => haftmann@27103: (add_stmt' base' stmts, names))) haftmann@24219: end; haftmann@27103: val hs_program = fold add_stmt (AList.make (fn name => haftmann@27103: (Graph.get_node program name, Graph.imm_succs program name)) haftmann@27103: (Graph.strong_conn program |> flat)) Symtab.empty; haftmann@27103: fun deresolver name = haftmann@27103: (fst o the o AList.lookup (op =) ((fst o snd o the haftmann@27103: o Symtab.lookup hs_program) ((mk_name_module o fst o dest_name) name))) name haftmann@27024: handle Option => error ("Unknown statement name: " ^ labelled_name name); haftmann@27103: in (deresolver, hs_program) end; haftmann@27103: haftmann@27304: fun serialize_haskell module_prefix raw_module_name string_classes labelled_name haftmann@27103: reserved_names includes raw_module_alias haftmann@27103: syntax_class syntax_tyco syntax_const program cs destination = haftmann@27103: let haftmann@27304: val stmt_names = stmt_names_of_destination destination; haftmann@27304: val module_name = if null stmt_names then raw_module_name else SOME "Code"; haftmann@27103: val (deresolver, hs_program) = haskell_program_of_program labelled_name haftmann@27103: module_name module_prefix reserved_names raw_module_alias program; haftmann@27103: val is_cons = CodeThingol.is_cons program; haftmann@27103: val contr_classparam_typs = CodeThingol.contr_classparam_typs program; haftmann@24219: fun deriving_show tyco = haftmann@24219: let haftmann@24219: fun deriv _ "fun" = false haftmann@24219: | deriv tycos tyco = member (op =) tycos tyco orelse haftmann@27103: case try (Graph.get_node program) tyco haftmann@27024: of SOME (CodeThingol.Datatype (_, cs)) => forall (deriv' (tyco :: tycos)) haftmann@24219: (maps snd cs) haftmann@27024: | NONE => true haftmann@24219: and deriv' tycos (tyco `%% tys) = deriv tycos tyco haftmann@24219: andalso forall (deriv' tycos) tys haftmann@24219: | deriv' _ (ITyVar _) = true haftmann@24219: in deriv [] tyco end; haftmann@27103: val reserved_names = CodeName.make_vars reserved_names; haftmann@27103: fun pr_stmt qualified = pr_haskell_stmt syntax_class syntax_tyco haftmann@27103: syntax_const labelled_name reserved_names haftmann@27103: (if qualified then deresolver else NameSpace.base o deresolver) haftmann@27103: is_cons contr_classparam_typs haftmann@24219: (if string_classes then deriving_show else K false); haftmann@27103: fun pr_module name content = haftmann@27103: (name, Pretty.chunks [ haftmann@27103: str ("module " ^ name ^ " where {"), haftmann@24992: str "", haftmann@24992: content, haftmann@24992: str "", haftmann@24992: str "}" haftmann@27001: ]); haftmann@27304: fun serialize_module1 (module_name', (deps, (stmts, _))) = haftmann@24219: let haftmann@27103: val stmt_names = map fst stmts; haftmann@27103: val deps' = subtract (op =) stmt_names deps haftmann@24219: |> distinct (op =) haftmann@27103: |> map_filter (try deresolver); haftmann@27103: val qualified = is_none module_name andalso haftmann@27103: map deresolver stmt_names @ deps' haftmann@24219: |> map NameSpace.base haftmann@24219: |> has_duplicates (op =); haftmann@27103: val imports = deps' haftmann@27103: |> map NameSpace.qualifier haftmann@27103: |> distinct (op =); haftmann@27103: fun pr_import_include (name, _) = str ("import " ^ name ^ ";"); haftmann@27103: val pr_import_module = str o (if qualified haftmann@24219: then prefix "import qualified " haftmann@24219: else prefix "import ") o suffix ";"; haftmann@24992: val content = Pretty.chunks ( haftmann@27103: map pr_import_include includes haftmann@27103: @ map pr_import_module imports haftmann@24992: @ str "" haftmann@24992: :: separate (str "") (map_filter haftmann@27103: (fn (name, (_, SOME stmt)) => SOME (pr_stmt qualified (name, stmt)) haftmann@27103: | (_, (_, NONE)) => NONE) stmts) haftmann@24992: ) haftmann@27103: in pr_module module_name' content end; haftmann@27304: fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks ( haftmann@27304: separate (str "") (map_filter haftmann@27304: (fn (name, (_, SOME stmt)) => if null stmt_names haftmann@27304: orelse member (op =) stmt_names name haftmann@27304: then SOME (pr_stmt false (name, stmt)) haftmann@27304: else NONE haftmann@27304: | (_, (_, NONE)) => NONE) stmts)); haftmann@27304: val serialize_module = case destination of String _ => pair "" o serialize_module2 haftmann@27304: | _ => serialize_module1; haftmann@27000: fun write_module destination (modlname, content) = haftmann@27000: let haftmann@27000: val filename = case modlname haftmann@27000: of "" => Path.explode "Main.hs" haftmann@27000: | _ => (Path.ext "hs" o Path.explode o implode o separate "/" haftmann@27000: o NameSpace.explode) modlname; haftmann@27000: val pathname = Path.append destination filename; haftmann@27000: val _ = File.mkdir (Path.dir pathname); haftmann@27103: in File.write pathname (code_of_pretty content) end haftmann@27103: fun output Compile = error ("Haskell: no internal compilation") haftmann@27304: | output Export = K NONE o map (code_writeln o snd) haftmann@27103: | output (File destination) = K NONE o map (write_module destination) haftmann@27436: | output (String _) = SOME o rpair [] o cat_lines o map (code_of_pretty o snd); haftmann@27103: in haftmann@27103: output destination (map (uncurry pr_module) includes haftmann@27103: @ map serialize_module (Symtab.dest hs_program)) haftmann@27000: end; haftmann@24219: haftmann@27000: fun isar_seri_haskell module = haftmann@27000: parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name) haftmann@27000: -- Scan.optional (Args.$$$ "string_classes" >> K true) false haftmann@27000: >> (fn (module_prefix, string_classes) => haftmann@27103: serialize_haskell module_prefix module string_classes)); haftmann@24219: haftmann@24219: haftmann@24219: (** optional pretty serialization **) haftmann@24219: haftmann@24219: local haftmann@24219: haftmann@26448: fun ocaml_char c = haftmann@26448: let haftmann@26448: fun chr i = haftmann@26448: let haftmann@26448: val xs = string_of_int i; haftmann@26448: val ys = replicate_string (3 - length (explode xs)) "0"; haftmann@26448: in "\\" ^ ys ^ xs end; haftmann@26448: val i = ord c; haftmann@26448: val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126 haftmann@26448: then chr i else c haftmann@26448: in s end; haftmann@26448: haftmann@26448: fun haskell_char c = haftmann@26448: let haftmann@26448: val s = ML_Syntax.print_char c; haftmann@26448: in if s = "'" then "\\'" else s end; haftmann@26448: haftmann@24219: val pretty : (string * { haftmann@24219: pretty_char: string -> string, haftmann@24219: pretty_string: string -> string, wenzelm@24630: pretty_numeral: bool -> int -> string, haftmann@24219: pretty_list: Pretty.T list -> Pretty.T, haftmann@24219: infix_cons: int * string haftmann@24219: }) list = [ haftmann@24219: ("SML", { pretty_char = prefix "#" o quote o ML_Syntax.print_char, haftmann@26448: pretty_string = quote o translate_string ML_Syntax.print_char, haftmann@24219: pretty_numeral = fn unbounded => fn k => haftmann@24750: if unbounded then "(" ^ string_of_int k ^ " : IntInf.int)" wenzelm@24630: else string_of_int k, haftmann@24219: pretty_list = Pretty.enum "," "[" "]", haftmann@24219: infix_cons = (7, "::")}), haftmann@26448: ("OCaml", { pretty_char = enclose "'" "'" o ocaml_char, haftmann@26448: pretty_string = quote o translate_string ocaml_char, wenzelm@24630: pretty_numeral = fn unbounded => fn k => if k >= 0 then haftmann@24219: if unbounded then wenzelm@24630: "(Big_int.big_int_of_int " ^ string_of_int k ^ ")" wenzelm@24630: else string_of_int k haftmann@24219: else haftmann@24219: if unbounded then haftmann@24219: "(Big_int.big_int_of_int " ^ (enclose "(" ")" o prefix "-" wenzelm@24630: o string_of_int o op ~) k ^ ")" wenzelm@24630: else (enclose "(" ")" o prefix "-" o string_of_int o op ~) k, haftmann@24219: pretty_list = Pretty.enum ";" "[" "]", haftmann@24219: infix_cons = (6, "::")}), haftmann@26448: ("Haskell", { pretty_char = enclose "'" "'" o haskell_char, haftmann@26448: pretty_string = quote o translate_string haskell_char, wenzelm@24630: pretty_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k wenzelm@24630: else enclose "(" ")" (signed_string_of_int k), haftmann@24219: pretty_list = Pretty.enum "," "[" "]", haftmann@24219: infix_cons = (5, ":")}) haftmann@24219: ]; haftmann@24219: haftmann@24219: in haftmann@24219: haftmann@24219: fun pr_pretty target = case AList.lookup (op =) pretty target haftmann@24219: of SOME x => x haftmann@24219: | NONE => error ("Unknown code target language: " ^ quote target); haftmann@24219: haftmann@24219: fun default_list (target_fxy, target_cons) pr fxy t1 t2 = haftmann@24219: brackify_infix (target_fxy, R) fxy [ haftmann@24219: pr (INFX (target_fxy, X)) t1, haftmann@24219: str target_cons, haftmann@24219: pr (INFX (target_fxy, R)) t2 haftmann@24219: ]; haftmann@24219: haftmann@24219: fun pretty_list c_nil c_cons target = haftmann@24219: let haftmann@24219: val pretty_ops = pr_pretty target; haftmann@24219: val mk_list = #pretty_list pretty_ops; haftmann@27304: fun pretty pr thm pat vars fxy [(t1, _), (t2, _)] = haftmann@24219: case Option.map (cons t1) (implode_list c_nil c_cons t2) haftmann@24219: of SOME ts => mk_list (map (pr vars NOBR) ts) haftmann@24219: | NONE => default_list (#infix_cons pretty_ops) (pr vars) fxy t1 t2; haftmann@24219: in (2, pretty) end; haftmann@24219: haftmann@24219: fun pretty_list_string c_nil c_cons c_char c_nibbles target = haftmann@24219: let haftmann@24219: val pretty_ops = pr_pretty target; haftmann@24219: val mk_list = #pretty_list pretty_ops; haftmann@24219: val mk_char = #pretty_char pretty_ops; haftmann@24219: val mk_string = #pretty_string pretty_ops; haftmann@27304: fun pretty pr thm pat vars fxy [(t1, _), (t2, _)] = haftmann@24219: case Option.map (cons t1) (implode_list c_nil c_cons t2) haftmann@26448: of SOME ts => (case implode_string c_char c_nibbles mk_char mk_string ts haftmann@24219: of SOME p => p haftmann@26448: | NONE => mk_list (map (pr vars NOBR) ts)) haftmann@24219: | NONE => default_list (#infix_cons pretty_ops) (pr vars) fxy t1 t2; haftmann@24219: in (2, pretty) end; haftmann@24219: haftmann@24219: fun pretty_char c_char c_nibbles target = haftmann@24219: let haftmann@24219: val mk_char = #pretty_char (pr_pretty target); haftmann@27304: fun pretty _ thm _ _ _ [(t1, _), (t2, _)] = haftmann@24219: case decode_char c_nibbles (t1, t2) haftmann@24219: of SOME c => (str o mk_char) c haftmann@27304: | NONE => nerror thm "Illegal character expression"; haftmann@24219: in (2, pretty) end; haftmann@24219: huffman@26086: fun pretty_numeral unbounded negative c_pls c_min c_bit0 c_bit1 target = haftmann@24219: let haftmann@24219: val mk_numeral = #pretty_numeral (pr_pretty target); haftmann@27304: fun pretty _ thm _ _ _ [(t, _)] = haftmann@27304: (str o mk_numeral unbounded o implode_numeral thm negative c_pls c_min c_bit0 c_bit1) t; haftmann@24219: in (1, pretty) end; haftmann@24219: haftmann@24992: fun pretty_message c_char c_nibbles c_nil c_cons target = haftmann@24219: let haftmann@24219: val pretty_ops = pr_pretty target; haftmann@24219: val mk_char = #pretty_char pretty_ops; haftmann@24219: val mk_string = #pretty_string pretty_ops; haftmann@27304: fun pretty _ thm _ _ _ [(t, _)] = haftmann@24219: case implode_list c_nil c_cons t haftmann@24219: of SOME ts => (case implode_string c_char c_nibbles mk_char mk_string ts haftmann@24219: of SOME p => p haftmann@27304: | NONE => nerror thm "Illegal message expression") haftmann@27304: | NONE => nerror thm "Illegal message expression"; haftmann@24219: in (1, pretty) end; haftmann@24219: haftmann@24219: end; (*local*) haftmann@24219: haftmann@27000: haftmann@27304: (** serializer use cases **) haftmann@27000: haftmann@27000: (* evaluation *) haftmann@27000: haftmann@27103: fun eval eval'' term_of reff thy ct args = haftmann@27000: let haftmann@27103: val _ = if null (term_frees (term_of ct)) then () else error ("Term " haftmann@27103: ^ quote (Syntax.string_of_term_global thy (term_of ct)) haftmann@27103: ^ " to be evaluated contains free variables"); haftmann@27103: fun eval' program ((vs, ty), t) deps = haftmann@27103: let haftmann@27103: val _ = if CodeThingol.contains_dictvar t then haftmann@27103: error "Term to be evaluated constains free dictionaries" else (); haftmann@27103: val program' = program haftmann@27103: |> Graph.new_node (CodeName.value_name, CodeThingol.Fun (([], ty), [(([], t), Drule.dummy_thm)])) haftmann@27103: |> fold (curry Graph.add_edge CodeName.value_name) deps; haftmann@27436: val (value_code, [value_name']) = ml_code_of thy program' [CodeName.value_name]; haftmann@27436: val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name' haftmann@27436: ^ space_implode " " (map (enclose "(" ")") args) ^ " end"; haftmann@27103: in ML_Context.evaluate Output.ml_output false reff sml_code end; haftmann@27103: in eval'' thy (fn t => (t, eval')) ct end; haftmann@27103: haftmann@27103: fun eval_conv reff = eval CodeThingol.eval_conv Thm.term_of reff; haftmann@27103: fun eval_term reff = eval CodeThingol.eval_term I reff; haftmann@27103: haftmann@27103: haftmann@27304: (* instrumentalization by antiquotation *) haftmann@27103: haftmann@27436: local haftmann@27436: haftmann@27436: structure CodeAntiqData = ProofDataFun haftmann@27436: ( haftmann@27437: type T = string list * (bool * (string * (string * (string * string) list) Susp.T)); haftmann@27437: fun init _ = ([], (true, ("", Susp.value ("", [])))); haftmann@27436: ); haftmann@27436: haftmann@27436: val is_first_occ = fst o snd o CodeAntiqData.get; haftmann@27436: haftmann@27437: fun delayed_code thy consts () = haftmann@27437: let haftmann@27437: val (consts', program) = CodeThingol.consts_program thy consts; haftmann@27437: val (ml_code, consts'') = ml_code_of thy program consts'; haftmann@27437: val _ = if length consts <> length consts'' then haftmann@27437: error ("One of the constants " ^ commas (map (quote o CodeUnit.string_of_const thy) consts) haftmann@27437: ^ "\nhas a user-defined serialization") else (); haftmann@27437: in (ml_code, consts ~~ consts'') end; haftmann@27437: haftmann@27436: fun register_const const ctxt = haftmann@27436: let haftmann@27437: val (consts, (_, (struct_name, _))) = CodeAntiqData.get ctxt; haftmann@27437: val consts' = insert (op =) const consts; haftmann@27436: val (struct_name', ctxt') = if struct_name = "" haftmann@27436: then ML_Antiquote.variant "Code" ctxt haftmann@27436: else (struct_name, ctxt); haftmann@27437: val acc_code = Susp.delay (delayed_code (ProofContext.theory_of ctxt) consts'); haftmann@27437: in CodeAntiqData.put (consts', (false, (struct_name', acc_code))) ctxt' end; haftmann@27436: haftmann@27437: fun print_code struct_name is_first const ctxt = haftmann@27304: let haftmann@27437: val (consts, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; haftmann@27437: val (raw_ml_code, consts_map) = Susp.force acc_code; haftmann@27436: val const'' = NameSpace.append (NameSpace.append struct_name struct_code_name) haftmann@27437: ((the o AList.lookup (op =) consts_map) const); haftmann@27436: val ml_code = if is_first then "\nstructure " ^ struct_code_name haftmann@27436: ^ " =\nstruct\n\n" ^ raw_ml_code ^ "\nend;\n\n" haftmann@27436: else ""; haftmann@27436: in (ml_code, const'') end; haftmann@27436: haftmann@27436: in haftmann@27436: haftmann@27436: fun ml_code_antiq raw_const {struct_name, background} = haftmann@27436: let haftmann@27437: val const = CodeUnit.check_const (ProofContext.theory_of background) raw_const; haftmann@27436: val is_first = is_first_occ background; haftmann@27436: val background' = register_const const background; haftmann@27437: in (print_code struct_name is_first const, background') end; haftmann@27436: haftmann@27436: end; (*local*) haftmann@27304: haftmann@27304: haftmann@27304: (* code presentation *) haftmann@27304: haftmann@27304: fun code_of thy target module_name cs stmt_names = haftmann@27103: let haftmann@27103: val (cs', program) = CodeThingol.consts_program thy cs; haftmann@27103: in haftmann@27304: string stmt_names (serialize thy target (SOME module_name) [] program cs') haftmann@27103: end; haftmann@27000: haftmann@27000: haftmann@27304: (* code generation *) haftmann@27304: haftmann@27304: fun read_const_exprs thy cs = haftmann@27304: let haftmann@27304: val (cs1, cs2) = CodeName.read_const_exprs thy cs; haftmann@27304: val (cs3, program) = CodeThingol.consts_program thy cs2; haftmann@27304: val cs4 = CodeThingol.transitivly_non_empty_funs program (abort_allowed thy); haftmann@27304: val cs5 = map_filter haftmann@27304: (fn (c, c') => if member (op =) cs4 c' then SOME c else NONE) (cs2 ~~ cs3); haftmann@27304: in fold (insert (op =)) cs5 cs1 end; haftmann@27304: haftmann@27304: fun cached_program thy = haftmann@27304: let haftmann@27304: val program = CodeThingol.cached_program thy; haftmann@27304: in (CodeThingol.transitivly_non_empty_funs program (abort_allowed thy), program) end haftmann@27304: haftmann@27304: fun export_code thy cs seris = haftmann@27304: let haftmann@27304: val (cs', program) = if null cs then cached_program thy haftmann@27304: else CodeThingol.consts_program thy cs; haftmann@27304: fun mk_seri_dest dest = case dest haftmann@27304: of NONE => compile haftmann@27304: | SOME "-" => export haftmann@27304: | SOME f => file (Path.explode f) haftmann@27304: val _ = map (fn (((target, module), dest), args) => haftmann@27304: (mk_seri_dest dest (serialize thy target module args program cs'))) seris; haftmann@27304: in () end; haftmann@27304: haftmann@27304: fun export_code_cmd raw_cs seris thy = export_code thy (read_const_exprs thy raw_cs) seris; haftmann@27304: haftmann@27304: haftmann@27304: (** serializer data **) haftmann@27304: haftmann@27000: (* infix syntax *) haftmann@27000: haftmann@27000: datatype 'a mixfix = haftmann@27000: Arg of fixity haftmann@27000: | Pretty of Pretty.T; haftmann@27000: haftmann@27000: fun mk_mixfix prep_arg (fixity_this, mfx) = haftmann@27000: let haftmann@27000: fun is_arg (Arg _) = true haftmann@27000: | is_arg _ = false; haftmann@27000: val i = (length o filter is_arg) mfx; haftmann@27000: fun fillin _ [] [] = haftmann@27000: [] haftmann@27000: | fillin pr (Arg fxy :: mfx) (a :: args) = haftmann@27000: (pr fxy o prep_arg) a :: fillin pr mfx args haftmann@27000: | fillin pr (Pretty p :: mfx) args = haftmann@27304: p :: fillin pr mfx args; haftmann@27000: in haftmann@27000: (i, fn pr => fn fixity_ctxt => fn args => haftmann@27000: gen_brackify (fixity fixity_this fixity_ctxt) (fillin pr mfx args)) haftmann@27000: end; haftmann@27000: haftmann@27000: fun parse_infix prep_arg (x, i) s = haftmann@27000: let haftmann@27000: val l = case x of L => INFX (i, L) | _ => INFX (i, X); haftmann@27000: val r = case x of R => INFX (i, R) | _ => INFX (i, X); haftmann@27000: in haftmann@27000: mk_mixfix prep_arg (INFX (i, x), haftmann@27000: [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r]) haftmann@27000: end; haftmann@27000: haftmann@27000: haftmann@27000: (* data access *) haftmann@24219: haftmann@24219: local haftmann@24219: haftmann@24992: fun cert_class thy class = haftmann@24992: let haftmann@24992: val _ = AxClass.get_info thy class; haftmann@24992: in class end; haftmann@24992: haftmann@27103: fun read_class thy = cert_class thy o Sign.intern_class thy; haftmann@24992: haftmann@24992: fun cert_tyco thy tyco = haftmann@24992: let haftmann@24992: val _ = if Sign.declared_tyname thy tyco then () haftmann@24992: else error ("No such type constructor: " ^ quote tyco); haftmann@24992: in tyco end; haftmann@24992: haftmann@27103: fun read_tyco thy = cert_tyco thy o Sign.intern_type thy; haftmann@24219: haftmann@24219: fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy = haftmann@24219: let haftmann@24841: val class = prep_class thy raw_class; haftmann@24841: val class' = CodeName.class thy class; haftmann@24841: fun mk_classparam c = case AxClass.class_of_param thy c haftmann@24841: of SOME class'' => if class = class'' then CodeName.const thy c haftmann@24219: else error ("Not a class operation for class " ^ quote class ^ ": " ^ quote c) haftmann@24219: | NONE => error ("Not a class operation: " ^ quote c); haftmann@24841: fun mk_syntax_params raw_params = AList.lookup (op =) haftmann@24841: ((map o apfst) (mk_classparam o prep_const thy) raw_params); haftmann@24219: in case raw_syn haftmann@24841: of SOME (syntax, raw_params) => haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apfst o apfst) haftmann@24841: (Symtab.update (class', (syntax, mk_syntax_params raw_params))) haftmann@24219: | NONE => haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apfst o apfst) haftmann@24841: (Symtab.delete_safe class') haftmann@24219: end; haftmann@24219: haftmann@24219: fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) add_del thy = haftmann@24219: let haftmann@24219: val inst = CodeName.instance thy (prep_class thy raw_class, prep_tyco thy raw_tyco); haftmann@24219: in if add_del then haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apfst o apsnd) haftmann@24219: (Symtab.update (inst, ())) haftmann@24219: else haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apfst o apsnd) haftmann@24219: (Symtab.delete_safe inst) haftmann@24219: end; haftmann@24219: haftmann@24219: fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy = haftmann@24219: let haftmann@24219: val tyco = prep_tyco thy raw_tyco; haftmann@24219: val tyco' = if tyco = "fun" then "fun" else CodeName.tyco thy tyco; haftmann@24219: fun check_args (syntax as (n, _)) = if n <> Sign.arity_number thy tyco haftmann@24219: then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco) haftmann@24219: else syntax haftmann@24219: in case raw_syn haftmann@24219: of SOME syntax => haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apsnd o apfst) haftmann@24219: (Symtab.update (tyco', check_args syntax)) haftmann@24219: | NONE => haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apsnd o apfst) haftmann@24219: (Symtab.delete_safe tyco') haftmann@24219: end; haftmann@24219: haftmann@27304: fun simple_const_syntax x = (Option.map o apsnd) haftmann@27304: (fn pretty => fn pr => fn thm => fn pat => fn vars => pretty (pr vars)) x; haftmann@27304: haftmann@24219: fun gen_add_syntax_const prep_const target raw_c raw_syn thy = haftmann@24219: let haftmann@24219: val c = prep_const thy raw_c; haftmann@24219: val c' = CodeName.const thy c; haftmann@24423: fun check_args (syntax as (n, _)) = if n > CodeUnit.no_args thy c haftmann@24423: then error ("Too many arguments in syntax for constant " ^ quote c) haftmann@24219: else syntax; haftmann@24219: in case raw_syn haftmann@24219: of SOME syntax => haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apsnd o apsnd) haftmann@24219: (Symtab.update (c', check_args syntax)) haftmann@24219: | NONE => haftmann@24219: thy haftmann@27024: |> (map_name_syntax target o apsnd o apsnd) haftmann@24219: (Symtab.delete_safe c') haftmann@24219: end; haftmann@24219: haftmann@24219: fun add_reserved target = haftmann@24219: let haftmann@24219: fun add sym syms = if member (op =) syms sym haftmann@24219: then error ("Reserved symbol " ^ quote sym ^ " already declared") haftmann@24219: else insert (op =) sym syms haftmann@27103: in map_reserved target o add end; haftmann@24992: haftmann@24992: fun add_include target = haftmann@24992: let haftmann@24992: fun add (name, SOME content) incls = haftmann@24992: let haftmann@24992: val _ = if Symtab.defined incls name haftmann@24992: then warning ("Overwriting existing include " ^ name) haftmann@24992: else (); haftmann@24992: in Symtab.update (name, str content) incls end haftmann@24992: | add (name, NONE) incls = haftmann@24992: Symtab.delete name incls; haftmann@27103: in map_includes target o add end; haftmann@24219: haftmann@27103: fun add_module_alias target = haftmann@24992: map_module_alias target o Symtab.update o apsnd CodeName.check_modulename; haftmann@24219: haftmann@27710: fun add_monad target raw_c_run raw_c_bind thy = haftmann@24992: let haftmann@27103: val c_run = CodeUnit.read_const thy raw_c_run; haftmann@27103: val c_bind = CodeUnit.read_const thy raw_c_bind; haftmann@27103: val c_bind' = CodeName.const thy c_bind; haftmann@24992: in if target = target_Haskell then haftmann@24992: thy haftmann@27103: |> gen_add_syntax_const (K I) target_Haskell c_run haftmann@27103: (SOME (pretty_haskell_monad c_bind')) haftmann@27103: |> gen_add_syntax_const (K I) target_Haskell c_bind haftmann@27304: (simple_const_syntax (SOME (parse_infix fst (L, 1) ">>="))) haftmann@27710: else error "Only Haskell target allows for monad syntax" end; haftmann@24219: haftmann@27103: fun gen_allow_abort prep_cs raw_c thy = haftmann@24841: let haftmann@24841: val c = prep_cs thy raw_c; haftmann@24841: val c' = CodeName.const thy c; haftmann@24841: in thy |> (CodeTargetData.map o apsnd) (insert (op =) c') end; haftmann@24841: haftmann@24219: fun zip_list (x::xs) f g = haftmann@24219: f haftmann@24219: #-> (fn y => haftmann@24219: fold_map (fn x => g |-- f >> pair x) xs haftmann@24219: #-> (fn xys => pair ((x, y) :: xys))); haftmann@24219: haftmann@27000: haftmann@27103: (* concrete syntax *) haftmann@27000: haftmann@24219: structure P = OuterParse haftmann@24219: and K = OuterKeyword haftmann@24219: haftmann@24219: fun parse_multi_syntax parse_thing parse_syntax = haftmann@24219: P.and_list1 parse_thing haftmann@24219: #-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name -- haftmann@24219: (zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")")); haftmann@24219: haftmann@24219: val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr"); haftmann@24219: haftmann@27000: fun parse_mixfix prep_arg s = haftmann@27000: let haftmann@27000: val sym_any = Scan.one Symbol.is_regular; haftmann@27000: val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat ( haftmann@27000: ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR)) haftmann@27000: || ($$ "_" >> K (Arg BR)) haftmann@27000: || ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length)) haftmann@27000: || (Scan.repeat1 haftmann@27000: ( $$ "'" |-- sym_any haftmann@27000: || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")") haftmann@27000: sym_any) >> (Pretty o str o implode))); haftmann@27000: in case Scan.finite Symbol.stopper parse (Symbol.explode s) haftmann@27000: of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p) haftmann@27000: | ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p) haftmann@27000: | _ => Scan.!! haftmann@27000: (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail () haftmann@27000: end; haftmann@27000: haftmann@24219: fun parse_syntax prep_arg xs = haftmann@24219: Scan.option (( haftmann@24219: ((P.$$$ infixK >> K X) haftmann@24219: || (P.$$$ infixlK >> K L) haftmann@24219: || (P.$$$ infixrK >> K R)) haftmann@24219: -- P.nat >> parse_infix prep_arg haftmann@24219: || Scan.succeed (parse_mixfix prep_arg)) haftmann@24219: -- P.string haftmann@24219: >> (fn (parse, s) => parse s)) xs; haftmann@24219: haftmann@24219: in haftmann@24219: haftmann@24219: val parse_syntax = parse_syntax; haftmann@24219: haftmann@24219: val add_syntax_class = gen_add_syntax_class cert_class (K I); haftmann@24219: val add_syntax_inst = gen_add_syntax_inst cert_class cert_tyco; haftmann@24219: val add_syntax_tyco = gen_add_syntax_tyco cert_tyco; haftmann@24219: val add_syntax_const = gen_add_syntax_const (K I); haftmann@27103: val allow_abort = gen_allow_abort (K I); haftmann@24219: haftmann@24219: val add_syntax_class_cmd = gen_add_syntax_class read_class CodeUnit.read_const; haftmann@24219: val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco; haftmann@24219: val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco; haftmann@24219: val add_syntax_const_cmd = gen_add_syntax_const CodeUnit.read_const; haftmann@27103: val allow_abort_cmd = gen_allow_abort CodeUnit.read_const; haftmann@24219: haftmann@24219: fun add_syntax_tycoP target tyco = parse_syntax I >> add_syntax_tyco_cmd target tyco; haftmann@27304: fun add_syntax_constP target c = parse_syntax fst >> (add_syntax_const_cmd target c o simple_const_syntax); haftmann@24219: haftmann@24219: fun add_undefined target undef target_undefined thy = haftmann@24219: let haftmann@27304: fun pr _ _ _ _ _ _ = str target_undefined; haftmann@24219: in haftmann@24219: thy haftmann@24423: |> add_syntax_const target undef (SOME (~1, pr)) haftmann@24219: end; haftmann@24219: haftmann@24219: fun add_pretty_list target nill cons thy = haftmann@24219: let haftmann@24423: val nil' = CodeName.const thy nill; haftmann@24423: val cons' = CodeName.const thy cons; haftmann@24423: val pr = pretty_list nil' cons' target; haftmann@24219: in haftmann@24219: thy haftmann@24423: |> add_syntax_const target cons (SOME pr) haftmann@24219: end; haftmann@24219: haftmann@24219: fun add_pretty_list_string target nill cons charr nibbles thy = haftmann@24219: let haftmann@24423: val nil' = CodeName.const thy nill; haftmann@24423: val cons' = CodeName.const thy cons; haftmann@24423: val charr' = CodeName.const thy charr; haftmann@24423: val nibbles' = map (CodeName.const thy) nibbles; haftmann@24423: val pr = pretty_list_string nil' cons' charr' nibbles' target; haftmann@24219: in haftmann@24219: thy haftmann@24423: |> add_syntax_const target cons (SOME pr) haftmann@24219: end; haftmann@24219: haftmann@24219: fun add_pretty_char target charr nibbles thy = haftmann@24219: let haftmann@24423: val charr' = CodeName.const thy charr; haftmann@24423: val nibbles' = map (CodeName.const thy) nibbles; haftmann@24423: val pr = pretty_char charr' nibbles' target; haftmann@24219: in haftmann@24219: thy haftmann@24423: |> add_syntax_const target charr (SOME pr) haftmann@24219: end; haftmann@24219: huffman@26086: fun add_pretty_numeral target unbounded negative number_of pls min bit0 bit1 thy = haftmann@24219: let haftmann@24423: val pls' = CodeName.const thy pls; haftmann@24423: val min' = CodeName.const thy min; huffman@26086: val bit0' = CodeName.const thy bit0; huffman@26086: val bit1' = CodeName.const thy bit1; huffman@26086: val pr = pretty_numeral unbounded negative pls' min' bit0' bit1' target; haftmann@24219: in haftmann@24219: thy haftmann@24423: |> add_syntax_const target number_of (SOME pr) haftmann@24219: end; haftmann@24219: haftmann@24992: fun add_pretty_message target charr nibbles nill cons str thy = haftmann@24219: let haftmann@24423: val charr' = CodeName.const thy charr; haftmann@24423: val nibbles' = map (CodeName.const thy) nibbles; haftmann@24423: val nil' = CodeName.const thy nill; haftmann@24423: val cons' = CodeName.const thy cons; haftmann@24992: val pr = pretty_message charr' nibbles' nil' cons' target; haftmann@24219: in haftmann@24219: thy haftmann@24423: |> add_syntax_const target str (SOME pr) haftmann@24219: end; haftmann@24219: wenzelm@24867: haftmann@27103: haftmann@27304: (** Isar setup **) haftmann@27103: haftmann@27103: val (inK, module_nameK, fileK) = ("in", "module_name", "file"); haftmann@27103: haftmann@27103: fun code_exprP cmd = haftmann@27103: (Scan.repeat P.term haftmann@27103: -- Scan.repeat (P.$$$ inK |-- P.name haftmann@27103: -- Scan.option (P.$$$ module_nameK |-- P.name) haftmann@27103: -- Scan.option (P.$$$ fileK |-- P.name) haftmann@27103: -- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") [] haftmann@27103: ) >> (fn (raw_cs, seris) => cmd raw_cs seris)); haftmann@27103: wenzelm@27353: val _ = List.app OuterKeyword.keyword [infixK, infixlK, infixrK, inK, module_nameK, fileK]; wenzelm@24867: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_class" "define code syntax for class" K.thy_decl ( haftmann@24219: parse_multi_syntax P.xname haftmann@24219: (Scan.option (P.string -- Scan.optional (P.$$$ "where" |-- Scan.repeat1 wenzelm@25985: (P.term --| (P.$$$ "\" || P.$$$ "==") -- P.string)) [])) haftmann@24219: >> (Toplevel.theory oo fold) (fn (target, syns) => haftmann@24219: fold (fn (raw_class, syn) => add_syntax_class_cmd target raw_class syn) syns) haftmann@24219: ); haftmann@24219: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_instance" "define code syntax for instance" K.thy_decl ( haftmann@24219: parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname) haftmann@24219: ((P.minus >> K true) || Scan.succeed false) haftmann@24219: >> (Toplevel.theory oo fold) (fn (target, syns) => haftmann@24219: fold (fn (raw_inst, add_del) => add_syntax_inst_cmd target raw_inst add_del) syns) haftmann@24219: ); haftmann@24219: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_type" "define code syntax for type constructor" K.thy_decl ( haftmann@24219: parse_multi_syntax P.xname (parse_syntax I) haftmann@24219: >> (Toplevel.theory oo fold) (fn (target, syns) => haftmann@24219: fold (fn (raw_tyco, syn) => add_syntax_tyco_cmd target raw_tyco syn) syns) haftmann@24219: ); haftmann@24219: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_const" "define code syntax for constant" K.thy_decl ( haftmann@24219: parse_multi_syntax P.term (parse_syntax fst) haftmann@24219: >> (Toplevel.theory oo fold) (fn (target, syns) => haftmann@27304: fold (fn (raw_const, syn) => add_syntax_const_cmd target raw_const (simple_const_syntax syn)) syns) haftmann@24219: ); haftmann@24219: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_monad" "define code syntax for monads" K.thy_decl ( haftmann@27710: P.term -- P.term -- P.name haftmann@27710: >> (fn ((raw_run, raw_bind), target) => Toplevel.theory haftmann@27710: (add_monad target raw_run raw_bind)) haftmann@24219: ); haftmann@24219: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_reserved" "declare words as reserved for target language" K.thy_decl ( haftmann@24219: P.name -- Scan.repeat1 P.name haftmann@24219: >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds) haftmann@24841: ); haftmann@24219: wenzelm@24867: val _ = haftmann@24992: OuterSyntax.command "code_include" "declare piece of code to be included in generated code" K.thy_decl ( haftmann@24992: P.name -- P.name -- (P.text >> (fn "-" => NONE | s => SOME s)) haftmann@24992: >> (fn ((target, name), content) => (Toplevel.theory o add_include target) haftmann@24992: (name, content)) haftmann@24992: ); haftmann@24992: haftmann@24992: val _ = haftmann@24992: OuterSyntax.command "code_modulename" "alias module to other name" K.thy_decl ( haftmann@24219: P.name -- Scan.repeat1 (P.name -- P.name) haftmann@27103: >> (fn (target, modlnames) => (Toplevel.theory o fold (add_module_alias target)) modlnames) haftmann@27103: ); haftmann@27103: haftmann@27103: val _ = haftmann@27103: OuterSyntax.command "code_abort" "permit constant to be implemented as program abort" K.thy_decl ( haftmann@27103: Scan.repeat1 P.term >> (Toplevel.theory o fold allow_abort_cmd) haftmann@24841: ); haftmann@24219: wenzelm@24867: val _ = haftmann@27103: OuterSyntax.command "export_code" "generate executable code for constants" haftmann@27103: K.diag (P.!!! (code_exprP export_code_cmd) >> (fn f => Toplevel.keep (f o Toplevel.theory_of))); haftmann@27103: haftmann@27103: fun shell_command thyname cmd = Toplevel.program (fn _ => haftmann@27103: (use_thy thyname; case Scan.read OuterLex.stopper (P.!!! (code_exprP export_code_cmd)) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd) haftmann@27103: of SOME f => (writeln "Now generating code..."; f (theory thyname)) haftmann@27103: | NONE => error ("Bad directive " ^ quote cmd))) haftmann@27103: handle TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure; haftmann@27103: haftmann@27436: val _ = ML_Context.add_antiq "code" (Args.term >> ml_code_antiq); haftmann@24219: haftmann@24219: haftmann@27000: (* serializer setup, including serializer defaults *) haftmann@27000: haftmann@24219: val setup = haftmann@27103: add_target (target_SML, isar_seri_sml) haftmann@27103: #> add_target (target_OCaml, isar_seri_ocaml) haftmann@27103: #> add_target (target_Haskell, isar_seri_haskell) haftmann@24219: #> add_syntax_tyco "SML" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => haftmann@27024: brackify_infix (1, R) fxy [ haftmann@24219: pr_typ (INFX (1, X)) ty1, haftmann@24219: str "->", haftmann@24219: pr_typ (INFX (1, R)) ty2 haftmann@24219: ])) haftmann@24219: #> add_syntax_tyco "OCaml" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => haftmann@27024: brackify_infix (1, R) fxy [ haftmann@24219: pr_typ (INFX (1, X)) ty1, haftmann@24219: str "->", haftmann@24219: pr_typ (INFX (1, R)) ty2 haftmann@24219: ])) haftmann@24219: #> add_syntax_tyco "Haskell" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => haftmann@24219: brackify_infix (1, R) fxy [ haftmann@24219: pr_typ (INFX (1, X)) ty1, haftmann@24219: str "->", haftmann@24219: pr_typ (INFX (1, R)) ty2 haftmann@24219: ])) haftmann@24219: #> fold (add_reserved "SML") ML_Syntax.reserved_names haftmann@24219: #> fold (add_reserved "SML") haftmann@24219: ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)] haftmann@24219: #> fold (add_reserved "OCaml") [ haftmann@24219: "and", "as", "assert", "begin", "class", haftmann@24219: "constraint", "do", "done", "downto", "else", "end", "exception", haftmann@24219: "external", "false", "for", "fun", "function", "functor", "if", haftmann@24219: "in", "include", "inherit", "initializer", "lazy", "let", "match", "method", haftmann@24219: "module", "mutable", "new", "object", "of", "open", "or", "private", "rec", haftmann@24219: "sig", "struct", "then", "to", "true", "try", "type", "val", haftmann@24219: "virtual", "when", "while", "with" haftmann@24219: ] haftmann@24219: #> fold (add_reserved "OCaml") ["failwith", "mod"] haftmann@24219: #> fold (add_reserved "Haskell") [ haftmann@24219: "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", haftmann@24219: "import", "default", "forall", "let", "in", "class", "qualified", "data", haftmann@24219: "newtype", "instance", "if", "then", "else", "type", "as", "do", "module" haftmann@24219: ] haftmann@24219: #> fold (add_reserved "Haskell") [ haftmann@24219: "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int", haftmann@24219: "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded", haftmann@24219: "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor", haftmann@24219: "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException", haftmann@24219: "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException", haftmann@24219: "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure", haftmann@24219: "ExitSuccess", "False", "GT", "HeapOverflow", haftmann@24219: "IOError", "IOException", "IllegalOperation", haftmann@24219: "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError", haftmann@24219: "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow", haftmann@24219: "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError", haftmann@24219: "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow", haftmann@24219: "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow", haftmann@24219: "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all", haftmann@24219: "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan", haftmann@24219: "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen", haftmann@24219: "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break", haftmann@24219: "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const", haftmann@24219: "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod", haftmann@24219: "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem", haftmann@24219: "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo", haftmann@24219: "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip", haftmann@24219: "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational", haftmann@24219: "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble", haftmann@24219: "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj", haftmann@24219: "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head", haftmann@24219: "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha", haftmann@24219: "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite", haftmann@24219: "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'", haftmann@24219: "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log", haftmann@24219: "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum", haftmann@24219: "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem", haftmann@24219: "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo", haftmann@24219: "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred", haftmann@24219: "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn", haftmann@24219: "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat", haftmann@24219: "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile", haftmann@24219: "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn", haftmann@24219: "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational", haftmann@24219: "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse", haftmann@24219: "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence", haftmann@24219: "sequence_", "show", "showChar", "showException", "showField", "showList", haftmann@24219: "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand", haftmann@24219: "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract", haftmann@24219: "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult", haftmann@24219: "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry", haftmann@24219: "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords", haftmann@24219: "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3" haftmann@24219: ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*); haftmann@24219: haftmann@24219: end; (*local*) haftmann@24219: haftmann@24219: end; (*struct*)