haftmann@34294: (* 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@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@34294: fun print_scala_stmt labelled_name syntax_tyco syntax_const reserved args_num is_singleton deresolve = haftmann@34294: let haftmann@34294: val deresolve_base = Long_Name.base_name o deresolve; haftmann@37243: val lookup_tyvar = first_upper oo lookup_var; haftmann@34294: fun print_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco haftmann@34294: of NONE => applify "[" "]" fxy haftmann@34294: ((str o deresolve) tyco) (map (print_typ tyvars NOBR) 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@34294: fun print_typed tyvars p ty = haftmann@34294: Pretty.block [p, str ":", Pretty.brk 1, print_typ tyvars NOBR ty] 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@34294: | _ => applify "(" ")" fxy haftmann@35228: (print_term tyvars is_pat some_thm vars BR t1) haftmann@35228: [print_term tyvars is_pat some_thm vars NOBR 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@34294: Pretty.block [str "(", print_typed tyvars (print_var vars' v) ty, str ")"], 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@35228: and print_app tyvars is_pat some_thm vars fxy (app as ((c, ((tys, _), tys_args)), ts)) = haftmann@34294: let haftmann@34294: val k = length ts; haftmann@34294: val l = case syntax_const c of NONE => args_num c | SOME (l, _) => l; haftmann@34294: val tys' = if is_pat orelse haftmann@34294: (is_none (syntax_const c) andalso is_singleton c) then [] else tys; haftmann@34294: val (no_syntax, print') = case syntax_const c haftmann@34294: of NONE => (true, fn ts => applify "(" ")" fxy haftmann@34294: (applify "[" "]" NOBR ((str o deresolve) c) (map (print_typ tyvars NOBR) tys')) haftmann@35228: (map (print_term tyvars is_pat some_thm vars NOBR) ts)) haftmann@34294: | SOME (_, print) => (false, fn ts => haftmann@35228: print (print_term tyvars is_pat some_thm) some_thm vars fxy (ts ~~ take l tys_args)); haftmann@34294: in if k = l then print' 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@34294: Pretty.block (print' 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@35228: and print_bind tyvars some_thm fxy p = 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@34294: fun print_match ((pat, ty), t) vars = haftmann@34294: vars haftmann@35228: |> print_bind tyvars some_thm BR pat haftmann@34294: |>> (fn p => semicolon [Pretty.block [str "val", Pretty.brk 1, p, haftmann@34294: str ":", Pretty.brk 1, print_typ tyvars NOBR ty], haftmann@35228: str "=", print_term tyvars false some_thm vars NOBR t]) haftmann@34294: val (ps, vars') = fold_map print_match binds vars; haftmann@34294: in haftmann@34294: brackify_block fxy haftmann@34294: (str "{") haftmann@35228: (ps @| print_term tyvars false some_thm vars' NOBR body) haftmann@34294: (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@35228: val (p, vars') = print_bind tyvars some_thm NOBR pat vars; haftmann@35228: in concat [str "case", p, str "=>", print_term tyvars false some_thm vars' NOBR 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@34294: fun implicit_arguments tyvars vs vars = haftmann@34294: let haftmann@34294: val implicit_typ_ps = maps (fn (v, sort) => map (fn class => Pretty.block haftmann@37243: [(str o deresolve) class, str "[", (str o lookup_tyvar tyvars) v, str "]"]) sort) vs; haftmann@37243: val implicit_names = Name.variant_list [] (maps (fn (v, sort) => map (fn class => haftmann@37243: lookup_tyvar tyvars v ^ "_" ^ (Long_Name.base_name o deresolve) class) sort) vs); haftmann@34294: val vars' = intro_vars implicit_names vars; haftmann@34294: val implicit_ps = map2 (fn v => fn p => concat [str (v ^ ":"), p]) haftmann@34294: implicit_names implicit_typ_ps; haftmann@34294: in ((implicit_names, implicit_ps), vars') end; haftmann@34294: fun print_defhead tyvars vars p vs params tys implicits ty = haftmann@37243: Pretty.block [str "def ", print_typed tyvars (applify "(implicit " ")" NOBR haftmann@34294: (applify "(" ")" NOBR haftmann@37243: (applify "[" "]" NOBR p (map (str o lookup_tyvar tyvars o fst) vs)) haftmann@34294: (map2 (fn param => fn ty => print_typed tyvars haftmann@34294: ((str o lookup_var vars) param) ty) haftmann@37243: params tys)) implicits) ty, str " ="] haftmann@34294: fun print_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) = (case filter (snd o snd) raw_eqs haftmann@34294: of [] => haftmann@34294: let haftmann@34294: val (tys, ty') = Code_Thingol.unfold_fun ty; haftmann@34294: val params = Name.invents (snd reserved) "a" (length tys); haftmann@34294: val tyvars = intro_vars (map fst vs) reserved; haftmann@34294: val vars = intro_vars params reserved; haftmann@34294: in haftmann@34294: concat [print_defhead tyvars vars ((str o deresolve) name) vs params tys [] ty', haftmann@34294: str ("error(\"" ^ name ^ "\")")] haftmann@34294: end haftmann@34294: | eqs => haftmann@34294: let haftmann@35228: val tycos = fold (fn ((ts, t), _) => haftmann@34294: fold Code_Thingol.add_tyconames (t :: ts)) eqs []; haftmann@34294: val tyvars = reserved haftmann@34294: |> intro_base_names haftmann@34294: (is_none o syntax_tyco) deresolve tycos haftmann@34294: |> intro_vars (map fst vs); haftmann@34294: val simple = case eqs haftmann@34294: of [((ts, _), _)] => forall Code_Thingol.is_IVar ts haftmann@34294: | _ => false; haftmann@34294: val consts = fold Code_Thingol.add_constnames haftmann@34294: (map (snd o fst) eqs) []; haftmann@34294: val vars1 = reserved haftmann@34294: |> intro_base_names haftmann@34294: (is_none o syntax_const) deresolve consts haftmann@34294: val ((_, implicit_ps), vars2) = implicit_arguments tyvars vs vars1; haftmann@34294: val params = if simple then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs haftmann@34294: else aux_params vars2 (map (fst o fst) eqs); haftmann@34294: val vars3 = intro_vars params vars2; haftmann@34294: val (tys, ty1) = Code_Thingol.unfold_fun ty; haftmann@34294: val (tys1, tys2) = chop (length params) tys; haftmann@34294: val ty2 = Library.foldr haftmann@34294: (fn (ty1, ty2) => Code_Thingol.fun_tyco `%% [ty1, ty2]) (tys2, ty1); haftmann@34294: fun print_tuple [p] = p haftmann@34294: | print_tuple ps = enum "," "(" ")" ps; haftmann@35228: fun print_rhs vars' ((_, t), (some_thm, _)) = print_term tyvars false some_thm vars' NOBR t; haftmann@35228: fun print_clause (eq as ((ts, _), (some_thm, _))) = haftmann@34294: let haftmann@34294: val vars' = intro_vars ((fold o Code_Thingol.fold_varnames) (insert (op =)) ts []) vars2; haftmann@34294: in haftmann@35228: concat [str "case", print_tuple (map (print_term tyvars true some_thm vars' NOBR) ts), haftmann@34294: str "=>", print_rhs vars' eq] haftmann@34294: end; haftmann@34294: val head = print_defhead tyvars vars3 ((str o deresolve) name) vs params tys1 implicit_ps ty2; haftmann@34294: in if simple then haftmann@34294: concat [head, print_rhs vars3 (hd eqs)] haftmann@34294: else haftmann@34294: Pretty.block_enclose haftmann@34294: (concat [head, print_tuple (map (str o lookup_var vars3) params), haftmann@34294: str "match", str "{"], str "}") haftmann@34294: (map print_clause eqs) haftmann@34294: end) haftmann@34294: | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) = haftmann@34294: let haftmann@34294: val tyvars = intro_vars (map fst vs) reserved; haftmann@34294: fun print_co (co, []) = haftmann@34294: concat [str "final", str "case", str "object", (str o deresolve_base) co, haftmann@34294: str "extends", applify "[" "]" NOBR ((str o deresolve_base) name) haftmann@34294: (replicate (length vs) (str "Nothing"))] haftmann@34294: | print_co (co, tys) = haftmann@34294: let haftmann@34294: val fields = Name.names (snd reserved) "a" tys; haftmann@34294: val vars = intro_vars (map fst fields) reserved; haftmann@34294: fun add_typargs p = applify "[" "]" NOBR p haftmann@37243: (map (str o lookup_tyvar tyvars o fst) vs); haftmann@34294: in haftmann@34294: concat [ haftmann@34294: applify "(" ")" NOBR haftmann@34294: (add_typargs ((concat o map str) ["final", "case", "class", deresolve_base co])) haftmann@34294: (map (uncurry (print_typed tyvars) o apfst str) fields), haftmann@34294: str "extends", haftmann@34294: add_typargs ((str o deresolve_base) name) haftmann@34294: ] haftmann@34294: end haftmann@34294: in haftmann@34294: Pretty.chunks ( haftmann@34294: applify "[" "]" NOBR ((concat o map str) ["sealed", "class", deresolve_base name]) haftmann@37243: (map (str o prefix "+" o lookup_tyvar tyvars o fst) vs) haftmann@34294: :: map print_co cos haftmann@34294: ) haftmann@34294: end haftmann@34294: | print_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) = haftmann@34294: let haftmann@34294: val tyvars = intro_vars [v] reserved; haftmann@34294: val vs = [(v, [name])]; haftmann@37243: fun add_typarg p = applify "[" "]" NOBR p [(str o lookup_tyvar tyvars) v]; haftmann@34294: fun print_superclasses [] = NONE haftmann@34294: | print_superclasses classes = SOME (concat (str "extends" haftmann@34294: :: separate (str "with") (map (add_typarg o str o deresolve o fst) classes))); haftmann@34294: fun print_tupled_typ ([], ty) = haftmann@34294: print_typ tyvars NOBR ty haftmann@34294: | print_tupled_typ ([ty1], ty2) = haftmann@34294: concat [print_typ tyvars BR ty1, str "=>", print_typ tyvars NOBR ty2] haftmann@34294: | print_tupled_typ (tys, ty) = haftmann@34294: concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys), haftmann@34294: str "=>", print_typ tyvars NOBR ty]; haftmann@34294: fun print_classparam_val (classparam, ty) = haftmann@37337: concat [str "val", (str o Library.enclose "`" "+`:" o deresolve_base) classparam, haftmann@34294: (print_tupled_typ 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@34294: val params = Name.invents (snd reserved) "a" (length tys); haftmann@34294: val vars = intro_vars params reserved; haftmann@34294: val (([implicit], [p_implicit]), vars') = implicit_arguments tyvars vs vars; haftmann@34294: val head = print_defhead tyvars vars' ((str o deresolve) classparam) vs params tys [p_implicit] ty; haftmann@34294: in haftmann@34294: concat [head, applify "(" ")" NOBR haftmann@37337: (Pretty.block [str implicit, str ".", (str o Library.enclose "`" "+`" o deresolve_base) classparam]) haftmann@34294: (map (str o lookup_var vars') params)] haftmann@34294: end; haftmann@34294: in haftmann@34294: Pretty.chunks ( haftmann@34294: (Pretty.block_enclose haftmann@34294: (concat ([str "trait", add_typarg ((str o deresolve_base) name)] haftmann@34294: @ the_list (print_superclasses superclasses) @ [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@34294: (super_instances, classparam_insts))) = haftmann@34294: let haftmann@34294: val tyvars = intro_vars (map fst vs) reserved; haftmann@34294: val insttyp = tyco `%% map (ITyVar o fst) vs; haftmann@34294: val p_inst_typ = print_typ tyvars NOBR insttyp; haftmann@37243: fun add_typ_params p = applify "[" "]" NOBR p (map (str o lookup_tyvar tyvars o fst) vs); haftmann@34294: fun add_inst_typ p = Pretty.block [p, str "[", p_inst_typ, str "]"]; haftmann@34294: val ((implicit_names, implicit_ps), vars) = implicit_arguments tyvars vs reserved; haftmann@34294: fun print_classparam_inst ((classparam, const as (_, (_, tys))), (thm, _)) = haftmann@34294: let haftmann@34294: val auxs = Name.invents (snd reserved) "a" (length tys); haftmann@34294: val vars = intro_vars auxs reserved; haftmann@34294: val args = if null auxs then [] else haftmann@34294: [concat [enum "," "(" ")" (map2 (fn aux => fn ty => print_typed tyvars ((str o lookup_var vars) aux) ty) haftmann@34294: auxs tys), str "=>"]]; haftmann@34294: in haftmann@37337: concat ([str "val", (str o Library.enclose "`" "+`" o deresolve_base) classparam, haftmann@35228: str "="] @ args @ [print_app tyvars false (SOME thm) vars NOBR (const, map (IVar o SOME) auxs)]) haftmann@34294: end; haftmann@34294: in haftmann@34294: Pretty.chunks [ haftmann@34294: Pretty.block_enclose haftmann@34294: (concat ([str "trait", haftmann@34294: add_typ_params ((str o deresolve_base) name), haftmann@34294: str "extends", haftmann@34294: Pretty.block [(str o deresolve) class, str "[", p_inst_typ, str "]"]] haftmann@34294: @ map (fn (_, (_, (superinst, _))) => add_typ_params (str ("with " ^ deresolve superinst))) haftmann@34294: super_instances @| str "{"), str "}") haftmann@34294: (map (fn p => Pretty.block [str "implicit val arg$", p]) implicit_ps haftmann@34294: @ map print_classparam_inst classparam_insts), haftmann@34294: concat [str "implicit", str (if null vs then "val" else "def"), haftmann@34294: applify "(implicit " ")" NOBR (applify "[" "]" NOBR haftmann@37243: ((str o deresolve_base) name) (map (str o lookup_tyvar tyvars o fst) vs)) haftmann@34294: implicit_ps, haftmann@34294: str "=", str "new", applify "[" "]" NOBR ((str o deresolve_base) name) haftmann@37243: (map (str o lookup_tyvar tyvars o fst) vs), haftmann@34294: Pretty.enum ";" "{ " " }" (map (str o (fn v => "val arg$" ^ v ^ " = " ^ v) o lookup_var vars) haftmann@34294: implicit_names)] haftmann@34294: ] haftmann@34294: end; haftmann@34294: in print_stmt end; haftmann@34294: haftmann@34294: fun scala_program_of_program labelled_name module_name reserved raw_module_alias program = haftmann@34294: let haftmann@34294: val the_module_name = the_default "Program" module_name; haftmann@34294: val module_alias = K (SOME the_module_name); haftmann@34294: val reserved = Name.make_context reserved; haftmann@34294: fun prepare_stmt (name, stmt) (nsps, stmts) = haftmann@34294: let haftmann@34294: val (_, base) = Code_Printer.dest_name name; haftmann@34294: val mk_name_stmt = yield_singleton Name.variants; haftmann@34294: fun add_class ((nsp_class, nsp_object), nsp_common) = haftmann@34294: let haftmann@34294: val (base', nsp_class') = mk_name_stmt base nsp_class haftmann@34294: in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end; haftmann@34294: fun add_object ((nsp_class, nsp_object), nsp_common) = haftmann@34294: let haftmann@34294: val (base', nsp_object') = mk_name_stmt base nsp_object haftmann@34294: in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end; haftmann@34294: fun add_common upper ((nsp_class, nsp_object), nsp_common) = haftmann@34294: let haftmann@34294: val (base', nsp_common') = haftmann@34294: mk_name_stmt (if upper then first_upper base else base) nsp_common haftmann@34294: in (base', ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common')) end; haftmann@34294: val add_name = case stmt haftmann@34294: of Code_Thingol.Fun _ => add_object haftmann@34294: | Code_Thingol.Datatype _ => add_class haftmann@34294: | Code_Thingol.Datatypecons _ => add_common true haftmann@34294: | Code_Thingol.Class _ => add_class haftmann@34294: | Code_Thingol.Classrel _ => add_object haftmann@34294: | Code_Thingol.Classparam _ => add_object haftmann@34294: | Code_Thingol.Classinst _ => add_common false; haftmann@34294: fun add_stmt base' = case stmt haftmann@34294: of Code_Thingol.Datatypecons _ => cons (name, (base', NONE)) haftmann@34294: | Code_Thingol.Classrel _ => cons (name, (base', NONE)) haftmann@34294: | Code_Thingol.Classparam _ => cons (name, (base', NONE)) haftmann@34294: | _ => cons (name, (base', SOME stmt)); haftmann@34294: in haftmann@34294: nsps haftmann@34294: |> add_name haftmann@34294: |-> (fn base' => rpair (add_stmt base' stmts)) haftmann@34294: end; haftmann@34294: val (_, sca_program) = fold prepare_stmt (AList.make (fn name => Graph.get_node program name) haftmann@34294: (Graph.strong_conn program |> flat)) (((reserved, reserved), reserved), []); haftmann@34294: fun deresolver name = (fst o the o AList.lookup (op =) sca_program) name haftmann@34294: handle Option => error ("Unknown statement name: " ^ labelled_name name); haftmann@34294: in (deresolver, (the_module_name, sca_program)) end; haftmann@34294: haftmann@34294: fun serialize_scala raw_module_name labelled_name haftmann@34294: raw_reserved includes raw_module_alias haftmann@36535: _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination = haftmann@34294: let haftmann@36535: val presentation_stmt_names = Code_Target.stmt_names_of_destination destination; haftmann@36535: val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code"; haftmann@34294: val reserved = fold (insert (op =) o fst) includes raw_reserved; haftmann@34294: val (deresolver, (the_module_name, sca_program)) = scala_program_of_program labelled_name haftmann@34294: module_name reserved raw_module_alias program; haftmann@34294: val reserved = make_vars reserved; haftmann@34294: fun args_num c = case Graph.get_node program c haftmann@34294: of Code_Thingol.Fun (_, ((_, ty), [])) => (length o fst o Code_Thingol.unfold_fun) ty haftmann@34294: | Code_Thingol.Fun (_, (_, ((ts, _), _) :: _)) => length ts haftmann@34294: | Code_Thingol.Datatypecons (_, tyco) => haftmann@35228: let val Code_Thingol.Datatype (_, (_, constrs)) = Graph.get_node program tyco haftmann@35228: in (length o the o AList.lookup (op =) constrs) c end haftmann@34294: | Code_Thingol.Classparam (_, class) => haftmann@34294: let val Code_Thingol.Class (_, (_, (_, classparams))) = Graph.get_node program class haftmann@34294: in (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =) classparams) c end; haftmann@34294: fun is_singleton c = case Graph.get_node program c haftmann@34294: of Code_Thingol.Datatypecons (_, tyco) => haftmann@35228: let val Code_Thingol.Datatype (_, (_, constrs)) = Graph.get_node program tyco haftmann@35228: in (null o the o AList.lookup (op =) constrs) c end haftmann@34294: | _ => false; haftmann@34294: val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const haftmann@34294: reserved args_num is_singleton deresolver; haftmann@34294: fun print_module name content = haftmann@34294: (name, Pretty.chunks [ haftmann@34294: str ("object " ^ name ^ " {"), haftmann@34294: str "", haftmann@34294: content, haftmann@34294: str "", haftmann@34294: str "}" haftmann@34294: ]); haftmann@34294: fun serialize_module the_module_name sca_program = haftmann@34294: let haftmann@34294: val content = Pretty.chunks2 (map_filter haftmann@34294: (fn (name, (_, SOME stmt)) => SOME (print_stmt (name, stmt)) haftmann@34294: | (_, (_, NONE)) => NONE) sca_program); haftmann@34294: in print_module the_module_name content end; haftmann@34294: fun check_destination destination = haftmann@34294: (File.check destination; destination); haftmann@34294: fun write_module destination (modlname, content) = haftmann@34294: let haftmann@34294: val filename = case modlname haftmann@34294: of "" => Path.explode "Main.scala" haftmann@34294: | _ => (Path.ext "scala" o Path.explode o implode o separate "/" haftmann@34294: o Long_Name.explode) modlname; haftmann@34294: val pathname = Path.append destination filename; haftmann@34294: val _ = File.mkdir (Path.dir pathname); haftmann@34294: in File.write pathname (code_of_pretty content) end haftmann@34294: in haftmann@34294: Code_Target.mk_serialization target NONE haftmann@34294: (fn NONE => K () o map (code_writeln o snd) | SOME file => K () o map haftmann@34294: (write_module (check_destination file))) haftmann@34294: (rpair [] o cat_lines o map (code_of_pretty o snd)) haftmann@34294: (map (uncurry print_module) includes haftmann@34294: @| serialize_module the_module_name sca_program) haftmann@34294: destination haftmann@34294: end; haftmann@34294: 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@34944: 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@34944: literal_positive_numeral = fn k => "Nat.Nat(" ^ numeral_scala k ^ ")", haftmann@34944: literal_naive_numeral = fn k => if k >= 0 haftmann@34944: then string_of_int k else "(- " ^ string_of_int (~ 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@34294: fun isar_seri_scala module_name = haftmann@34294: Code_Target.parse_args (Scan.succeed ()) haftmann@34294: #> (fn () => serialize_scala module_name); haftmann@34294: haftmann@34294: val setup = haftmann@34294: Code_Target.add_target (target, (isar_seri_scala, literals)) haftmann@34294: #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] => haftmann@37243: brackify_infix (1, R) fxy ( haftmann@34900: print_typ BR ty1 (*product type vs. tupled arguments!*), haftmann@34294: str "=>", haftmann@34294: print_typ (INFX (1, R)) ty2 haftmann@37243: ))) 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@34944: "error", "apply", "List", "Nil", "BigInt" haftmann@34294: ]; haftmann@34294: haftmann@34294: end; (*struct*)