haftmann@37745: (* Title: Tools/Code/code_scala.ML haftmann@37745: Author: Florian Haftmann, TU Muenchen haftmann@34294: haftmann@34294: Serializer for Scala. haftmann@34294: *) haftmann@34294: haftmann@34294: signature CODE_SCALA = haftmann@34294: sig haftmann@37745: val target: string haftmann@34294: val setup: theory -> theory haftmann@34294: end; haftmann@34294: haftmann@34294: structure Code_Scala : CODE_SCALA = haftmann@34294: struct haftmann@34294: haftmann@34294: val target = "Scala"; haftmann@34294: haftmann@34294: open Basic_Code_Thingol; haftmann@34294: open Code_Printer; haftmann@34294: haftmann@34294: infixr 5 @@; haftmann@34294: infixr 5 @|; haftmann@34294: haftmann@34294: haftmann@34294: (** Scala serializer **) haftmann@34294: haftmann@37464: fun print_scala_stmt labelled_name syntax_tyco syntax_const reserved haftmann@38780: args_num is_singleton_constr (deresolve, deresolve_full) = haftmann@34294: let haftmann@34294: val deresolve_base = Long_Name.base_name o deresolve; haftmann@37639: fun lookup_tyvar tyvars = lookup_var tyvars o first_upper; haftmann@37639: fun intro_tyvars vs = intro_vars (map (first_upper o fst) vs); haftmann@37639: fun print_tyco_expr tyvars fxy (tyco, tys) = applify "[" "]" haftmann@37639: (print_typ tyvars NOBR) fxy ((str o deresolve) tyco) tys haftmann@37639: and print_typ tyvars fxy (tyco `%% tys) = (case syntax_tyco tyco haftmann@37639: of NONE => print_tyco_expr tyvars fxy (tyco, tys) haftmann@34294: | SOME (i, print) => print (print_typ tyvars) fxy tys) haftmann@37243: | print_typ tyvars fxy (ITyVar v) = (str o lookup_tyvar tyvars) v; haftmann@37639: fun print_dicttyp tyvars (class, ty) = print_tyco_expr tyvars NOBR (class, [ty]); haftmann@37639: fun print_tupled_typ tyvars ([], ty) = haftmann@37639: print_typ tyvars NOBR ty haftmann@37639: | print_tupled_typ tyvars ([ty1], ty2) = haftmann@37639: concat [print_typ tyvars BR ty1, str "=>", print_typ tyvars NOBR ty2] haftmann@37639: | print_tupled_typ tyvars (tys, ty) = haftmann@37639: concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys), haftmann@37639: str "=>", print_typ tyvars NOBR ty]; haftmann@37639: fun constraint p1 p2 = Pretty.block [p1, str ":", Pretty.brk 1, p2]; haftmann@34294: fun print_var vars NONE = str "_" haftmann@34294: | print_var vars (SOME v) = (str o lookup_var vars) v haftmann@35228: fun print_term tyvars is_pat some_thm vars fxy (IConst c) = haftmann@35228: print_app tyvars is_pat some_thm vars fxy (c, []) haftmann@35228: | print_term tyvars is_pat some_thm vars fxy (t as (t1 `$ t2)) = haftmann@34294: (case Code_Thingol.unfold_const_app t haftmann@35228: of SOME app => print_app tyvars is_pat some_thm vars fxy app haftmann@37639: | _ => applify "(" ")" (print_term tyvars is_pat some_thm vars NOBR) fxy haftmann@37639: (print_term tyvars is_pat some_thm vars BR t1) [t2]) haftmann@35228: | print_term tyvars is_pat some_thm vars fxy (IVar v) = haftmann@34294: print_var vars v haftmann@35228: | print_term tyvars is_pat some_thm vars fxy ((v, ty) `|=> t) = haftmann@34294: let haftmann@34294: val vars' = intro_vars (the_list v) vars; haftmann@34294: in haftmann@34294: concat [ haftmann@37639: enclose "(" ")" [constraint (print_var vars' v) (print_typ tyvars NOBR ty)], haftmann@34294: str "=>", haftmann@35228: print_term tyvars false some_thm vars' NOBR t haftmann@34294: ] haftmann@34294: end haftmann@35228: | print_term tyvars is_pat some_thm vars fxy (ICase (cases as (_, t0))) = haftmann@34294: (case Code_Thingol.unfold_const_app t0 haftmann@34294: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@35228: then print_case tyvars some_thm vars fxy cases haftmann@35228: else print_app tyvars is_pat some_thm vars fxy c_ts haftmann@35228: | NONE => print_case tyvars some_thm vars fxy cases) haftmann@37464: and print_app tyvars is_pat some_thm vars fxy haftmann@37464: (app as ((c, ((arg_typs, _), function_typs)), ts)) = haftmann@34294: let haftmann@34294: val k = length ts; haftmann@37450: val arg_typs' = if is_pat orelse haftmann@37639: (is_none (syntax_const c) andalso is_singleton_constr c) then [] else arg_typs; haftmann@37881: val (l, print') = case syntax_const c haftmann@38059: of NONE => (args_num c, fn fxy => fn ts => applify "(" ")" haftmann@37639: (print_term tyvars is_pat some_thm vars NOBR) fxy haftmann@37639: (applify "[" "]" (print_typ tyvars NOBR) haftmann@37639: NOBR ((str o deresolve) c) arg_typs') ts) haftmann@38059: | SOME (Plain_const_syntax (k, s)) => (k, fn fxy => fn ts => applify "(" ")" haftmann@37881: (print_term tyvars is_pat some_thm vars NOBR) fxy haftmann@37881: (applify "[" "]" (print_typ tyvars NOBR) haftmann@37881: NOBR (str s) arg_typs') ts) haftmann@37881: | SOME (Complex_const_syntax (k, print)) => haftmann@38059: (k, fn fxy => fn ts => print (print_term tyvars is_pat some_thm) some_thm vars fxy haftmann@37881: (ts ~~ take k function_typs)) haftmann@38059: in if k = l then print' fxy ts haftmann@34294: else if k < l then haftmann@35228: print_term tyvars is_pat some_thm vars fxy (Code_Thingol.eta_expand l app) haftmann@34294: else let haftmann@34294: val (ts1, ts23) = chop l ts; haftmann@34294: in haftmann@38059: Pretty.block (print' BR ts1 :: map (fn t => Pretty.block haftmann@35228: [str ".apply(", print_term tyvars is_pat some_thm vars NOBR t, str ")"]) ts23) haftmann@34294: end end haftmann@37464: and print_bind tyvars some_thm fxy p = haftmann@37464: gen_print_bind (print_term tyvars true) some_thm fxy p haftmann@35228: and print_case tyvars some_thm vars fxy (cases as ((_, [_]), _)) = haftmann@34294: let haftmann@34294: val (binds, body) = Code_Thingol.unfold_let (ICase cases); haftmann@38059: fun print_match ((IVar NONE, _), t) vars = haftmann@38059: ((true, print_term tyvars false some_thm vars NOBR t), vars) haftmann@38059: | print_match ((pat, ty), t) vars = haftmann@38059: vars haftmann@38059: |> print_bind tyvars some_thm BR pat haftmann@38059: |>> (fn p => (false, concat [str "val", constraint p (print_typ tyvars NOBR ty), haftmann@38059: str "=", print_term tyvars false some_thm vars NOBR t])) haftmann@38059: val (seps_ps, vars') = fold_map print_match binds vars; haftmann@38059: val all_seps_ps = seps_ps @ [(true, print_term tyvars false some_thm vars' NOBR body)]; haftmann@38059: fun insert_seps [(_, p)] = [p] haftmann@38059: | insert_seps ((_, p) :: (seps_ps as (sep, _) :: _)) = haftmann@38059: (if sep then Pretty.block [p, str ";"] else p) :: insert_seps seps_ps haftmann@38059: in brackify_block fxy (str "{") (insert_seps all_seps_ps) (str "}") haftmann@34294: end haftmann@35228: | print_case tyvars some_thm vars fxy (((t, ty), clauses as _ :: _), _) = haftmann@34294: let haftmann@34294: fun print_select (pat, body) = haftmann@34294: let haftmann@37464: val (p_pat, vars') = print_bind tyvars some_thm NOBR pat vars; haftmann@37464: val p_body = print_term tyvars false some_thm vars' NOBR body haftmann@37464: in concat [str "case", p_pat, str "=>", p_body] end; haftmann@34294: in brackify_block fxy haftmann@35228: (concat [print_term tyvars false some_thm vars NOBR t, str "match", str "{"]) haftmann@34294: (map print_select clauses) haftmann@34294: (str "}") haftmann@34294: end haftmann@35228: | print_case tyvars some_thm vars fxy ((_, []), _) = haftmann@34294: (brackify fxy o Pretty.breaks o map str) ["error(\"empty case\")"]; haftmann@37639: fun print_context tyvars vs name = applify "[" "]" haftmann@37639: (fn (v, sort) => (Pretty.block o map str) haftmann@37639: (lookup_tyvar tyvars v :: maps (fn sort => [": ", deresolve sort]) sort)) haftmann@38769: NOBR ((str o deresolve_base) name) vs; haftmann@37639: fun print_defhead tyvars vars name vs params tys ty = haftmann@37639: Pretty.block [str "def ", constraint (applify "(" ")" (fn (param, ty) => haftmann@37639: constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty)) haftmann@37639: NOBR (print_context tyvars vs name) (params ~~ tys)) (print_typ tyvars NOBR ty), haftmann@37639: str " ="]; haftmann@37639: fun print_def name (vs, ty) [] = haftmann@37639: let haftmann@37639: val (tys, ty') = Code_Thingol.unfold_fun ty; haftmann@37639: val params = Name.invents (snd reserved) "a" (length tys); haftmann@37639: val tyvars = intro_tyvars vs reserved; haftmann@37639: val vars = intro_vars params reserved; haftmann@37639: in haftmann@37639: concat [print_defhead tyvars vars name vs params tys ty', haftmann@37639: str ("error(\"" ^ name ^ "\")")] haftmann@37639: end haftmann@37639: | print_def name (vs, ty) eqs = haftmann@37639: let haftmann@37639: val tycos = fold (fn ((ts, t), _) => haftmann@37639: fold Code_Thingol.add_tyconames (t :: ts)) eqs []; haftmann@37639: val tyvars = reserved haftmann@37639: |> intro_base_names haftmann@37639: (is_none o syntax_tyco) deresolve tycos haftmann@37639: |> intro_tyvars vs; haftmann@37639: val simple = case eqs haftmann@37639: of [((ts, _), _)] => forall Code_Thingol.is_IVar ts haftmann@37639: | _ => false; haftmann@37639: val consts = fold Code_Thingol.add_constnames haftmann@37639: (map (snd o fst) eqs) []; haftmann@37639: val vars1 = reserved haftmann@37639: |> intro_base_names haftmann@37639: (is_none o syntax_const) deresolve consts haftmann@37639: val params = if simple haftmann@37639: then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs haftmann@37639: else aux_params vars1 (map (fst o fst) eqs); haftmann@37639: val vars2 = intro_vars params vars1; haftmann@37639: val (tys', ty') = Code_Thingol.unfold_fun_n (length params) ty; haftmann@37639: fun print_tuple [p] = p haftmann@37639: | print_tuple ps = enum "," "(" ")" ps; haftmann@37639: fun print_rhs vars' ((_, t), (some_thm, _)) = haftmann@37639: print_term tyvars false some_thm vars' NOBR t; haftmann@37639: fun print_clause (eq as ((ts, _), (some_thm, _))) = haftmann@34294: let haftmann@37639: val vars' = intro_vars ((fold o Code_Thingol.fold_varnames) haftmann@37639: (insert (op =)) ts []) vars1; haftmann@37639: in haftmann@37639: concat [str "case", haftmann@37639: print_tuple (map (print_term tyvars true some_thm vars' NOBR) ts), haftmann@37639: str "=>", print_rhs vars' eq] haftmann@37639: end; haftmann@37639: val head = print_defhead tyvars vars2 name vs params tys' ty'; haftmann@37639: in if simple then haftmann@37639: concat [head, print_rhs vars2 (hd eqs)] haftmann@37639: else haftmann@37639: Pretty.block_enclose haftmann@37639: (concat [head, print_tuple (map (str o lookup_var vars2) params), haftmann@37639: str "match", str "{"], str "}") haftmann@37639: (map print_clause eqs) haftmann@37639: end; haftmann@38769: val print_method = str o Library.enclose "`" "`" o space_implode "+" haftmann@38780: o Long_Name.explode o deresolve_full; haftmann@37639: fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) = haftmann@37639: print_def name (vs, ty) (filter (snd o snd) raw_eqs) haftmann@34294: | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) = haftmann@34294: let haftmann@37639: val tyvars = intro_tyvars vs reserved; haftmann@37450: fun print_co ((co, _), []) = haftmann@34294: concat [str "final", str "case", str "object", (str o deresolve_base) co, haftmann@37639: str "extends", applify "[" "]" I NOBR ((str o deresolve_base) name) haftmann@34294: (replicate (length vs) (str "Nothing"))] haftmann@37450: | print_co ((co, vs_args), tys) = haftmann@37639: concat [applify "(" ")" haftmann@37639: (fn (v, arg) => constraint (str v) (print_typ tyvars NOBR arg)) NOBR haftmann@37639: (applify "[" "]" (str o lookup_tyvar tyvars) NOBR ((concat o map str) haftmann@37639: ["final", "case", "class", deresolve_base co]) vs_args) haftmann@37639: (Name.names (snd reserved) "a" tys), haftmann@37639: str "extends", haftmann@37639: applify "[" "]" (str o lookup_tyvar tyvars o fst) NOBR haftmann@37639: ((str o deresolve_base) name) vs haftmann@37639: ]; haftmann@34294: in haftmann@37639: Pretty.chunks (applify "[" "]" (str o prefix "+" o lookup_tyvar tyvars o fst) haftmann@37893: NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve_base name]) vs haftmann@37639: :: map print_co cos) haftmann@34294: end haftmann@37447: | print_stmt (name, Code_Thingol.Class (_, (v, (super_classes, classparams)))) = haftmann@34294: let haftmann@37639: val tyvars = intro_tyvars [(v, [name])] reserved; haftmann@37639: fun add_typarg s = Pretty.block haftmann@37639: [str s, str "[", (str o lookup_tyvar tyvars) v, str "]"]; haftmann@37384: fun print_super_classes [] = NONE haftmann@37384: | print_super_classes classes = SOME (concat (str "extends" haftmann@37639: :: separate (str "with") (map (add_typarg o deresolve o fst) classes))); haftmann@34294: fun print_classparam_val (classparam, ty) = haftmann@37639: concat [str "val", constraint (print_method classparam) haftmann@37639: ((print_tupled_typ tyvars o Code_Thingol.unfold_fun) ty)]; haftmann@34294: fun print_classparam_def (classparam, ty) = haftmann@34294: let haftmann@34294: val (tys, ty) = Code_Thingol.unfold_fun ty; haftmann@37639: val [implicit_name] = Name.invents (snd reserved) (lookup_tyvar tyvars v) 1; haftmann@37639: val proto_vars = intro_vars [implicit_name] reserved; haftmann@37639: val auxs = Name.invents (snd proto_vars) "a" (length tys); haftmann@37639: val vars = intro_vars auxs proto_vars; haftmann@34294: in haftmann@37639: concat [str "def", constraint (Pretty.block [applify "(" ")" haftmann@37639: (fn (aux, ty) => constraint ((str o lookup_var vars) aux) haftmann@38769: (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve_base classparam)) haftmann@37639: (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ", haftmann@37639: add_typarg (deresolve name), str ")"]) (print_typ tyvars NOBR ty), str "=", haftmann@37639: applify "(" ")" (str o lookup_var vars) NOBR haftmann@37639: (Pretty.block [str implicit_name, str ".", print_method classparam]) auxs] haftmann@34294: end; haftmann@34294: in haftmann@34294: Pretty.chunks ( haftmann@34294: (Pretty.block_enclose haftmann@37639: (concat ([str "trait", (add_typarg o deresolve_base) name] haftmann@37384: @ the_list (print_super_classes super_classes) @ [str "{"]), str "}") haftmann@34294: (map print_classparam_val classparams)) haftmann@34294: :: map print_classparam_def classparams haftmann@34294: ) haftmann@34294: end haftmann@34294: | print_stmt (name, Code_Thingol.Classinst ((class, (tyco, vs)), haftmann@37450: (super_instances, (classparam_instances, further_classparam_instances)))) = haftmann@34294: let haftmann@37639: val tyvars = intro_tyvars vs reserved; haftmann@37639: val classtyp = (class, tyco `%% map (ITyVar o fst) vs); haftmann@37450: fun print_classparam_instance ((classparam, const as (_, (_, tys))), (thm, _)) = haftmann@37450: let haftmann@37639: val aux_tys = Name.names (snd reserved) "a" tys; haftmann@37639: val auxs = map fst aux_tys; haftmann@37450: val vars = intro_vars auxs reserved; haftmann@37639: val aux_abstr = if null auxs then [] else [enum "," "(" ")" haftmann@37639: (map (fn (aux, ty) => constraint ((str o lookup_var vars) aux) haftmann@37639: (print_typ tyvars NOBR ty)) aux_tys), str "=>"]; haftmann@37450: in haftmann@37639: concat ([str "val", print_method classparam, str "="] haftmann@37639: @ aux_abstr @| print_app tyvars false (SOME thm) vars NOBR haftmann@37639: (const, map (IVar o SOME) auxs)) haftmann@37450: end; haftmann@37639: in haftmann@37639: Pretty.block_enclose (concat [str "implicit def", haftmann@37639: constraint (print_context tyvars vs name) (print_dicttyp tyvars classtyp), haftmann@37639: str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}") haftmann@37639: (map print_classparam_instance (classparam_instances @ further_classparam_instances)) haftmann@37639: end; haftmann@34294: in print_stmt end; haftmann@34294: haftmann@38769: local haftmann@38769: haftmann@38769: (* hierarchical module name space *) haftmann@38769: haftmann@38769: datatype node = haftmann@38769: Dummy haftmann@38769: | Stmt of Code_Thingol.stmt haftmann@38769: | Module of ((Name.context * Name.context) * Name.context) * (string list * (string * node) Graph.T); haftmann@38769: haftmann@38769: in haftmann@38769: haftmann@38779: fun scala_program_of_program labelled_name reserved module_alias program = haftmann@34294: let haftmann@38769: haftmann@38769: (* building module name hierarchy *) haftmann@38769: fun alias_fragments name = case module_alias name haftmann@38769: of SOME name' => Long_Name.explode name' haftmann@38769: | NONE => map (fn name => fst (yield_singleton Name.variants name reserved)) haftmann@38769: (Long_Name.explode name); haftmann@38769: val module_names = Graph.fold (insert (op =) o fst o dest_name o fst) program []; haftmann@38769: val fragments_tab = fold (fn name => Symtab.update haftmann@38769: (name, alias_fragments name)) module_names Symtab.empty; haftmann@38769: val dest_name = Code_Printer.dest_name #>> (the o Symtab.lookup fragments_tab); haftmann@38769: haftmann@38769: (* building empty module hierarchy *) haftmann@38769: val empty_module = (((reserved, reserved), reserved), ([], Graph.empty)); haftmann@38769: fun map_module f (Module content) = Module (f content); haftmann@38769: fun declare_module name_fragement ((nsp_class, nsp_object), nsp_common) = haftmann@34294: let haftmann@38769: val declare = Name.declare name_fragement; haftmann@38769: in ((declare nsp_class, declare nsp_object), declare nsp_common) end; haftmann@38769: fun ensure_module name_fragement (nsps, (implicits, nodes)) = haftmann@38769: if can (Graph.get_node nodes) name_fragement then (nsps, (implicits, nodes)) haftmann@38769: else haftmann@38769: (nsps |> declare_module name_fragement, (implicits, haftmann@38769: nodes |> Graph.new_node (name_fragement, (name_fragement, Module empty_module)))); haftmann@38769: fun allocate_module [] = I haftmann@38769: | allocate_module (name_fragment :: name_fragments) = haftmann@38769: ensure_module name_fragment haftmann@38769: #> (apsnd o apsnd o Graph.map_node name_fragment o apsnd o map_module o allocate_module) name_fragments; haftmann@38769: val empty_program = Symtab.fold (fn (_, fragments) => allocate_module fragments) haftmann@38769: fragments_tab empty_module; haftmann@38769: fun change_module [] = I haftmann@38769: | change_module (name_fragment :: name_fragments) = haftmann@38769: apsnd o apsnd o Graph.map_node name_fragment o apsnd o map_module haftmann@38769: o change_module name_fragments; haftmann@38769: haftmann@38769: (* statement declaration *) haftmann@38769: fun namify_class base ((nsp_class, nsp_object), nsp_common) = haftmann@38769: let haftmann@38769: val (base', nsp_class') = yield_singleton Name.variants base nsp_class haftmann@38769: in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end; haftmann@38769: fun namify_object base ((nsp_class, nsp_object), nsp_common) = haftmann@38769: let haftmann@38769: val (base', nsp_object') = yield_singleton Name.variants base nsp_object haftmann@38769: in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end; haftmann@38769: fun namify_common upper base ((nsp_class, nsp_object), nsp_common) = haftmann@38769: let haftmann@38769: val (base', nsp_common') = haftmann@38769: yield_singleton Name.variants (if upper then first_upper base else base) nsp_common haftmann@38769: in haftmann@38769: (base', haftmann@38769: ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common')) haftmann@38769: end; haftmann@38769: fun declare_stmt name stmt = haftmann@38769: let haftmann@38769: val (name_fragments, base) = dest_name name; haftmann@38769: val namify = case stmt haftmann@38769: of Code_Thingol.Fun _ => namify_object haftmann@38769: | Code_Thingol.Datatype _ => namify_class haftmann@38769: | Code_Thingol.Datatypecons _ => namify_common true haftmann@38769: | Code_Thingol.Class _ => namify_class haftmann@38769: | Code_Thingol.Classrel _ => namify_object haftmann@38769: | Code_Thingol.Classparam _ => namify_object haftmann@38769: | Code_Thingol.Classinst _ => namify_common false; haftmann@38769: val stmt' = case stmt haftmann@38769: of Code_Thingol.Datatypecons _ => Dummy haftmann@38769: | Code_Thingol.Classrel _ => Dummy haftmann@38769: | Code_Thingol.Classparam _ => Dummy haftmann@38769: | _ => Stmt stmt; haftmann@38769: fun is_classinst stmt = case stmt haftmann@38769: of Code_Thingol.Classinst _ => true haftmann@38769: | _ => false; haftmann@38769: val implicit_deps = filter (is_classinst o Graph.get_node program) haftmann@38769: (Graph.imm_succs program name); haftmann@38769: fun declaration (nsps, (implicits, nodes)) = haftmann@34294: let haftmann@38769: val (base', nsps') = namify base nsps; haftmann@38769: val implicits' = union (op =) implicit_deps implicits; haftmann@38769: val nodes' = Graph.new_node (name, (base', stmt')) nodes; haftmann@38769: in (nsps', (implicits', nodes')) end; haftmann@38769: in change_module name_fragments declaration end; haftmann@38769: haftmann@38769: (* dependencies *) haftmann@38769: fun add_dependency name name' = haftmann@38769: let haftmann@38769: val (name_fragments, base) = dest_name name; haftmann@38769: val (name_fragments', base') = dest_name name'; haftmann@38769: val (name_fragments_common, (diff, diff')) = haftmann@38769: chop_prefix (op =) (name_fragments, name_fragments'); haftmann@38769: val dep = if null diff then (name, name') else (hd diff, hd diff') haftmann@38769: in (change_module name_fragments_common o apsnd o apsnd) (Graph.add_edge dep) end; haftmann@38769: haftmann@38769: (* producing program *) haftmann@38769: val (_, (_, sca_program)) = empty_program haftmann@38769: |> Graph.fold (fn (name, (stmt, _)) => declare_stmt name stmt) program haftmann@38769: |> Graph.fold (fn (name, (_, (_, names))) => fold (add_dependency name) names) program; haftmann@38769: haftmann@38769: (* deresolving *) haftmann@38769: fun deresolve name = haftmann@38769: let haftmann@38769: val (name_fragments, _) = dest_name name; haftmann@38769: val nodes = fold (fn name_fragement => fn nodes => case Graph.get_node nodes name_fragement haftmann@38769: of (_, Module (_, (_, nodes))) => nodes) name_fragments sca_program; haftmann@38769: val (base', _) = Graph.get_node nodes name; haftmann@38769: in Long_Name.implode (name_fragments @ [base']) end haftmann@38769: handle Graph.UNDEF _ => error ("Unknown statement name: " ^ labelled_name name); haftmann@38769: haftmann@38769: in (deresolve, sca_program) end; haftmann@34294: haftmann@38779: fun serialize_scala labelled_name raw_reserved includes module_alias haftmann@37464: _ syntax_tyco syntax_const (code_of_pretty, code_writeln) haftmann@38779: program (stmt_names, presentation_stmt_names) destination = haftmann@34294: let haftmann@38769: haftmann@38769: (* preprocess program *) haftmann@34294: val reserved = fold (insert (op =) o fst) includes raw_reserved; haftmann@38769: val (deresolve, sca_program) = scala_program_of_program labelled_name haftmann@38779: (Name.make_context reserved) module_alias program; haftmann@38769: haftmann@38769: (* print statements *) haftmann@37639: fun lookup_constr tyco constr = case Graph.get_node program tyco haftmann@37639: of Code_Thingol.Datatype (_, (_, constrs)) => haftmann@37639: the (AList.lookup (op = o apsnd fst) constrs constr); haftmann@37639: fun classparams_of_class class = case Graph.get_node program class haftmann@37639: of Code_Thingol.Class (_, (_, (_, classparams))) => classparams; haftmann@34294: fun args_num c = case Graph.get_node program c haftmann@37464: of Code_Thingol.Fun (_, (((_, ty), []), _)) => haftmann@37464: (length o fst o Code_Thingol.unfold_fun) ty haftmann@37437: | Code_Thingol.Fun (_, ((_, ((ts, _), _) :: _), _)) => length ts haftmann@37639: | Code_Thingol.Datatypecons (_, tyco) => length (lookup_constr tyco c) haftmann@34294: | Code_Thingol.Classparam (_, class) => haftmann@37639: (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =) haftmann@37639: (classparams_of_class class)) c; haftmann@37639: fun is_singleton_constr c = case Graph.get_node program c haftmann@37639: of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c) haftmann@34294: | _ => false; haftmann@34294: val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const haftmann@38780: (make_vars reserved) args_num is_singleton_constr haftmann@38780: (deresolve, Long_Name.implode o fst o split_last o Long_Name.explode (*FIXME full*)); haftmann@38769: haftmann@38769: (* print nodes *) haftmann@38769: fun print_implicits [] = NONE haftmann@38769: | print_implicits implicits = (SOME o Pretty.block) haftmann@38778: (str "import /*implicits*/" :: Pretty.brk 1 :: commas (map (str o deresolve) implicits)); haftmann@38769: fun print_module base implicits p = Pretty.chunks2 haftmann@38769: ([str ("object " ^ base ^ " {")] @ the_list (print_implicits implicits) haftmann@38769: @ [p, str ("} /* object " ^ base ^ " */")]); haftmann@38769: fun print_node (_, Dummy) = NONE haftmann@38772: | print_node (name, Stmt stmt) = if null presentation_stmt_names haftmann@38772: orelse member (op =) presentation_stmt_names name haftmann@38769: then SOME (print_stmt (name, stmt)) haftmann@38769: else NONE haftmann@38769: | print_node (name, Module (_, (implicits, nodes))) = if null presentation_stmt_names haftmann@38769: then case print_nodes nodes haftmann@38769: of NONE => NONE haftmann@38769: | SOME p => SOME (print_module (Long_Name.base_name name) implicits p) haftmann@38769: else print_nodes nodes haftmann@38769: and print_nodes nodes = let haftmann@38769: val ps = map_filter (fn name => print_node (name, haftmann@38769: snd (Graph.get_node nodes name))) haftmann@38769: ((rev o flat o Graph.strong_conn) nodes); haftmann@38769: in if null ps then NONE else SOME (Pretty.chunks2 ps) end; haftmann@38769: haftmann@38769: (* serialization *) haftmann@38772: val p_includes = if null presentation_stmt_names haftmann@38772: then map (fn (base, p) => print_module base [] p) includes else []; haftmann@38772: val p = Pretty.chunks2 (p_includes @ the_list (print_nodes sca_program)); haftmann@34294: in haftmann@37748: Code_Target.mk_serialization target haftmann@38769: (fn NONE => code_writeln | SOME file => File.write file o code_of_pretty) haftmann@38769: (rpair [] o code_of_pretty) p destination haftmann@34294: end; haftmann@34294: haftmann@38769: end; (*local*) haftmann@38769: haftmann@34294: val literals = let haftmann@37224: fun char_scala c = if c = "'" then "\\'" haftmann@37224: else if c = "\"" then "\\\"" haftmann@37224: else if c = "\\" then "\\\\" haftmann@37224: else let val k = ord c haftmann@37224: in if k < 32 orelse k > 126 then "\\" ^ radixstring (8, "0", k) else c end haftmann@34944: fun numeral_scala k = if k < 0 haftmann@37958: then if k > ~ 2147483647 then "- " ^ string_of_int (~ k) haftmann@34944: else quote ("- " ^ string_of_int (~ k)) haftmann@34944: else if k <= 2147483647 then string_of_int k haftmann@34944: else quote (string_of_int k) haftmann@34294: in Literals { haftmann@34294: literal_char = Library.enclose "'" "'" o char_scala, haftmann@34294: literal_string = quote o translate_string char_scala, haftmann@34944: literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")", haftmann@38771: literal_positive_numeral = fn k => "Nat.Nat(" ^ numeral_scala k ^ ")", haftmann@38771: literal_alternative_numeral = fn k => "Natural.Nat(" ^ numeral_scala k ^ ")", haftmann@37958: literal_naive_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")", haftmann@34888: literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps], haftmann@34294: infix_cons = (6, "::") haftmann@34294: } end; haftmann@34294: haftmann@34294: haftmann@34294: (** Isar setup **) haftmann@34294: haftmann@38779: fun isar_serializer _ = haftmann@34294: Code_Target.parse_args (Scan.succeed ()) haftmann@38779: #> (fn () => serialize_scala); haftmann@34294: haftmann@34294: val setup = haftmann@37821: Code_Target.add_target haftmann@37822: (target, { serializer = isar_serializer, literals = literals, haftmann@38769: check = { env_var = "SCALA_HOME", make_destination = fn p => Path.append p (Path.explode "ROOT.scala"), haftmann@37822: make_command = fn scala_home => fn p => fn _ => haftmann@37932: "export JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' && " haftmann@38769: ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " " ^ File.shell_path p } }) haftmann@37464: #> Code_Target.add_syntax_tyco target "fun" haftmann@37464: (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] => haftmann@37464: brackify_infix (1, R) fxy ( haftmann@37464: print_typ BR ty1 (*product type vs. tupled arguments!*), haftmann@37464: str "=>", haftmann@37464: print_typ (INFX (1, R)) ty2 haftmann@37464: ))) haftmann@34294: #> fold (Code_Target.add_reserved target) [ haftmann@34294: "abstract", "case", "catch", "class", "def", "do", "else", "extends", "false", haftmann@34294: "final", "finally", "for", "forSome", "if", "implicit", "import", "lazy", haftmann@34294: "match", "new", "null", "object", "override", "package", "private", "protected", haftmann@34294: "requires", "return", "sealed", "super", "this", "throw", "trait", "try", haftmann@37243: "true", "type", "val", "var", "while", "with", "yield" haftmann@34294: ] haftmann@34294: #> fold (Code_Target.add_reserved target) [ haftmann@37639: "apply", "error", "BigInt", "Nil", "List" haftmann@34294: ]; haftmann@34294: haftmann@34294: end; (*struct*)