src/Tools/Code/code_haskell.ML
changeset 55679 59244fc1a7ca
parent 55374 636a8523876f
child 55681 7714287dc044
equal deleted inserted replaced
55678:aec339197a40 55679:59244fc1a7ca
   353     fun print_stmt deresolve = print_haskell_stmt
   353     fun print_stmt deresolve = print_haskell_stmt
   354       class_syntax tyco_syntax const_syntax (make_vars reserved)
   354       class_syntax tyco_syntax const_syntax (make_vars reserved)
   355       deresolve (if string_classes then deriving_show else K false);
   355       deresolve (if string_classes then deriving_show else K false);
   356 
   356 
   357     (* print modules *)
   357     (* print modules *)
   358     fun print_module_frame module_name ps =
   358     fun print_module_frame module_name exports ps =
   359       (module_name, Pretty.chunks2 (
   359       (module_name, Pretty.chunks2 (
   360         str ("module " ^ module_name ^ " where {")
   360         concat [str "module",
       
   361           case exports of
       
   362             SOME ps => Pretty.block [str module_name, enclose "(" ")" (commas ps)]
       
   363           | NONE => str module_name,
       
   364           str "where {"
       
   365         ]
   361         :: ps
   366         :: ps
   362         @| str "}"
   367         @| str "}"
   363       ));
   368       ));
   364     fun print_qualified_import module_name = semicolon [str "import qualified", str module_name];
   369     fun print_qualified_import module_name =
       
   370       semicolon [str "import qualified", str module_name];
   365     val import_common_ps =
   371     val import_common_ps =
   366       enclose "import Prelude (" ");" (commas (map str
   372       enclose "import Prelude (" ");" (commas (map str
   367         (map (Library.enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
   373         (map (Library.enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
   368           @ map (fn (tyco, constrs) => (enclose (tyco ^ "(") ")" o commas o map str) constrs) prelude_import_unqualified_constr))
   374           @ map (fn (tyco, constrs) => (enclose (tyco ^ "(") ")" o commas o map str) constrs) prelude_import_unqualified_constr))
   369       :: print_qualified_import "Prelude"
   375       :: print_qualified_import "Prelude"
   370       :: map (print_qualified_import o fst) includes;
   376       :: map (print_qualified_import o fst) includes;
   371     fun print_module module_name (gr, imports) =
   377     fun print_module module_name (gr, imports) =
   372       let
   378       let
   373         val deresolve = deresolver module_name;
   379         val deresolve = deresolver module_name;
   374         fun print_import module_name = (semicolon o map str) ["import qualified", module_name];
   380         val deresolve_import = SOME o str o deresolve;
       
   381         val deresolve_import_attached = SOME o str o suffix "(..)" o deresolve;
       
   382         fun print_import (sym, Code_Thingol.Fun _) = deresolve_import sym
       
   383           | print_import (sym, Code_Thingol.Datatype _) = deresolve_import_attached sym
       
   384           | print_import (sym, Code_Thingol.Class _) = deresolve_import_attached sym
       
   385           | print_import (sym, Code_Thingol.Classinst _) = NONE;
   375         val import_ps = import_common_ps @ map (print_qualified_import o fst) imports;
   386         val import_ps = import_common_ps @ map (print_qualified_import o fst) imports;
   376         fun print_stmt' sym = case Code_Symbol.Graph.get_node gr sym
   387         fun print_stmt' sym = case Code_Symbol.Graph.get_node gr sym
   377          of (_, NONE) => NONE
   388          of (_, NONE) => NONE
   378           | (_, SOME stmt) => SOME (markup_stmt sym (print_stmt deresolve (sym, stmt)));
   389           | (_, SOME (export, stmt)) =>
   379         val body_ps = map_filter print_stmt' ((flat o rev o Code_Symbol.Graph.strong_conn) gr);
   390               SOME (if export then print_import (sym, stmt) else NONE, markup_stmt sym (print_stmt deresolve (sym, stmt)));
       
   391         val (export_ps, body_ps) = (flat o rev o Code_Symbol.Graph.strong_conn) gr
       
   392           |> map_filter print_stmt'
       
   393           |> split_list
       
   394           |> apfst (map_filter I);
   380       in
   395       in
   381         print_module_frame module_name
   396         print_module_frame module_name (SOME export_ps)
   382           ((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)
   397           ((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)
   383       end;
   398       end;
   384 
   399 
   385     (*serialization*)
   400     (*serialization*)
   386     fun write_module width (SOME destination) (module_name, content) =
   401     fun write_module width (SOME destination) (module_name, content) =
   397   in
   412   in
   398     Code_Target.serialization
   413     Code_Target.serialization
   399       (fn width => fn destination => K () o map (write_module width destination))
   414       (fn width => fn destination => K () o map (write_module width destination))
   400       (fn present => fn width => rpair (try (deresolver ""))
   415       (fn present => fn width => rpair (try (deresolver ""))
   401         o (map o apsnd) (format present width))
   416         o (map o apsnd) (format present width))
   402       (map (uncurry print_module_frame o apsnd single) includes
   417       (map (fn (module_name, content) => print_module_frame module_name NONE [content]) includes
   403         @ map (fn module_name => print_module module_name (Graph.get_node haskell_program module_name))
   418         @ map (fn module_name => print_module module_name (Graph.get_node haskell_program module_name))
   404           ((flat o rev o Graph.strong_conn) haskell_program))
   419           ((flat o rev o Graph.strong_conn) haskell_program))
   405   end;
   420   end;
   406 
   421 
   407 val serializer : Code_Target.serializer =
   422 val serializer : Code_Target.serializer =