src/Tools/Code/code_scala.ML
changeset 38769 317e64c886d2
parent 38059 72f4630d4c43
child 38771 f9cd27cbe8a4
equal deleted inserted replaced
38768:ecc713816e33 38769:317e64c886d2
   133       | print_case tyvars some_thm vars fxy ((_, []), _) =
   133       | print_case tyvars some_thm vars fxy ((_, []), _) =
   134           (brackify fxy o Pretty.breaks o map str) ["error(\"empty case\")"];
   134           (brackify fxy o Pretty.breaks o map str) ["error(\"empty case\")"];
   135     fun print_context tyvars vs name = applify "[" "]"
   135     fun print_context tyvars vs name = applify "[" "]"
   136       (fn (v, sort) => (Pretty.block o map str)
   136       (fn (v, sort) => (Pretty.block o map str)
   137         (lookup_tyvar tyvars v :: maps (fn sort => [": ", deresolve sort]) sort))
   137         (lookup_tyvar tyvars v :: maps (fn sort => [": ", deresolve sort]) sort))
   138           NOBR ((str o deresolve) name) vs;
   138           NOBR ((str o deresolve_base) name) vs;
   139     fun print_defhead tyvars vars name vs params tys ty =
   139     fun print_defhead tyvars vars name vs params tys ty =
   140       Pretty.block [str "def ", constraint (applify "(" ")" (fn (param, ty) =>
   140       Pretty.block [str "def ", constraint (applify "(" ")" (fn (param, ty) =>
   141         constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty))
   141         constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty))
   142           NOBR (print_context tyvars vs name) (params ~~ tys)) (print_typ tyvars NOBR ty),
   142           NOBR (print_context tyvars vs name) (params ~~ tys)) (print_typ tyvars NOBR ty),
   143             str " ="];
   143             str " ="];
   192             Pretty.block_enclose
   192             Pretty.block_enclose
   193               (concat [head, print_tuple (map (str o lookup_var vars2) params),
   193               (concat [head, print_tuple (map (str o lookup_var vars2) params),
   194                 str "match", str "{"], str "}")
   194                 str "match", str "{"], str "}")
   195               (map print_clause eqs)
   195               (map print_clause eqs)
   196           end;
   196           end;
   197     val print_method = (str o Library.enclose "`" "+`" o deresolve_base);
   197     val print_method = str o Library.enclose "`" "`" o space_implode "+"
       
   198       o fst o split_last o Long_Name.explode;
   198     fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
   199     fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
   199           print_def name (vs, ty) (filter (snd o snd) raw_eqs)
   200           print_def name (vs, ty) (filter (snd o snd) raw_eqs)
   200       | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) =
   201       | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) =
   201           let
   202           let
   202             val tyvars = intro_tyvars vs reserved;
   203             val tyvars = intro_tyvars vs reserved;
   238                 val auxs = Name.invents (snd proto_vars) "a" (length tys);
   239                 val auxs = Name.invents (snd proto_vars) "a" (length tys);
   239                 val vars = intro_vars auxs proto_vars;
   240                 val vars = intro_vars auxs proto_vars;
   240               in
   241               in
   241                 concat [str "def", constraint (Pretty.block [applify "(" ")"
   242                 concat [str "def", constraint (Pretty.block [applify "(" ")"
   242                   (fn (aux, ty) => constraint ((str o lookup_var vars) aux)
   243                   (fn (aux, ty) => constraint ((str o lookup_var vars) aux)
   243                   (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve classparam))
   244                   (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve_base classparam))
   244                   (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ",
   245                   (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ",
   245                   add_typarg (deresolve name), str ")"]) (print_typ tyvars NOBR ty), str "=",
   246                   add_typarg (deresolve name), str ")"]) (print_typ tyvars NOBR ty), str "=",
   246                   applify "(" ")" (str o lookup_var vars) NOBR
   247                   applify "(" ")" (str o lookup_var vars) NOBR
   247                   (Pretty.block [str implicit_name, str ".", print_method classparam]) auxs]
   248                   (Pretty.block [str implicit_name, str ".", print_method classparam]) auxs]
   248               end;
   249               end;
   279               str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}")
   280               str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}")
   280                 (map print_classparam_instance (classparam_instances @ further_classparam_instances))
   281                 (map print_classparam_instance (classparam_instances @ further_classparam_instances))
   281           end;
   282           end;
   282   in print_stmt end;
   283   in print_stmt end;
   283 
   284 
       
   285 local
       
   286 
       
   287 (* hierarchical module name space *)
       
   288 
       
   289 datatype node =
       
   290     Dummy
       
   291   | Stmt of Code_Thingol.stmt
       
   292   | Module of ((Name.context * Name.context) * Name.context) * (string list * (string * node) Graph.T);
       
   293 
       
   294 in
       
   295 
   284 fun scala_program_of_program labelled_name module_name reserved raw_module_alias program =
   296 fun scala_program_of_program labelled_name module_name reserved raw_module_alias program =
   285   let
   297   let
   286     val the_module_name = the_default "Program" module_name;
   298 
   287     val module_alias = K (SOME the_module_name);
   299     (* building module name hierarchy *)
   288     val reserved = Name.make_context reserved;
   300     val module_alias = if is_some module_name then K module_name else raw_module_alias;
   289     fun prepare_stmt (name, stmt) (nsps, stmts) =
   301     fun alias_fragments name = case module_alias name
   290       let
   302      of SOME name' => Long_Name.explode name'
   291         val (_, base) = Code_Printer.dest_name name;
   303       | NONE => map (fn name => fst (yield_singleton Name.variants name reserved))
   292         val mk_name_stmt = yield_singleton Name.variants;
   304           (Long_Name.explode name);
   293         fun add_class ((nsp_class, nsp_object), nsp_common) =
   305     val module_names = Graph.fold (insert (op =) o fst o dest_name o fst) program [];
   294           let
   306     val fragments_tab = fold (fn name => Symtab.update
   295             val (base', nsp_class') = mk_name_stmt base nsp_class
   307       (name, alias_fragments name)) module_names Symtab.empty;
   296           in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
   308     val dest_name = Code_Printer.dest_name #>> (the o Symtab.lookup fragments_tab);
   297         fun add_object ((nsp_class, nsp_object), nsp_common) =
   309 
   298           let
   310     (* building empty module hierarchy *)
   299             val (base', nsp_object') = mk_name_stmt base nsp_object
   311     val empty_module = (((reserved, reserved), reserved), ([], Graph.empty));
   300           in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end;
   312     fun map_module f (Module content) = Module (f content);
   301         fun add_common upper ((nsp_class, nsp_object), nsp_common) =
   313     fun declare_module name_fragement ((nsp_class, nsp_object), nsp_common) =
   302           let
   314       let
   303             val (base', nsp_common') =
   315         val declare = Name.declare name_fragement;
   304               mk_name_stmt (if upper then first_upper base else base) nsp_common
   316       in ((declare nsp_class, declare nsp_object), declare nsp_common) end;
   305           in
   317     fun ensure_module name_fragement (nsps, (implicits, nodes)) =
   306             (base',
   318       if can (Graph.get_node nodes) name_fragement then (nsps, (implicits, nodes))
   307               ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
   319       else
   308           end;
   320         (nsps |> declare_module name_fragement, (implicits,
   309         val add_name = case stmt
   321           nodes |> Graph.new_node (name_fragement, (name_fragement, Module empty_module))));
   310          of Code_Thingol.Fun _ => add_object
   322     fun allocate_module [] = I
   311           | Code_Thingol.Datatype _ => add_class
   323       | allocate_module (name_fragment :: name_fragments) =
   312           | Code_Thingol.Datatypecons _ => add_common true
   324           ensure_module name_fragment
   313           | Code_Thingol.Class _ => add_class
   325           #> (apsnd o apsnd o Graph.map_node name_fragment o apsnd o map_module o allocate_module) name_fragments;
   314           | Code_Thingol.Classrel _ => add_object
   326     val empty_program = Symtab.fold (fn (_, fragments) => allocate_module fragments)
   315           | Code_Thingol.Classparam _ => add_object
   327       fragments_tab empty_module;
   316           | Code_Thingol.Classinst _ => add_common false;
   328     fun change_module [] = I
   317         fun add_stmt base' = case stmt
   329       | change_module (name_fragment :: name_fragments) =
   318          of Code_Thingol.Datatypecons _ => cons (name, (base', NONE))
   330           apsnd o apsnd o Graph.map_node name_fragment o apsnd o map_module
   319           | Code_Thingol.Classrel _ => cons (name, (base', NONE))
   331             o change_module name_fragments;
   320           | Code_Thingol.Classparam _ => cons (name, (base', NONE))
   332 
   321           | _ => cons (name, (base', SOME stmt));
   333     (* statement declaration *)
       
   334     fun namify_class base ((nsp_class, nsp_object), nsp_common) =
       
   335       let
       
   336         val (base', nsp_class') = yield_singleton Name.variants base nsp_class
       
   337       in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
       
   338     fun namify_object base ((nsp_class, nsp_object), nsp_common) =
       
   339       let
       
   340         val (base', nsp_object') = yield_singleton Name.variants base nsp_object
       
   341       in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end;
       
   342     fun namify_common upper base ((nsp_class, nsp_object), nsp_common) =
       
   343       let
       
   344         val (base', nsp_common') =
       
   345           yield_singleton Name.variants (if upper then first_upper base else base) nsp_common
   322       in
   346       in
   323         nsps
   347         (base',
   324         |> add_name
   348           ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
   325         |-> (fn base' => rpair (add_stmt base' stmts))
       
   326       end;
   349       end;
   327     val stmts = AList.make (Graph.get_node program) (Graph.strong_conn program |> flat)
   350     fun declare_stmt name stmt =
   328       |> filter_out (Code_Thingol.is_case o snd);
   351       let
   329     val (_, sca_program) = fold prepare_stmt stmts (((reserved, reserved), reserved), []);
   352         val (name_fragments, base) = dest_name name;
   330     fun deresolver name = (fst o the o AList.lookup (op =) sca_program) name
   353         val namify = case stmt
   331       handle Option => error ("Unknown statement name: " ^ labelled_name name);
   354          of Code_Thingol.Fun _ => namify_object
   332   in (deresolver, (the_module_name, sca_program)) end;
   355           | Code_Thingol.Datatype _ => namify_class
       
   356           | Code_Thingol.Datatypecons _ => namify_common true
       
   357           | Code_Thingol.Class _ => namify_class
       
   358           | Code_Thingol.Classrel _ => namify_object
       
   359           | Code_Thingol.Classparam _ => namify_object
       
   360           | Code_Thingol.Classinst _ => namify_common false;
       
   361         val stmt' = case stmt
       
   362          of Code_Thingol.Datatypecons _ => Dummy
       
   363           | Code_Thingol.Classrel _ => Dummy
       
   364           | Code_Thingol.Classparam _ => Dummy
       
   365           | _ => Stmt stmt;
       
   366         fun is_classinst stmt = case stmt
       
   367          of Code_Thingol.Classinst _ => true
       
   368           | _ => false;
       
   369         val implicit_deps = filter (is_classinst o Graph.get_node program)
       
   370           (Graph.imm_succs program name);
       
   371         fun declaration (nsps, (implicits, nodes)) =
       
   372           let
       
   373             val (base', nsps') = namify base nsps;
       
   374             val implicits' = union (op =) implicit_deps implicits;
       
   375             val nodes' = Graph.new_node (name, (base', stmt')) nodes;
       
   376           in (nsps', (implicits', nodes')) end;
       
   377       in change_module name_fragments declaration end;
       
   378 
       
   379     (* dependencies *)
       
   380     fun add_dependency name name' =
       
   381       let
       
   382         val (name_fragments, base) = dest_name name;
       
   383         val (name_fragments', base') = dest_name name';
       
   384         val (name_fragments_common, (diff, diff')) =
       
   385           chop_prefix (op =) (name_fragments, name_fragments');
       
   386         val dep = if null diff then (name, name') else (hd diff, hd diff')
       
   387       in (change_module name_fragments_common o apsnd o apsnd) (Graph.add_edge dep) end;
       
   388 
       
   389     (* producing program *)
       
   390     val (_, (_, sca_program)) = empty_program
       
   391       |> Graph.fold (fn (name, (stmt, _)) => declare_stmt name stmt) program
       
   392       |> Graph.fold (fn (name, (_, (_, names))) => fold (add_dependency name) names) program;
       
   393 
       
   394     (* deresolving *)
       
   395     fun deresolve name =
       
   396       let
       
   397         val (name_fragments, _) = dest_name name;
       
   398         val nodes = fold (fn name_fragement => fn nodes => case Graph.get_node nodes name_fragement
       
   399          of (_, Module (_, (_, nodes))) => nodes) name_fragments sca_program;
       
   400         val (base', _) = Graph.get_node nodes name;
       
   401       in Long_Name.implode (name_fragments @ [base']) end
       
   402         handle Graph.UNDEF _ => error ("Unknown statement name: " ^ labelled_name name);
       
   403 
       
   404   in (deresolve, sca_program) end;
   333 
   405 
   334 fun serialize_scala raw_module_name labelled_name
   406 fun serialize_scala raw_module_name labelled_name
   335     raw_reserved includes raw_module_alias
   407     raw_reserved includes raw_module_alias
   336     _ syntax_tyco syntax_const (code_of_pretty, code_writeln)
   408     _ syntax_tyco syntax_const (code_of_pretty, code_writeln)
   337     program stmt_names destination =
   409     program stmt_names destination =
   338   let
   410   let
       
   411 
       
   412     (* generic nonsense *)
   339     val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
   413     val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
   340     val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
   414     val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
       
   415 
       
   416     (* preprocess program *)
   341     val reserved = fold (insert (op =) o fst) includes raw_reserved;
   417     val reserved = fold (insert (op =) o fst) includes raw_reserved;
   342     val (deresolver, (the_module_name, sca_program)) = scala_program_of_program labelled_name
   418     val (deresolve, sca_program) = scala_program_of_program labelled_name
   343       module_name reserved raw_module_alias program;
   419       module_name (Name.make_context reserved) raw_module_alias program;
   344     val reserved = make_vars reserved;
   420 
       
   421     (* print statements *)
   345     fun lookup_constr tyco constr = case Graph.get_node program tyco
   422     fun lookup_constr tyco constr = case Graph.get_node program tyco
   346      of Code_Thingol.Datatype (_, (_, constrs)) =>
   423      of Code_Thingol.Datatype (_, (_, constrs)) =>
   347           the (AList.lookup (op = o apsnd fst) constrs constr);
   424           the (AList.lookup (op = o apsnd fst) constrs constr);
   348     fun classparams_of_class class = case Graph.get_node program class
   425     fun classparams_of_class class = case Graph.get_node program class
   349      of Code_Thingol.Class (_, (_, (_, classparams))) => classparams;
   426      of Code_Thingol.Class (_, (_, (_, classparams))) => classparams;
   357             (classparams_of_class class)) c;
   434             (classparams_of_class class)) c;
   358     fun is_singleton_constr c = case Graph.get_node program c
   435     fun is_singleton_constr c = case Graph.get_node program c
   359      of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
   436      of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
   360       | _ => false;
   437       | _ => false;
   361     val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const
   438     val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const
   362       reserved args_num is_singleton_constr deresolver;
   439       (make_vars reserved) args_num is_singleton_constr deresolve;
   363     fun print_module name imports content =
   440 
   364       (name, Pretty.chunks (
   441     (* print nodes *)
   365         str ("object " ^ name ^ " {")
   442     fun print_implicits [] = NONE
   366         :: (if null imports then []
   443       | print_implicits implicits = (SOME o Pretty.block)
   367           else str "" :: map (fn name => str ("import " ^ name ^ "._")) imports)
   444           (str "import /*implicits*/" :: Pretty.brk 1 :: Pretty.commas (map (str o deresolve) implicits));
   368         @ [str "",
   445     fun print_module base implicits p = Pretty.chunks2
   369         content,
   446       ([str ("object " ^ base ^ " {")] @ the_list (print_implicits implicits)
   370         str "",
   447         @ [p, str ("} /* object " ^ base ^ " */")]);
   371         str "}"]
   448     fun print_node (_, Dummy) = NONE
   372       ));
   449       | print_node (name, Stmt stmt) = if not (not (null presentation_stmt_names)
   373     fun serialize_module the_module_name sca_program =
   450           andalso member (op =) presentation_stmt_names name)
   374       let
   451           then SOME (print_stmt (name, stmt))
   375         val content = Pretty.chunks2 (map_filter
   452           else NONE
   376           (fn (name, (_, SOME stmt)) => SOME (print_stmt (name, stmt))
   453       | print_node (name, Module (_, (implicits, nodes))) = if null presentation_stmt_names
   377             | (_, (_, NONE)) => NONE) sca_program);
   454           then case print_nodes nodes
   378       in print_module the_module_name (map fst includes) content end;
   455            of NONE => NONE
   379     fun check_destination destination =
   456             | SOME p => SOME (print_module (Long_Name.base_name name) implicits p)
   380       (File.check destination; destination);
   457           else print_nodes nodes
   381     fun write_module destination (modlname, content) =
   458     and print_nodes nodes = let
   382       let
   459         val ps = map_filter (fn name => print_node (name,
   383         val filename = case modlname
   460           snd (Graph.get_node nodes name)))
   384          of "" => Path.explode "Main.scala"
   461             ((rev o flat o Graph.strong_conn) nodes);
   385           | _ => (Path.ext "scala" o Path.explode o implode o separate "/"
   462       in if null ps then NONE else SOME (Pretty.chunks2 ps) end;
   386                 o Long_Name.explode) modlname;
   463 
   387         val pathname = Path.append destination filename;
   464     (* serialization *)
   388         val _ = File.mkdir_leaf (Path.dir pathname);
   465     val p = Pretty.chunks2 (map (fn (base, p) => print_module base [] p) includes
   389       in File.write pathname (code_of_pretty content) end
   466       @ the_list (print_nodes sca_program));
   390   in
   467   in
   391     Code_Target.mk_serialization target
   468     Code_Target.mk_serialization target
   392       (fn NONE => K () o map (code_writeln o snd) | SOME file => K () o map
   469       (fn NONE => code_writeln | SOME file => File.write file o code_of_pretty)
   393         (write_module (check_destination file)))
   470       (rpair [] o code_of_pretty) p destination
   394       (rpair [] o cat_lines o map (code_of_pretty o snd))
       
   395       (map (fn (name, content) => print_module name [] content) includes
       
   396         @| serialize_module the_module_name sca_program)
       
   397       destination
       
   398   end;
   471   end;
       
   472 
       
   473 end; (*local*)
   399 
   474 
   400 val literals = let
   475 val literals = let
   401   fun char_scala c = if c = "'" then "\\'"
   476   fun char_scala c = if c = "'" then "\\'"
   402     else if c = "\"" then "\\\""
   477     else if c = "\"" then "\\\""
   403     else if c = "\\" then "\\\\"
   478     else if c = "\\" then "\\\\"
   427   #> (fn () => serialize_scala module_name);
   502   #> (fn () => serialize_scala module_name);
   428 
   503 
   429 val setup =
   504 val setup =
   430   Code_Target.add_target
   505   Code_Target.add_target
   431     (target, { serializer = isar_serializer, literals = literals,
   506     (target, { serializer = isar_serializer, literals = literals,
   432       check = { env_var = "SCALA_HOME", make_destination = I,
   507       check = { env_var = "SCALA_HOME", make_destination = fn p => Path.append p (Path.explode "ROOT.scala"),
   433         make_command = fn scala_home => fn p => fn _ =>
   508         make_command = fn scala_home => fn p => fn _ =>
   434           "export JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' && "
   509           "export JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' && "
   435             ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " *.scala" } })
   510             ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " " ^ File.shell_path p } })
   436   #> Code_Target.add_syntax_tyco target "fun"
   511   #> Code_Target.add_syntax_tyco target "fun"
   437      (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
   512      (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
   438         brackify_infix (1, R) fxy (
   513         brackify_infix (1, R) fxy (
   439           print_typ BR ty1 (*product type vs. tupled arguments!*),
   514           print_typ BR ty1 (*product type vs. tupled arguments!*),
   440           str "=>",
   515           str "=>",