src/Tools/Code/code_scala.ML
changeset 39148 b6530978c14d
parent 39142 f63715f00fdd
parent 39147 3c284a152bd6
child 39781 2053638a2bf2
equal deleted inserted replaced
39146:142f1425e074 39148:b6530978c14d
   331 fun serialize_scala { labelled_name, reserved_syms, includes,
   331 fun serialize_scala { labelled_name, reserved_syms, includes,
   332     module_alias, class_syntax, tyco_syntax, const_syntax, program } =
   332     module_alias, class_syntax, tyco_syntax, const_syntax, program } =
   333   let
   333   let
   334 
   334 
   335     (* build program *)
   335     (* build program *)
   336     val { deresolver, hierarchical_program = sca_program } =
   336     val { deresolver, hierarchical_program = scala_program } =
   337       scala_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program;
   337       scala_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program;
   338 
   338 
   339     (* print statements *)
   339     (* print statements *)
   340     fun lookup_constr tyco constr = case Graph.get_node program tyco
   340     fun lookup_constr tyco constr = case Graph.get_node program tyco
   341      of Code_Thingol.Datatype (_, (_, constrs)) =>
   341      of Code_Thingol.Datatype (_, (_, constrs)) =>
   351           (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =)
   351           (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =)
   352             (classparams_of_class class)) c;
   352             (classparams_of_class class)) c;
   353     fun is_singleton_constr c = case Graph.get_node program c
   353     fun is_singleton_constr c = case Graph.get_node program c
   354      of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
   354      of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
   355       | _ => false;
   355       | _ => false;
   356     val print_stmt = print_scala_stmt labelled_name tyco_syntax const_syntax
   356     fun print_stmt prefix_fragments = print_scala_stmt labelled_name
   357       (make_vars reserved_syms) args_num is_singleton_constr;
   357       tyco_syntax const_syntax (make_vars reserved_syms) args_num
   358 
   358       is_singleton_constr (deresolver prefix_fragments, deresolver []);
   359     (* print nodes *)
   359 
   360     fun print_module base implicit_ps p = Pretty.chunks2
   360     (* print modules *)
   361       ([str ("object " ^ base ^ " {")]
       
   362         @ (if null implicit_ps then [] else (single o Pretty.block)
       
   363             (str "import /*implicits*/" :: Pretty.brk 1 :: commas implicit_ps))
       
   364         @ [p, str ("} /* object " ^ base ^ " */")]);
       
   365     fun print_implicit prefix_fragments implicit =
   361     fun print_implicit prefix_fragments implicit =
   366       let
   362       let
   367         val s = deresolver prefix_fragments implicit;
   363         val s = deresolver prefix_fragments implicit;
   368       in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end;
   364       in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end;
   369     fun print_node _ (_, Code_Namespace.Dummy) = NONE
   365     fun print_module prefix_fragments base implicits ps = Pretty.chunks2
   370       | print_node prefix_fragments (name, Code_Namespace.Stmt stmt) =
   366       ([str ("object " ^ base ^ " {")]
   371           SOME (markup_stmt name (print_stmt (deresolver prefix_fragments, deresolver []) (name, stmt)))
   367         @ (case map_filter (print_implicit prefix_fragments) implicits
   372       | print_node prefix_fragments (name_fragment, Code_Namespace.Module (implicits, nodes)) =
   368             of [] => [] | implicit_ps => (single o Pretty.block)
   373           let
   369             (str "import /*implicits*/" :: Pretty.brk 1 :: commas implicit_ps))
   374             val prefix_fragments' = prefix_fragments @ [name_fragment];
   370         @ ps @ [str ("} /* object " ^ base ^ " */")]);
   375           in
       
   376             Option.map (print_module name_fragment
       
   377               (map_filter (print_implicit prefix_fragments') implicits))
       
   378                 (print_nodes prefix_fragments' nodes)
       
   379           end
       
   380     and print_nodes prefix_fragments nodes = let
       
   381         val ps = map_filter (fn name => print_node prefix_fragments (name,
       
   382           snd (Graph.get_node nodes name)))
       
   383             ((rev o flat o Graph.strong_conn) nodes);
       
   384       in if null ps then NONE else SOME (Pretty.chunks2 ps) end;
       
   385 
   371 
   386     (* serialization *)
   372     (* serialization *)
   387     val p = Pretty.chunks2 (map snd includes @ the_list (print_nodes [] sca_program));
   373     val p = Pretty.chunks2 (map snd includes
       
   374       @ Code_Namespace.print_hierarchical {
       
   375         print_module = print_module, print_stmt = print_stmt,
       
   376         lift_markup = I } scala_program);
   388     fun write width NONE = writeln o format [] width
   377     fun write width NONE = writeln o format [] width
   389       | write width (SOME p) = File.write p o format [] width;
   378       | write width (SOME p) = File.write p o format [] width;
   390   in
   379   in
   391     Code_Target.serialization write (rpair (try (deresolver [])) ooo format) p
   380     Code_Target.serialization write (rpair (try (deresolver [])) ooo format) p
   392   end;
   381   end;