haftmann@28054: (* Title: Tools/code/code_ml.ML haftmann@28054: Author: Florian Haftmann, TU Muenchen haftmann@28054: haftmann@28054: Serializer for SML and OCaml. haftmann@28054: *) haftmann@28054: haftmann@28054: signature CODE_ML = haftmann@28054: sig haftmann@30947: val eval_term: string option -> string * (unit -> term) option ref haftmann@30947: -> theory -> term -> string list -> term haftmann@30947: val eval: string option -> string * (unit -> 'a) option ref haftmann@28724: -> theory -> term -> string list -> 'a haftmann@30947: val target_Eval: string haftmann@28054: val setup: theory -> theory haftmann@28054: end; haftmann@28054: haftmann@28054: structure Code_ML : CODE_ML = haftmann@28054: struct haftmann@28054: haftmann@28054: open Basic_Code_Thingol; haftmann@28054: open Code_Printer; haftmann@28054: haftmann@28054: infixr 5 @@; haftmann@28054: infixr 5 @|; haftmann@28054: haftmann@28054: val target_SML = "SML"; haftmann@28054: val target_OCaml = "OCaml"; haftmann@30947: val target_Eval = "Eval"; haftmann@28054: haftmann@28054: datatype ml_stmt = haftmann@29189: MLExc of string * int haftmann@29189: | MLVal of string * ((typscheme * iterm) * (thm * bool)) haftmann@29189: | MLFuns of (string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)) list * string list haftmann@28054: | MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list haftmann@28054: | MLClass of string * (vname * ((class * string) list * (string * itype) list)) haftmann@28054: | MLClassinst of string * ((class * (string * (vname * sort) list)) haftmann@28054: * ((class * (string * (string * dict list list))) list haftmann@28350: * ((string * const) * (thm * bool)) list)); haftmann@28054: haftmann@29189: fun stmt_names_of (MLExc (name, _)) = [name] haftmann@29189: | stmt_names_of (MLVal (name, _)) = [name] haftmann@29189: | stmt_names_of (MLFuns (fs, _)) = map fst fs haftmann@28054: | stmt_names_of (MLDatas ds) = map fst ds haftmann@29189: | stmt_names_of (MLClass (name, _)) = [name] haftmann@29189: | stmt_names_of (MLClassinst (name, _)) = [name]; haftmann@28054: haftmann@28054: haftmann@28054: (** SML serailizer **) haftmann@28054: haftmann@28663: fun pr_sml_stmt naming labelled_name syntax_tyco syntax_const reserved_names deresolve is_cons = haftmann@28054: let haftmann@28054: val pr_label_classrel = translate_string (fn "." => "__" | c => c) wenzelm@30364: o Long_Name.qualifier; wenzelm@30364: val pr_label_classparam = Long_Name.base_name o Long_Name.qualifier; haftmann@28054: fun pr_dicts fxy ds = haftmann@28054: let haftmann@30648: fun pr_dictvar (v, (_, 1)) = Code_Printer.first_upper v ^ "_" haftmann@30648: | pr_dictvar (v, (i, _)) = Code_Printer.first_upper v ^ string_of_int (i+1) ^ "_"; haftmann@28054: fun pr_proj [] p = haftmann@28054: p haftmann@28054: | pr_proj [p'] p = haftmann@28054: brackets [p', p] haftmann@28054: | pr_proj (ps as _ :: _) p = haftmann@28054: brackets [Pretty.enum " o" "(" ")" ps, p]; haftmann@28054: fun pr_dict fxy (DictConst (inst, dss)) = haftmann@28054: brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss) haftmann@28054: | pr_dict fxy (DictVar (classrels, v)) = haftmann@28054: pr_proj (map (str o deresolve) classrels) ((str o pr_dictvar) v) haftmann@28054: in case ds haftmann@28054: of [] => str "()" haftmann@28054: | [d] => pr_dict fxy d haftmann@28054: | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds haftmann@28054: end; haftmann@28054: fun pr_tyvar_dicts vs = haftmann@28054: vs haftmann@28054: |> map (fn (v, sort) => map_index (fn (i, _) => haftmann@28054: DictVar ([], (v, (i, length sort)))) sort) haftmann@28054: |> map (pr_dicts BR); haftmann@28054: fun pr_tycoexpr fxy (tyco, tys) = haftmann@28054: let haftmann@28054: val tyco' = (str o deresolve) tyco haftmann@28054: in case map (pr_typ BR) tys haftmann@28054: of [] => tyco' haftmann@28054: | [p] => Pretty.block [p, Pretty.brk 1, tyco'] haftmann@28054: | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco'] haftmann@28054: end haftmann@28054: and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco haftmann@28054: of NONE => pr_tycoexpr fxy (tyco, tys) haftmann@28054: | SOME (i, pr) => pr pr_typ fxy tys) haftmann@28054: | pr_typ fxy (ITyVar v) = str ("'" ^ v); haftmann@29175: fun pr_term is_closure thm vars fxy (IConst c) = haftmann@29175: pr_app is_closure thm vars fxy (c, []) haftmann@29175: | pr_term is_closure thm vars fxy (IVar v) = haftmann@30648: str (Code_Printer.lookup_var vars v) haftmann@29175: | pr_term is_closure thm vars fxy (t as t1 `$ t2) = haftmann@28054: (case Code_Thingol.unfold_const_app t haftmann@29175: of SOME c_ts => pr_app is_closure thm vars fxy c_ts haftmann@29175: | NONE => brackify fxy haftmann@29175: [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2]) haftmann@29175: | pr_term is_closure thm vars fxy (t as _ `|-> _) = haftmann@28054: let haftmann@28054: val (binds, t') = Code_Thingol.unfold_abs t; haftmann@28054: fun pr ((v, pat), ty) = haftmann@29175: pr_bind is_closure thm NOBR ((SOME v, pat), ty) haftmann@28054: #>> (fn p => concat [str "fn", p, str "=>"]); haftmann@28054: val (ps, vars') = fold_map pr binds vars; haftmann@29175: in brackets (ps @ [pr_term is_closure thm vars' NOBR t']) end haftmann@29175: | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) = haftmann@28054: (case Code_Thingol.unfold_const_app t0 haftmann@28054: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@29175: then pr_case is_closure thm vars fxy cases haftmann@29175: else pr_app is_closure thm vars fxy c_ts haftmann@29175: | NONE => pr_case is_closure thm vars fxy cases) haftmann@29175: and pr_app' is_closure thm vars (app as ((c, (iss, tys)), ts)) = haftmann@29175: if is_cons c then haftmann@29175: let haftmann@29175: val k = length tys haftmann@29175: in if k < 2 then haftmann@29175: (str o deresolve) c :: map (pr_term is_closure thm vars BR) ts haftmann@29175: else if k = length ts then haftmann@29175: [(str o deresolve) c, Pretty.enum "," "(" ")" (map (pr_term is_closure thm vars NOBR) ts)] haftmann@29175: else [pr_term is_closure thm vars BR (Code_Thingol.eta_expand k app)] end haftmann@29175: else if is_closure c haftmann@29175: then (str o deresolve) c @@ str "()" haftmann@29175: else haftmann@28054: (str o deresolve) c haftmann@29175: :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts haftmann@29175: and pr_app is_closure thm vars = gen_pr_app (pr_app' is_closure) (pr_term is_closure) haftmann@29175: syntax_const naming thm vars haftmann@28054: and pr_bind' ((NONE, NONE), _) = str "_" haftmann@28054: | pr_bind' ((SOME v, NONE), _) = str v haftmann@28054: | pr_bind' ((NONE, SOME p), _) = p haftmann@28054: | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p] haftmann@29175: and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure) haftmann@29175: and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) = haftmann@28054: let haftmann@29952: val (binds, body) = Code_Thingol.unfold_let (ICase cases); haftmann@28054: fun pr ((pat, ty), t) vars = haftmann@28054: vars haftmann@29175: |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) haftmann@29175: |>> (fn p => semicolon [str "val", p, str "=", pr_term is_closure thm vars NOBR t]) haftmann@28054: val (ps, vars') = fold_map pr binds vars; haftmann@28054: in haftmann@28054: Pretty.chunks [ haftmann@28054: [str ("let"), Pretty.fbrk, Pretty.chunks ps] |> Pretty.block, haftmann@29952: [str ("in"), Pretty.fbrk, pr_term is_closure thm vars' NOBR body] |> Pretty.block, haftmann@28054: str ("end") haftmann@28054: ] haftmann@28054: end haftmann@29952: | pr_case is_closure thm vars fxy (((t, ty), clause :: clauses), _) = haftmann@28054: let haftmann@29952: fun pr delim (pat, body) = haftmann@28054: let haftmann@29175: val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars; haftmann@28054: in haftmann@29952: concat [str delim, p, str "=>", pr_term is_closure thm vars' NOBR body] haftmann@28054: end; haftmann@28054: in haftmann@28054: (Pretty.enclose "(" ")" o single o brackify fxy) ( haftmann@28054: str "case" haftmann@29952: :: pr_term is_closure thm vars NOBR t haftmann@29952: :: pr "of" clause haftmann@29952: :: map (pr "|") clauses haftmann@28054: ) haftmann@28054: end haftmann@29175: | pr_case is_closure thm vars fxy ((_, []), _) = str "raise Fail \"empty case\""; haftmann@29189: fun pr_stmt (MLExc (name, n)) = haftmann@29189: let haftmann@29189: val exc_str = wenzelm@30364: (ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name; haftmann@29189: in haftmann@29189: concat ( haftmann@29189: str (if n = 0 then "val" else "fun") haftmann@29189: :: (str o deresolve) name haftmann@29189: :: map str (replicate n "_") haftmann@29189: @ str "=" haftmann@29189: :: str "raise" haftmann@29189: :: str "(Fail" haftmann@29189: @@ str (exc_str ^ ")") haftmann@29189: ) haftmann@29189: end haftmann@29189: | pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) = haftmann@28054: let haftmann@29189: val consts = map_filter haftmann@29189: (fn c => if (is_some o syntax_const) c wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c) haftmann@29189: (Code_Thingol.fold_constnames (insert (op =)) t []); haftmann@29189: val vars = reserved_names haftmann@30648: |> Code_Printer.intro_vars consts; haftmann@29189: in haftmann@29189: concat [ haftmann@29189: str "val", haftmann@29189: (str o deresolve) name, haftmann@29189: str ":", haftmann@29189: pr_typ NOBR ty, haftmann@29189: str "=", haftmann@29189: pr_term (K false) thm vars NOBR t haftmann@29189: ] haftmann@29189: end haftmann@29189: | pr_stmt (MLFuns (funn :: funns, pseudo_funs)) = haftmann@29189: let haftmann@29189: fun pr_funn definer (name, ((vs, ty), eqs as eq :: eqs')) = haftmann@28054: let haftmann@29189: val vs_dict = filter_out (null o snd) vs; haftmann@29189: val shift = if null eqs' then I else haftmann@29189: map (Pretty.block o single o Pretty.block o single); haftmann@29189: fun pr_eq definer ((ts, t), (thm, _)) = haftmann@28054: let haftmann@29189: val consts = map_filter haftmann@29189: (fn c => if (is_some o syntax_const) c wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c) haftmann@29189: ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@29189: val vars = reserved_names haftmann@30648: |> Code_Printer.intro_vars consts haftmann@30648: |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) haftmann@29189: (insert (op =)) ts []); haftmann@28054: in haftmann@28054: concat ( haftmann@28054: str definer haftmann@28054: :: (str o deresolve) name haftmann@29189: :: (if member (op =) pseudo_funs name then [str "()"] haftmann@29189: else pr_tyvar_dicts vs_dict haftmann@29189: @ map (pr_term (member (op =) pseudo_funs) thm vars BR) ts) haftmann@28054: @ str "=" haftmann@29189: @@ pr_term (member (op =) pseudo_funs) thm vars NOBR t haftmann@28054: ) haftmann@28054: end haftmann@29189: in haftmann@29189: (Pretty.block o Pretty.fbreaks o shift) ( haftmann@29189: pr_eq definer eq haftmann@29189: :: map (pr_eq "|") eqs' haftmann@29189: ) haftmann@29189: end; haftmann@29189: fun pr_pseudo_fun name = concat [ haftmann@29189: str "val", haftmann@29189: (str o deresolve) name, haftmann@29189: str "=", haftmann@29189: (str o deresolve) name, haftmann@29189: str "();" haftmann@29189: ]; haftmann@29189: val (ps, p) = split_last (pr_funn "fun" funn :: map (pr_funn "and") funns); haftmann@29189: val pseudo_ps = map pr_pseudo_fun pseudo_funs; haftmann@29189: in Pretty.chunks (ps @ Pretty.block ([p, str ";"]) :: pseudo_ps) end haftmann@28054: | pr_stmt (MLDatas (datas as (data :: datas'))) = haftmann@28054: let haftmann@28054: fun pr_co (co, []) = haftmann@28054: str (deresolve co) haftmann@28054: | pr_co (co, tys) = haftmann@28054: concat [ haftmann@28054: str (deresolve co), haftmann@28054: str "of", haftmann@28054: Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys) haftmann@28054: ]; haftmann@28054: fun pr_data definer (tyco, (vs, [])) = haftmann@28054: concat ( haftmann@28054: str definer haftmann@28054: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@28054: :: str "=" haftmann@28054: @@ str "EMPTY__" haftmann@28054: ) haftmann@28054: | pr_data definer (tyco, (vs, cos)) = haftmann@28054: concat ( haftmann@28054: str definer haftmann@28054: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@28054: :: str "=" haftmann@28054: :: separate (str "|") (map pr_co cos) haftmann@28054: ); haftmann@28054: val (ps, p) = split_last haftmann@28054: (pr_data "datatype" data :: map (pr_data "and") datas'); haftmann@29189: in Pretty.chunks (ps @| Pretty.block ([p, str ";"])) end haftmann@28054: | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) = haftmann@28054: let haftmann@30648: val w = Code_Printer.first_upper v ^ "_"; haftmann@28054: fun pr_superclass_field (class, classrel) = haftmann@28054: (concat o map str) [ haftmann@28054: pr_label_classrel classrel, ":", "'" ^ v, deresolve class haftmann@28054: ]; haftmann@28054: fun pr_classparam_field (classparam, ty) = haftmann@28054: concat [ haftmann@28054: (str o pr_label_classparam) classparam, str ":", pr_typ NOBR ty haftmann@28054: ]; haftmann@28054: fun pr_classparam_proj (classparam, _) = haftmann@28054: semicolon [ haftmann@28054: str "fun", haftmann@28054: (str o deresolve) classparam, haftmann@28054: Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)], haftmann@28054: str "=", haftmann@28054: str ("#" ^ pr_label_classparam classparam), haftmann@28054: str w haftmann@28054: ]; haftmann@28054: fun pr_superclass_proj (_, classrel) = haftmann@28054: semicolon [ haftmann@28054: str "fun", haftmann@28054: (str o deresolve) classrel, haftmann@28054: Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)], haftmann@28054: str "=", haftmann@28054: str ("#" ^ pr_label_classrel classrel), haftmann@28054: str w haftmann@28054: ]; haftmann@28054: in haftmann@28054: Pretty.chunks ( haftmann@28054: concat [ haftmann@28054: str ("type '" ^ v), haftmann@28054: (str o deresolve) class, haftmann@28054: str "=", haftmann@28054: Pretty.enum "," "{" "};" ( haftmann@28054: map pr_superclass_field superclasses @ map pr_classparam_field classparams haftmann@28054: ) haftmann@28054: ] haftmann@28054: :: map pr_superclass_proj superclasses haftmann@28054: @ map pr_classparam_proj classparams haftmann@28054: ) haftmann@28054: end haftmann@28054: | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) = haftmann@28054: let haftmann@28054: fun pr_superclass (_, (classrel, dss)) = haftmann@28054: concat [ haftmann@28054: (str o pr_label_classrel) classrel, haftmann@28054: str "=", haftmann@28054: pr_dicts NOBR [DictConst dss] haftmann@28054: ]; haftmann@28350: fun pr_classparam ((classparam, c_inst), (thm, _)) = haftmann@28054: concat [ haftmann@28054: (str o pr_label_classparam) classparam, haftmann@28054: str "=", haftmann@29175: pr_app (K false) thm reserved_names NOBR (c_inst, []) haftmann@28054: ]; haftmann@28054: in haftmann@28054: semicolon ([ haftmann@28054: str (if null arity then "val" else "fun"), haftmann@28054: (str o deresolve) inst ] @ haftmann@28054: pr_tyvar_dicts arity @ [ haftmann@28054: str "=", haftmann@28054: Pretty.enum "," "{" "}" haftmann@28054: (map pr_superclass superarities @ map pr_classparam classparam_insts), haftmann@28054: str ":", haftmann@28054: pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) haftmann@28054: ]) haftmann@28054: end; haftmann@28054: in pr_stmt end; haftmann@28054: haftmann@28054: fun pr_sml_module name content = haftmann@28054: Pretty.chunks ( haftmann@28054: str ("structure " ^ name ^ " = ") haftmann@28054: :: str "struct" haftmann@28054: :: str "" haftmann@28054: :: content haftmann@28054: @ str "" haftmann@28054: @@ str ("end; (*struct " ^ name ^ "*)") haftmann@28054: ); haftmann@28054: haftmann@28064: val literals_sml = Literals { haftmann@28064: literal_char = prefix "#" o quote o ML_Syntax.print_char, haftmann@28064: literal_string = quote o translate_string ML_Syntax.print_char, haftmann@28064: literal_numeral = fn unbounded => fn k => haftmann@28064: if unbounded then "(" ^ string_of_int k ^ " : IntInf.int)" haftmann@28064: else string_of_int k, haftmann@28064: literal_list = Pretty.enum "," "[" "]", haftmann@28064: infix_cons = (7, "::") haftmann@28064: }; haftmann@28064: haftmann@28054: haftmann@28054: (** OCaml serializer **) haftmann@28054: haftmann@28663: fun pr_ocaml_stmt naming labelled_name syntax_tyco syntax_const reserved_names deresolve is_cons = haftmann@28054: let haftmann@28054: fun pr_dicts fxy ds = haftmann@28054: let haftmann@30648: fun pr_dictvar (v, (_, 1)) = "_" ^ Code_Printer.first_upper v haftmann@30648: | pr_dictvar (v, (i, _)) = "_" ^ Code_Printer.first_upper v ^ string_of_int (i+1); haftmann@28054: fun pr_proj ps p = haftmann@28054: fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p haftmann@28054: fun pr_dict fxy (DictConst (inst, dss)) = haftmann@28054: brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss) haftmann@28054: | pr_dict fxy (DictVar (classrels, v)) = haftmann@28054: pr_proj (map deresolve classrels) ((str o pr_dictvar) v) haftmann@28054: in case ds haftmann@28054: of [] => str "()" haftmann@28054: | [d] => pr_dict fxy d haftmann@28054: | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds haftmann@28054: end; haftmann@28054: fun pr_tyvar_dicts vs = haftmann@28054: vs haftmann@28054: |> map (fn (v, sort) => map_index (fn (i, _) => haftmann@28054: DictVar ([], (v, (i, length sort)))) sort) haftmann@28054: |> map (pr_dicts BR); haftmann@28054: fun pr_tycoexpr fxy (tyco, tys) = haftmann@28054: let haftmann@28054: val tyco' = (str o deresolve) tyco haftmann@28054: in case map (pr_typ BR) tys haftmann@28054: of [] => tyco' haftmann@28054: | [p] => Pretty.block [p, Pretty.brk 1, tyco'] haftmann@28054: | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco'] haftmann@28054: end haftmann@28054: and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco haftmann@28054: of NONE => pr_tycoexpr fxy (tyco, tys) haftmann@28054: | SOME (i, pr) => pr pr_typ fxy tys) haftmann@28054: | pr_typ fxy (ITyVar v) = str ("'" ^ v); haftmann@29175: fun pr_term is_closure thm vars fxy (IConst c) = haftmann@29175: pr_app is_closure thm vars fxy (c, []) haftmann@29175: | pr_term is_closure thm vars fxy (IVar v) = haftmann@30648: str (Code_Printer.lookup_var vars v) haftmann@29175: | pr_term is_closure thm vars fxy (t as t1 `$ t2) = haftmann@28054: (case Code_Thingol.unfold_const_app t haftmann@29175: of SOME c_ts => pr_app is_closure thm vars fxy c_ts haftmann@28054: | NONE => haftmann@29175: brackify fxy [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2]) haftmann@29175: | pr_term is_closure thm vars fxy (t as _ `|-> _) = haftmann@28054: let haftmann@28054: val (binds, t') = Code_Thingol.unfold_abs t; haftmann@29175: fun pr ((v, pat), ty) = pr_bind is_closure thm BR ((SOME v, pat), ty); haftmann@28054: val (ps, vars') = fold_map pr binds vars; haftmann@29175: in brackets (str "fun" :: ps @ str "->" @@ pr_term is_closure thm vars' NOBR t') end haftmann@29175: | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) = (case Code_Thingol.unfold_const_app t0 haftmann@28054: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@29175: then pr_case is_closure thm vars fxy cases haftmann@29175: else pr_app is_closure thm vars fxy c_ts haftmann@29175: | NONE => pr_case is_closure thm vars fxy cases) haftmann@29175: and pr_app' is_closure thm vars (app as ((c, (iss, tys)), ts)) = haftmann@28054: if is_cons c then haftmann@28054: if length tys = length ts haftmann@28054: then case ts haftmann@28054: of [] => [(str o deresolve) c] haftmann@29175: | [t] => [(str o deresolve) c, pr_term is_closure thm vars BR t] haftmann@28054: | _ => [(str o deresolve) c, Pretty.enum "," "(" ")" haftmann@29175: (map (pr_term is_closure thm vars NOBR) ts)] haftmann@29175: else [pr_term is_closure thm vars BR (Code_Thingol.eta_expand (length tys) app)] haftmann@29175: else if is_closure c haftmann@29175: then (str o deresolve) c @@ str "()" haftmann@28054: else (str o deresolve) c haftmann@29175: :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts) haftmann@29175: and pr_app is_closure = gen_pr_app (pr_app' is_closure) (pr_term is_closure) haftmann@29175: syntax_const naming haftmann@28054: and pr_bind' ((NONE, NONE), _) = str "_" haftmann@28054: | pr_bind' ((SOME v, NONE), _) = str v haftmann@28054: | pr_bind' ((NONE, SOME p), _) = p haftmann@28054: | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v] haftmann@29175: and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure) haftmann@29175: and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) = haftmann@28054: let haftmann@29952: val (binds, body) = Code_Thingol.unfold_let (ICase cases); haftmann@28054: fun pr ((pat, ty), t) vars = haftmann@28054: vars haftmann@29175: |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) haftmann@28054: |>> (fn p => concat haftmann@29175: [str "let", p, str "=", pr_term is_closure thm vars NOBR t, str "in"]) haftmann@28054: val (ps, vars') = fold_map pr binds vars; haftmann@29952: in Pretty.chunks (ps @| pr_term is_closure thm vars' NOBR body) end haftmann@29952: | pr_case is_closure thm vars fxy (((t, ty), clause :: clauses), _) = haftmann@28054: let haftmann@29952: fun pr delim (pat, body) = haftmann@28054: let haftmann@29175: val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars; haftmann@29952: in concat [str delim, p, str "->", pr_term is_closure thm vars' NOBR body] end; haftmann@28054: in haftmann@28054: (Pretty.enclose "(" ")" o single o brackify fxy) ( haftmann@28054: str "match" haftmann@29952: :: pr_term is_closure thm vars NOBR t haftmann@29952: :: pr "with" clause haftmann@29952: :: map (pr "|") clauses haftmann@28054: ) haftmann@28054: end haftmann@29175: | pr_case is_closure thm vars fxy ((_, []), _) = str "failwith \"empty case\""; haftmann@28054: fun fish_params vars eqs = haftmann@28054: let haftmann@28054: fun fish_param _ (w as SOME _) = w haftmann@28054: | fish_param (IVar v) NONE = SOME v haftmann@28054: | fish_param _ NONE = NONE; haftmann@28054: fun fillup_param _ (_, SOME v) = v haftmann@28054: | fillup_param x (i, NONE) = x ^ string_of_int i; haftmann@28054: val fished1 = fold (map2 fish_param) eqs (replicate (length (hd eqs)) NONE); haftmann@28054: val x = Name.variant (map_filter I fished1) "x"; haftmann@28054: val fished2 = map_index (fillup_param x) fished1; haftmann@28054: val (fished3, _) = Name.variants fished2 Name.context; haftmann@30648: val vars' = Code_Printer.intro_vars fished3 vars; haftmann@30648: in map (Code_Printer.lookup_var vars') fished3 end; haftmann@29189: fun pr_stmt (MLExc (name, n)) = haftmann@29189: let haftmann@29189: val exc_str = wenzelm@30364: (ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name; haftmann@29189: in haftmann@29189: concat ( haftmann@29189: str "let" haftmann@29189: :: (str o deresolve) name haftmann@29189: :: map str (replicate n "_") haftmann@29189: @ str "=" haftmann@29189: :: str "failwith" haftmann@29189: @@ str exc_str haftmann@29189: ) haftmann@29189: end haftmann@29189: | pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) = haftmann@29189: let haftmann@29189: val consts = map_filter haftmann@29189: (fn c => if (is_some o syntax_const) c wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c) haftmann@29189: (Code_Thingol.fold_constnames (insert (op =)) t []); haftmann@29189: val vars = reserved_names haftmann@30648: |> Code_Printer.intro_vars consts; haftmann@29189: in haftmann@29189: concat [ haftmann@29189: str "let", haftmann@29189: (str o deresolve) name, haftmann@29189: str ":", haftmann@29189: pr_typ NOBR ty, haftmann@29189: str "=", haftmann@29189: pr_term (K false) thm vars NOBR t haftmann@29189: ] haftmann@29189: end haftmann@29189: | pr_stmt (MLFuns (funn :: funns, pseudo_funs)) = haftmann@28054: let haftmann@28350: fun pr_eq ((ts, t), (thm, _)) = haftmann@28054: let haftmann@28054: val consts = map_filter haftmann@28054: (fn c => if (is_some o syntax_const) c wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c) haftmann@28054: ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@28054: val vars = reserved_names haftmann@30648: |> Code_Printer.intro_vars consts haftmann@30648: |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) haftmann@28054: (insert (op =)) ts []); haftmann@28054: in concat [ haftmann@29189: (Pretty.block o Pretty.commas) haftmann@29189: (map (pr_term (member (op =) pseudo_funs) thm vars NOBR) ts), haftmann@28054: str "->", haftmann@29189: pr_term (member (op =) pseudo_funs) thm vars NOBR t haftmann@28054: ] end; haftmann@29189: fun pr_eqs is_pseudo [((ts, t), (thm, _))] = haftmann@28054: let haftmann@28054: val consts = map_filter haftmann@28054: (fn c => if (is_some o syntax_const) c wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c) haftmann@28054: ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); haftmann@28054: val vars = reserved_names haftmann@30648: |> Code_Printer.intro_vars consts haftmann@30648: |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) haftmann@28054: (insert (op =)) ts []); haftmann@28054: in haftmann@28054: concat ( haftmann@29189: (if is_pseudo then [str "()"] haftmann@29189: else map (pr_term (member (op =) pseudo_funs) thm vars BR) ts) haftmann@28054: @ str "=" haftmann@29189: @@ pr_term (member (op =) pseudo_funs) thm vars NOBR t haftmann@28054: ) haftmann@28054: end haftmann@29189: | pr_eqs _ (eqs as (eq as (([_], _), _)) :: eqs') = haftmann@28054: Pretty.block ( haftmann@28054: str "=" haftmann@28054: :: Pretty.brk 1 haftmann@28054: :: str "function" haftmann@28054: :: Pretty.brk 1 haftmann@28054: :: pr_eq eq haftmann@28054: :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] haftmann@28054: o single o pr_eq) eqs' haftmann@28054: ) haftmann@29189: | pr_eqs _ (eqs as eq :: eqs') = haftmann@28054: let haftmann@28054: val consts = map_filter haftmann@28054: (fn c => if (is_some o syntax_const) c wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c) haftmann@28054: ((fold o Code_Thingol.fold_constnames) haftmann@28054: (insert (op =)) (map (snd o fst) eqs) []); haftmann@28054: val vars = reserved_names haftmann@30648: |> Code_Printer.intro_vars consts; haftmann@28054: val dummy_parms = (map str o fish_params vars o map (fst o fst)) eqs; haftmann@28054: in haftmann@28054: Pretty.block ( haftmann@28054: Pretty.breaks dummy_parms haftmann@28054: @ Pretty.brk 1 haftmann@28054: :: str "=" haftmann@28054: :: Pretty.brk 1 haftmann@28054: :: str "match" haftmann@28054: :: Pretty.brk 1 haftmann@28054: :: (Pretty.block o Pretty.commas) dummy_parms haftmann@28054: :: Pretty.brk 1 haftmann@28054: :: str "with" haftmann@28054: :: Pretty.brk 1 haftmann@28054: :: pr_eq eq haftmann@28054: :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] haftmann@28054: o single o pr_eq) eqs' haftmann@28054: ) haftmann@28054: end; haftmann@29189: fun pr_funn definer (name, ((vs, ty), eqs)) = haftmann@28054: concat ( haftmann@28054: str definer haftmann@28054: :: (str o deresolve) name haftmann@28054: :: pr_tyvar_dicts (filter_out (null o snd) vs) haftmann@29189: @| pr_eqs (member (op =) pseudo_funs name) eqs haftmann@28054: ); haftmann@29189: fun pr_pseudo_fun name = concat [ haftmann@29189: str "let", haftmann@29189: (str o deresolve) name, haftmann@29189: str "=", haftmann@29189: (str o deresolve) name, haftmann@29189: str "();;" haftmann@29189: ]; haftmann@29189: val (ps, p) = split_last (pr_funn "fun" funn :: map (pr_funn "and") funns); haftmann@28054: val (ps, p) = split_last haftmann@29189: (pr_funn "let rec" funn :: map (pr_funn "and") funns); haftmann@29189: val pseudo_ps = map pr_pseudo_fun pseudo_funs; haftmann@29189: in Pretty.chunks (ps @ Pretty.block ([p, str ";;"]) :: pseudo_ps) end haftmann@28054: | pr_stmt (MLDatas (datas as (data :: datas'))) = haftmann@28054: let haftmann@28054: fun pr_co (co, []) = haftmann@28054: str (deresolve co) haftmann@28054: | pr_co (co, tys) = haftmann@28054: concat [ haftmann@28054: str (deresolve co), haftmann@28054: str "of", haftmann@28054: Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys) haftmann@28054: ]; haftmann@28054: fun pr_data definer (tyco, (vs, [])) = haftmann@28054: concat ( haftmann@28054: str definer haftmann@28054: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@28054: :: str "=" haftmann@28054: @@ str "EMPTY_" haftmann@28054: ) haftmann@28054: | pr_data definer (tyco, (vs, cos)) = haftmann@28054: concat ( haftmann@28054: str definer haftmann@28054: :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) haftmann@28054: :: str "=" haftmann@28054: :: separate (str "|") (map pr_co cos) haftmann@28054: ); haftmann@28054: val (ps, p) = split_last haftmann@28054: (pr_data "type" data :: map (pr_data "and") datas'); haftmann@29189: in Pretty.chunks (ps @| Pretty.block ([p, str ";;"])) end haftmann@28054: | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) = haftmann@28054: let haftmann@30648: val w = "_" ^ Code_Printer.first_upper v; haftmann@28054: fun pr_superclass_field (class, classrel) = haftmann@28054: (concat o map str) [ haftmann@28054: deresolve classrel, ":", "'" ^ v, deresolve class haftmann@28054: ]; haftmann@28054: fun pr_classparam_field (classparam, ty) = haftmann@28054: concat [ haftmann@28054: (str o deresolve) classparam, str ":", pr_typ NOBR ty haftmann@28054: ]; haftmann@28054: fun pr_classparam_proj (classparam, _) = haftmann@28054: concat [ haftmann@28054: str "let", haftmann@28054: (str o deresolve) classparam, haftmann@28054: str w, haftmann@28054: str "=", haftmann@28054: str (w ^ "." ^ deresolve classparam ^ ";;") haftmann@28054: ]; haftmann@28054: in Pretty.chunks ( haftmann@28054: concat [ haftmann@28054: str ("type '" ^ v), haftmann@28054: (str o deresolve) class, haftmann@28054: str "=", haftmann@28054: enum_default "();;" ";" "{" "};;" ( haftmann@28054: map pr_superclass_field superclasses haftmann@28054: @ map pr_classparam_field classparams haftmann@28054: ) haftmann@28054: ] haftmann@28054: :: map pr_classparam_proj classparams haftmann@28054: ) end haftmann@28054: | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) = haftmann@28054: let haftmann@28054: fun pr_superclass (_, (classrel, dss)) = haftmann@28054: concat [ haftmann@28054: (str o deresolve) classrel, haftmann@28054: str "=", haftmann@28054: pr_dicts NOBR [DictConst dss] haftmann@28054: ]; haftmann@28350: fun pr_classparam_inst ((classparam, c_inst), (thm, _)) = haftmann@28054: concat [ haftmann@28054: (str o deresolve) classparam, haftmann@28054: str "=", haftmann@29175: pr_app (K false) thm reserved_names NOBR (c_inst, []) haftmann@28054: ]; haftmann@28054: in haftmann@28054: concat ( haftmann@28054: str "let" haftmann@28054: :: (str o deresolve) inst haftmann@28054: :: pr_tyvar_dicts arity haftmann@28054: @ str "=" haftmann@28054: @@ (Pretty.enclose "(" ");;" o Pretty.breaks) [ haftmann@28054: enum_default "()" ";" "{" "}" (map pr_superclass superarities haftmann@28054: @ map pr_classparam_inst classparam_insts), haftmann@28054: str ":", haftmann@28054: pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) haftmann@28054: ] haftmann@28054: ) haftmann@28054: end; haftmann@28054: in pr_stmt end; haftmann@28054: haftmann@28054: fun pr_ocaml_module name content = haftmann@28054: Pretty.chunks ( haftmann@28054: str ("module " ^ name ^ " = ") haftmann@28054: :: str "struct" haftmann@28054: :: str "" haftmann@28054: :: content haftmann@28054: @ str "" haftmann@28054: @@ str ("end;; (*struct " ^ name ^ "*)") haftmann@28054: ); haftmann@28054: haftmann@28064: val literals_ocaml = let haftmann@28064: fun chr i = haftmann@28064: let haftmann@28064: val xs = string_of_int i; haftmann@28064: val ys = replicate_string (3 - length (explode xs)) "0"; haftmann@28064: in "\\" ^ ys ^ xs end; haftmann@28064: fun char_ocaml c = haftmann@28064: let haftmann@28064: val i = ord c; haftmann@28064: val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126 haftmann@28064: then chr i else c haftmann@28064: in s end; haftmann@28064: in Literals { haftmann@28064: literal_char = enclose "'" "'" o char_ocaml, haftmann@28064: literal_string = quote o translate_string char_ocaml, haftmann@28064: literal_numeral = fn unbounded => fn k => if k >= 0 then haftmann@28064: if unbounded then haftmann@28064: "(Big_int.big_int_of_int " ^ string_of_int k ^ ")" haftmann@28064: else string_of_int k haftmann@28064: else haftmann@28064: if unbounded then haftmann@28064: "(Big_int.big_int_of_int " ^ (enclose "(" ")" o prefix "-" haftmann@28064: o string_of_int o op ~) k ^ ")" haftmann@28064: else (enclose "(" ")" o prefix "-" o string_of_int o op ~) k, haftmann@28064: literal_list = Pretty.enum ";" "[" "]", haftmann@28064: infix_cons = (6, "::") haftmann@28064: } end; haftmann@28064: haftmann@28064: haftmann@28054: haftmann@28054: (** SML/OCaml generic part **) haftmann@28054: haftmann@28054: local haftmann@28054: haftmann@28054: datatype ml_node = haftmann@28054: Dummy of string haftmann@28054: | Stmt of string * ml_stmt haftmann@28054: | Module of string * ((Name.context * Name.context) * ml_node Graph.T); haftmann@28054: haftmann@28054: in haftmann@28054: haftmann@28054: fun ml_node_of_program labelled_name module_name reserved_names raw_module_alias program = haftmann@28054: let haftmann@28054: val module_alias = if is_some module_name then K module_name else raw_module_alias; haftmann@28054: val reserved_names = Name.make_context reserved_names; haftmann@28054: val empty_module = ((reserved_names, reserved_names), Graph.empty); haftmann@28054: fun map_node [] f = f haftmann@28054: | map_node (m::ms) f = haftmann@28054: Graph.default_node (m, Module (m, empty_module)) haftmann@28054: #> Graph.map_node m (fn (Module (module_name, (nsp, nodes))) => haftmann@28054: Module (module_name, (nsp, map_node ms f nodes))); haftmann@28054: fun map_nsp_yield [] f (nsp, nodes) = haftmann@28054: let haftmann@28054: val (x, nsp') = f nsp haftmann@28054: in (x, (nsp', nodes)) end haftmann@28054: | map_nsp_yield (m::ms) f (nsp, nodes) = haftmann@28054: let haftmann@28054: val (x, nodes') = haftmann@28054: nodes haftmann@28054: |> Graph.default_node (m, Module (m, empty_module)) haftmann@28054: |> Graph.map_node_yield m (fn Module (d_module_name, nsp_nodes) => haftmann@28054: let haftmann@28054: val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes haftmann@28054: in (x, Module (d_module_name, nsp_nodes')) end) haftmann@28054: in (x, (nsp, nodes')) end; haftmann@28054: fun map_nsp_fun_yield f (nsp_fun, nsp_typ) = haftmann@28054: let haftmann@28054: val (x, nsp_fun') = f nsp_fun haftmann@28054: in (x, (nsp_fun', nsp_typ)) end; haftmann@28054: fun map_nsp_typ_yield f (nsp_fun, nsp_typ) = haftmann@28054: let haftmann@28054: val (x, nsp_typ') = f nsp_typ haftmann@28054: in (x, (nsp_fun, nsp_typ')) end; haftmann@30648: val mk_name_module = Code_Printer.mk_name_module reserved_names NONE module_alias program; haftmann@28054: fun mk_name_stmt upper name nsp = haftmann@28054: let haftmann@30648: val (_, base) = Code_Printer.dest_name name; haftmann@30648: val base' = if upper then Code_Printer.first_upper base else base; haftmann@28054: val ([base''], nsp') = Name.variants [base'] nsp; haftmann@28054: in (base'', nsp') end; haftmann@29189: fun rearrange_fun name (tysm as (vs, ty), raw_eqs) = haftmann@29189: let haftmann@29189: val eqs = filter (snd o snd) raw_eqs; haftmann@29189: val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs haftmann@29189: of [(([], t), thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty haftmann@29189: then ([(([IVar "x"], t `$ IVar "x"), thm)], false) haftmann@29189: else (eqs, not (Code_Thingol.fold_constnames haftmann@29189: (fn name' => fn b => b orelse name = name') t false)) haftmann@29189: | _ => (eqs, false) haftmann@29189: else (eqs, false) haftmann@29189: in ((name, (tysm, eqs')), is_value) end; haftmann@29189: fun check_kind [((name, (tysm, [(([], t), thm)])), true)] = MLVal (name, ((tysm, t), thm)) haftmann@29189: | check_kind [((name, ((vs, ty), [])), _)] = haftmann@29189: MLExc (name, (length o filter_out (null o snd)) vs + (length o fst o Code_Thingol.unfold_fun) ty) haftmann@29189: | check_kind funns = haftmann@29189: MLFuns (map fst funns, map_filter haftmann@29189: (fn ((name, ((vs, _), [(([], _), _)])), _) => haftmann@29189: if null (filter_out (null o snd) vs) then SOME name else NONE haftmann@29189: | _ => NONE) funns); haftmann@29189: fun add_funs stmts = fold_map haftmann@28663: (fn (name, Code_Thingol.Fun (_, stmt)) => haftmann@29189: map_nsp_fun_yield (mk_name_stmt false name) haftmann@29189: #>> rpair (rearrange_fun name stmt) haftmann@28054: | (name, _) => haftmann@28054: error ("Function block containing illegal statement: " ^ labelled_name name) haftmann@28054: ) stmts haftmann@29189: #>> (split_list #> apsnd check_kind); haftmann@28054: fun add_datatypes stmts = haftmann@28054: fold_map haftmann@28663: (fn (name, Code_Thingol.Datatype (_, stmt)) => haftmann@28054: map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt)) haftmann@28054: | (name, Code_Thingol.Datatypecons _) => haftmann@28054: map_nsp_fun_yield (mk_name_stmt true name) #>> rpair NONE haftmann@28054: | (name, _) => haftmann@28054: error ("Datatype block containing illegal statement: " ^ labelled_name name) haftmann@28054: ) stmts haftmann@28054: #>> (split_list #> apsnd (map_filter I haftmann@28054: #> (fn [] => error ("Datatype block without data statement: " haftmann@28054: ^ (commas o map (labelled_name o fst)) stmts) haftmann@28054: | stmts => MLDatas stmts))); haftmann@28054: fun add_class stmts = haftmann@28054: fold_map haftmann@28663: (fn (name, Code_Thingol.Class (_, stmt)) => haftmann@28663: map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt)) haftmann@28054: | (name, Code_Thingol.Classrel _) => haftmann@28054: map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE haftmann@28054: | (name, Code_Thingol.Classparam _) => haftmann@28054: map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE haftmann@28054: | (name, _) => haftmann@28054: error ("Class block containing illegal statement: " ^ labelled_name name) haftmann@28054: ) stmts haftmann@28054: #>> (split_list #> apsnd (map_filter I haftmann@28054: #> (fn [] => error ("Class block without class statement: " haftmann@28054: ^ (commas o map (labelled_name o fst)) stmts) haftmann@28054: | [stmt] => MLClass stmt))); haftmann@28054: fun add_inst [(name, Code_Thingol.Classinst stmt)] = haftmann@28054: map_nsp_fun_yield (mk_name_stmt false name) haftmann@28054: #>> (fn base => ([base], MLClassinst (name, stmt))); haftmann@28054: fun add_stmts ((stmts as (_, Code_Thingol.Fun _)::_)) = haftmann@28054: add_funs stmts haftmann@28054: | add_stmts ((stmts as (_, Code_Thingol.Datatypecons _)::_)) = haftmann@28054: add_datatypes stmts haftmann@28054: | add_stmts ((stmts as (_, Code_Thingol.Datatype _)::_)) = haftmann@28054: add_datatypes stmts haftmann@28054: | add_stmts ((stmts as (_, Code_Thingol.Class _)::_)) = haftmann@28054: add_class stmts haftmann@28054: | add_stmts ((stmts as (_, Code_Thingol.Classrel _)::_)) = haftmann@28054: add_class stmts haftmann@28054: | add_stmts ((stmts as (_, Code_Thingol.Classparam _)::_)) = haftmann@28054: add_class stmts haftmann@28054: | add_stmts ((stmts as [(_, Code_Thingol.Classinst _)])) = haftmann@28054: add_inst stmts haftmann@28054: | add_stmts stmts = error ("Illegal mutual dependencies: " ^ haftmann@28054: (commas o map (labelled_name o fst)) stmts); haftmann@28054: fun add_stmts' stmts nsp_nodes = haftmann@28054: let haftmann@28054: val names as (name :: names') = map fst stmts; haftmann@28054: val deps = haftmann@28054: [] haftmann@28054: |> fold (fold (insert (op =)) o Graph.imm_succs program) names haftmann@28054: |> subtract (op =) names; haftmann@30648: val (module_names, _) = (split_list o map Code_Printer.dest_name) names; haftmann@28054: val module_name = (the_single o distinct (op =) o map mk_name_module) module_names haftmann@28054: handle Empty => haftmann@28054: error ("Different namespace prefixes for mutual dependencies:\n" haftmann@28054: ^ commas (map labelled_name names) haftmann@28054: ^ "\n" haftmann@28054: ^ commas module_names); wenzelm@30364: val module_name_path = Long_Name.explode module_name; haftmann@28054: fun add_dep name name' = haftmann@28054: let haftmann@30648: val module_name' = (mk_name_module o fst o Code_Printer.dest_name) name'; haftmann@28054: in if module_name = module_name' then haftmann@28054: map_node module_name_path (Graph.add_edge (name, name')) haftmann@28054: else let haftmann@28705: val (common, (diff1 :: _, diff2 :: _)) = chop_prefix (op =) wenzelm@30364: (module_name_path, Long_Name.explode module_name'); haftmann@28054: in haftmann@28054: map_node common haftmann@28054: (fn node => Graph.add_edge_acyclic (diff1, diff2) node haftmann@28054: handle Graph.CYCLES _ => error ("Dependency " haftmann@28054: ^ quote name ^ " -> " ^ quote name' haftmann@28054: ^ " would result in module dependency cycle")) haftmann@28054: end end; haftmann@28054: in haftmann@28054: nsp_nodes haftmann@28054: |> map_nsp_yield module_name_path (add_stmts stmts) haftmann@28054: |-> (fn (base' :: bases', stmt') => haftmann@28054: apsnd (map_node module_name_path (Graph.new_node (name, (Stmt (base', stmt'))) haftmann@28054: #> fold2 (fn name' => fn base' => haftmann@28054: Graph.new_node (name', (Dummy base'))) names' bases'))) haftmann@28054: |> apsnd (fold (fn name => fold (add_dep name) deps) names) haftmann@28054: |> apsnd (fold_product (curry (map_node module_name_path o Graph.add_edge)) names names) haftmann@28054: end; haftmann@28054: val (_, nodes) = empty_module haftmann@28054: |> fold add_stmts' (map (AList.make (Graph.get_node program)) haftmann@28054: (rev (Graph.strong_conn program))); haftmann@28054: fun deresolver prefix name = haftmann@28054: let haftmann@30648: val module_name = (fst o Code_Printer.dest_name) name; wenzelm@30364: val module_name' = (Long_Name.explode o mk_name_module) module_name; haftmann@28054: val (_, (_, remainder)) = chop_prefix (op =) (prefix, module_name'); haftmann@28054: val stmt_name = haftmann@28054: nodes haftmann@28054: |> fold (fn name => fn node => case Graph.get_node node name haftmann@28054: of Module (_, (_, node)) => node) module_name' haftmann@28054: |> (fn node => case Graph.get_node node name of Stmt (stmt_name, _) => stmt_name haftmann@28054: | Dummy stmt_name => stmt_name); haftmann@28054: in wenzelm@30364: Long_Name.implode (remainder @ [stmt_name]) haftmann@28054: end handle Graph.UNDEF _ => haftmann@28054: error ("Unknown statement name: " ^ labelled_name name); haftmann@28054: in (deresolver, nodes) end; haftmann@28054: haftmann@28054: fun serialize_ml target compile pr_module pr_stmt raw_module_name labelled_name reserved_names includes raw_module_alias haftmann@30962: _ syntax_tyco syntax_const naming program stmt_names destination = haftmann@28054: let haftmann@28054: val is_cons = Code_Thingol.is_cons program; haftmann@30962: val present_stmt_names = Code_Target.stmt_names_of_destination destination; haftmann@30962: val is_present = not (null present_stmt_names); haftmann@30962: val module_name = if is_present then SOME "Code" else raw_module_name; haftmann@28054: val (deresolver, nodes) = ml_node_of_program labelled_name module_name haftmann@28054: reserved_names raw_module_alias program; haftmann@30648: val reserved_names = Code_Printer.make_vars reserved_names; haftmann@28054: fun pr_node prefix (Dummy _) = haftmann@28054: NONE haftmann@30962: | pr_node prefix (Stmt (_, stmt)) = if is_present andalso haftmann@30962: (null o filter (member (op =) present_stmt_names) o stmt_names_of) stmt haftmann@30962: then NONE haftmann@30962: else SOME haftmann@28663: (pr_stmt naming labelled_name syntax_tyco syntax_const reserved_names haftmann@28054: (deresolver prefix) is_cons stmt) haftmann@28054: | pr_node prefix (Module (module_name, (_, nodes))) = haftmann@28054: separate (str "") haftmann@28054: ((map_filter (pr_node (prefix @ [module_name]) o Graph.get_node nodes) haftmann@28054: o rev o flat o Graph.strong_conn) nodes) haftmann@30962: |> (if is_present then Pretty.chunks else pr_module module_name) haftmann@28054: |> SOME; haftmann@30962: val stmt_names' = (map o try) haftmann@30962: (deresolver (if is_some module_name then the_list module_name else [])) stmt_names; haftmann@28054: val p = Pretty.chunks (separate (str "") (map snd includes @ (map_filter haftmann@28054: (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes)); haftmann@28054: in haftmann@28054: Code_Target.mk_serialization target haftmann@28054: (case compile of SOME compile => SOME (compile o Code_Target.code_of_pretty) | NONE => NONE) haftmann@28054: (fn NONE => Code_Target.code_writeln | SOME file => File.write file o Code_Target.code_of_pretty) haftmann@30962: (rpair stmt_names' o Code_Target.code_of_pretty) p destination haftmann@28054: end; haftmann@28054: haftmann@28054: end; (*local*) haftmann@28054: haftmann@28054: haftmann@28054: (** ML (system language) code for evaluation and instrumentalization **) haftmann@28054: haftmann@30947: fun eval_code_of some_target thy = Code_Target.serialize_custom thy (the_default target_Eval some_target, haftmann@28663: (fn _ => fn [] => serialize_ml target_SML (SOME (K ())) (K Pretty.chunks) pr_sml_stmt (SOME ""), haftmann@28064: literals_sml)); haftmann@28054: haftmann@28054: haftmann@28054: (* evaluation *) haftmann@28054: haftmann@30947: fun gen_eval eval some_target reff thy t args = haftmann@28054: let wenzelm@28275: val ctxt = ProofContext.init thy; haftmann@30941: val _ = if null (Term.add_frees t []) then () else error ("Term " haftmann@30941: ^ quote (Syntax.string_of_term_global thy t) haftmann@28054: ^ " to be evaluated contains free variables"); haftmann@30947: fun evaluator naming program (((_, (_, ty)), _), t) deps = haftmann@28054: let haftmann@28054: val _ = if Code_Thingol.contains_dictvar t then haftmann@28724: error "Term to be evaluated contains free dictionaries" else (); haftmann@28663: val value_name = "Value.VALUE.value" haftmann@28054: val program' = program haftmann@28663: |> Graph.new_node (value_name, haftmann@28663: Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (Drule.dummy_thm, true))]))) haftmann@28663: |> fold (curry Graph.add_edge value_name) deps; haftmann@30947: val (value_code, [SOME value_name']) = eval_code_of some_target thy naming program' [value_name]; haftmann@28054: val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name' haftmann@28054: ^ space_implode " " (map (enclose "(" ")") args) ^ " end"; wenzelm@30672: in ML_Context.evaluate ctxt false reff sml_code end; haftmann@30947: in eval thy I evaluator t end; haftmann@30947: haftmann@30947: fun eval_term thy = gen_eval Code_Thingol.eval_term thy; haftmann@30947: fun eval thy = gen_eval Code_Thingol.eval thy; haftmann@28054: haftmann@28054: haftmann@28054: (* instrumentalization by antiquotation *) haftmann@28054: haftmann@28054: local haftmann@28054: haftmann@28054: structure CodeAntiqData = ProofDataFun haftmann@28054: ( haftmann@30962: type T = (string list * string list) * (bool * (string haftmann@30962: * (string * ((string * string) list * (string * string) list)) lazy)); haftmann@30962: fun init _ = (([], []), (true, ("", Lazy.value ("", ([], []))))); haftmann@28054: ); haftmann@28054: haftmann@28054: val is_first_occ = fst o snd o CodeAntiqData.get; haftmann@28054: haftmann@30962: fun delayed_code thy tycos consts () = haftmann@28054: let haftmann@28663: val (consts', (naming, program)) = Code_Thingol.consts_program thy consts; haftmann@30962: val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos; haftmann@30962: val (ml_code, target_names) = eval_code_of NONE thy naming program (consts' @ tycos'); haftmann@30962: val (consts'', tycos'') = chop (length consts') target_names; haftmann@30962: val consts_map = map2 (fn const => fn NONE => haftmann@30962: error ("Constant " ^ (quote o Code_Unit.string_of_const thy) const haftmann@30962: ^ "\nhas a user-defined serialization") haftmann@30962: | SOME const'' => (const, const'')) consts consts'' haftmann@30962: val tycos_map = map2 (fn tyco => fn NONE => haftmann@30962: error ("Type " ^ (quote o Sign.extern_type thy) tyco haftmann@30962: ^ "\nhas a user-defined serialization") haftmann@30962: | SOME tyco'' => (tyco, tyco'')) tycos tycos''; haftmann@30962: in (ml_code, (tycos_map, consts_map)) end; haftmann@28054: haftmann@30962: fun register_code new_tycos new_consts ctxt = haftmann@28054: let haftmann@30962: val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt; haftmann@30962: val tycos' = fold (insert (op =)) new_tycos tycos; haftmann@30962: val consts' = fold (insert (op =)) new_consts consts; haftmann@28054: val (struct_name', ctxt') = if struct_name = "" haftmann@28054: then ML_Antiquote.variant "Code" ctxt haftmann@28054: else (struct_name, ctxt); haftmann@30962: val acc_code = Lazy.lazy (delayed_code (ProofContext.theory_of ctxt) tycos' consts'); haftmann@30962: in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end; haftmann@30962: haftmann@30962: fun register_const const = register_code [] [const]; haftmann@28054: haftmann@30962: fun register_datatype tyco constrs = register_code [tyco] constrs; haftmann@30962: haftmann@30962: fun print_const const all_struct_name tycos_map consts_map = haftmann@30962: (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const; haftmann@30962: haftmann@30962: fun print_datatype tyco constrs all_struct_name tycos_map consts_map = haftmann@28054: let haftmann@30962: val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode; haftmann@30962: fun check_base name name'' = haftmann@30962: if upperize (Long_Name.base_name name) = upperize name'' haftmann@30962: then () else error ("Name as printed " ^ quote name'' haftmann@30962: ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry."); haftmann@30962: val tyco'' = (the o AList.lookup (op =) tycos_map) tyco; haftmann@30962: val constrs'' = map (the o AList.lookup (op =) consts_map) constrs; haftmann@30962: val _ = check_base tyco tyco''; haftmann@30962: val _ = map2 check_base constrs constrs''; haftmann@30962: in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end; haftmann@30962: haftmann@30962: fun print_code struct_name is_first print_it ctxt = haftmann@30962: let haftmann@30962: val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; haftmann@30962: val (raw_ml_code, (tycos_map, consts_map)) = Lazy.force acc_code; haftmann@28054: val ml_code = if is_first then "\nstructure " ^ struct_code_name haftmann@28054: ^ " =\nstruct\n\n" ^ raw_ml_code ^ "\nend;\n\n" haftmann@28054: else ""; haftmann@30962: val all_struct_name = Long_Name.append struct_name struct_code_name; haftmann@30962: in (ml_code, print_it all_struct_name tycos_map consts_map) end; haftmann@28054: haftmann@28054: in haftmann@28054: haftmann@28054: fun ml_code_antiq raw_const {struct_name, background} = haftmann@28054: let haftmann@28054: val const = Code_Unit.check_const (ProofContext.theory_of background) raw_const; haftmann@28054: val is_first = is_first_occ background; haftmann@28054: val background' = register_const const background; haftmann@30962: in (print_code struct_name is_first (print_const const), background') end; haftmann@30962: haftmann@30962: fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} = haftmann@30962: let haftmann@30962: val thy = ProofContext.theory_of background; haftmann@30962: val tyco = Sign.intern_type thy raw_tyco; haftmann@30962: val constrs = map (Code_Unit.check_const thy) raw_constrs; haftmann@30962: val constrs' = (map fst o snd o Code.get_datatype thy) tyco; haftmann@30962: val _ = if gen_eq_set (op =) (constrs, constrs') then () haftmann@30962: else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors") haftmann@30962: val is_first = is_first_occ background; haftmann@30962: val background' = register_datatype tyco constrs background; haftmann@30962: in (print_code struct_name is_first (print_datatype tyco constrs), background') end; haftmann@28054: haftmann@28054: end; (*local*) haftmann@28054: haftmann@28054: haftmann@28054: (** Isar setup **) haftmann@28054: haftmann@28054: val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq); haftmann@30962: val _ = ML_Context.add_antiq "code_datatype" (fn _ => haftmann@30962: (Args.tyname --| Scan.lift (Args.$$$ "=") haftmann@30962: -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term))) haftmann@30962: >> ml_code_datatype_antiq); haftmann@28054: haftmann@28054: fun isar_seri_sml module_name = haftmann@28054: Code_Target.parse_args (Scan.succeed ()) wenzelm@28275: #> (fn () => serialize_ml target_SML wenzelm@30675: (SOME (use_text ML_Context.local_context (1, "generated code") false)) haftmann@28054: pr_sml_module pr_sml_stmt module_name); haftmann@28054: haftmann@28054: fun isar_seri_ocaml module_name = haftmann@28054: Code_Target.parse_args (Scan.succeed ()) haftmann@28054: #> (fn () => serialize_ml target_OCaml NONE haftmann@28054: pr_ocaml_module pr_ocaml_stmt module_name); haftmann@28054: haftmann@28054: val setup = haftmann@28064: Code_Target.add_target (target_SML, (isar_seri_sml, literals_sml)) haftmann@28064: #> Code_Target.add_target (target_OCaml, (isar_seri_ocaml, literals_ocaml)) haftmann@30947: #> Code_Target.extend_target (target_Eval, (target_SML, K I)) haftmann@28054: #> Code_Target.add_syntax_tyco target_SML "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => haftmann@28054: brackify_infix (1, R) fxy [ haftmann@28054: pr_typ (INFX (1, X)) ty1, haftmann@28054: str "->", haftmann@28054: pr_typ (INFX (1, R)) ty2 haftmann@28054: ])) haftmann@28054: #> Code_Target.add_syntax_tyco target_OCaml "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => haftmann@28054: brackify_infix (1, R) fxy [ haftmann@28054: pr_typ (INFX (1, X)) ty1, haftmann@28054: str "->", haftmann@28054: pr_typ (INFX (1, R)) ty2 haftmann@28054: ])) haftmann@28054: #> fold (Code_Target.add_reserved target_SML) ML_Syntax.reserved_names haftmann@28054: #> fold (Code_Target.add_reserved target_SML) haftmann@28054: ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)] haftmann@28054: #> fold (Code_Target.add_reserved target_OCaml) [ haftmann@28054: "and", "as", "assert", "begin", "class", haftmann@28054: "constraint", "do", "done", "downto", "else", "end", "exception", haftmann@28054: "external", "false", "for", "fun", "function", "functor", "if", haftmann@28054: "in", "include", "inherit", "initializer", "lazy", "let", "match", "method", haftmann@28054: "module", "mutable", "new", "object", "of", "open", "or", "private", "rec", haftmann@28054: "sig", "struct", "then", "to", "true", "try", "type", "val", haftmann@28054: "virtual", "when", "while", "with" haftmann@28054: ] haftmann@28054: #> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod"]; haftmann@28054: haftmann@28054: end; (*struct*)