src/Tools/Code/code_haskell.ML
changeset 82380 ceb4f33d3073
parent 81712 97987036f051
child 82447 741f6f6df144
equal deleted inserted replaced
82379:3f875966c3e1 82380:ceb4f33d3073
    49     fun class_name class = case class_syntax class
    49     fun class_name class = case class_syntax class
    50      of NONE => deresolve_class class
    50      of NONE => deresolve_class class
    51       | SOME class => class;
    51       | SOME class => class;
    52     fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
    52     fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
    53      of [] => []
    53      of [] => []
    54       | constraints => enum "," "(" ")" (
    54       | constraints => Pretty.enum "," "(" ")" (
    55           map (fn (v, class) =>
    55           map (fn (v, class) =>
    56             str (class_name class ^ " " ^ lookup_var tyvars v)) constraints)
    56             Pretty.str (class_name class ^ " " ^ lookup_var tyvars v)) constraints)
    57           @@ str " => ";
    57           @@ Pretty.str " => ";
    58     fun print_typforall tyvars vs = case map fst vs
    58     fun print_typforall tyvars vs = case map fst vs
    59      of [] => []
    59      of [] => []
    60       | vnames => str "forall " :: Pretty.breaks
    60       | vnames => Pretty.str "forall " :: Pretty.breaks
    61           (map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
    61           (map (Pretty.str o lookup_var tyvars) vnames) @ Pretty.str "." @@ Pretty.brk 1;
    62     fun print_tyco_expr tyvars fxy (tyco, tys) =
    62     fun print_tyco_expr tyvars fxy (tyco, tys) =
    63       brackify fxy (str tyco :: map (print_typ tyvars BR) tys)
    63       brackify fxy (Pretty.str tyco :: map (print_typ tyvars BR) tys)
    64     and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco
    64     and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco
    65          of NONE => print_tyco_expr tyvars fxy (deresolve_tyco tyco, tys)
    65          of NONE => print_tyco_expr tyvars fxy (deresolve_tyco tyco, tys)
    66           | SOME (_, print) => print (print_typ tyvars) fxy tys)
    66           | SOME (_, print) => print (print_typ tyvars) fxy tys)
    67       | print_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v;
    67       | print_typ tyvars fxy (ITyVar v) = (Pretty.str o lookup_var tyvars) v;
    68     fun print_typdecl tyvars (tyco, vs) =
    68     fun print_typdecl tyvars (tyco, vs) =
    69       print_tyco_expr tyvars NOBR (tyco, map ITyVar vs);
    69       print_tyco_expr tyvars NOBR (tyco, map ITyVar vs);
    70     fun print_typscheme tyvars (vs, ty) =
    70     fun print_typscheme tyvars (vs, ty) =
    71       Pretty.block (print_typforall tyvars vs @ print_typcontext tyvars vs @| print_typ tyvars NOBR ty);
    71       Pretty.block (print_typforall tyvars vs @ print_typcontext tyvars vs @| print_typ tyvars NOBR ty);
    72     fun print_term tyvars some_thm vars fxy (IConst const) =
    72     fun print_term tyvars some_thm vars fxy (IConst const) =
    78                 brackify fxy [
    78                 brackify fxy [
    79                   print_term tyvars some_thm vars NOBR t1,
    79                   print_term tyvars some_thm vars NOBR t1,
    80                   print_term tyvars some_thm vars BR t2
    80                   print_term tyvars some_thm vars BR t2
    81                 ])
    81                 ])
    82       | print_term tyvars some_thm vars fxy (IVar NONE) =
    82       | print_term tyvars some_thm vars fxy (IVar NONE) =
    83           str "_"
    83           Pretty.str "_"
    84       | print_term tyvars some_thm vars fxy (IVar (SOME v)) =
    84       | print_term tyvars some_thm vars fxy (IVar (SOME v)) =
    85           (str o lookup_var vars) v
    85           (Pretty.str o lookup_var vars) v
    86       | print_term tyvars some_thm vars fxy (t as _ `|=> _) =
    86       | print_term tyvars some_thm vars fxy (t as _ `|=> _) =
    87           let
    87           let
    88             val (binds, t') = Code_Thingol.unfold_pat_abs t;
    88             val (binds, t') = Code_Thingol.unfold_pat_abs t;
    89             val (ps, vars') = fold_map (print_bind tyvars some_thm BR o fst) binds vars;
    89             val (ps, vars') = fold_map (print_bind tyvars some_thm BR o fst) binds vars;
    90           in brackets (str "\\" :: ps @ str "->" @@ print_term tyvars some_thm vars' NOBR t') end
    90           in brackets (Pretty.str "\\" :: ps @ Pretty.str "->" @@ print_term tyvars some_thm vars' NOBR t') end
    91       | print_term tyvars some_thm vars fxy (ICase case_expr) =
    91       | print_term tyvars some_thm vars fxy (ICase case_expr) =
    92           (case Code_Thingol.unfold_const_app (#primitive case_expr)
    92           (case Code_Thingol.unfold_const_app (#primitive case_expr)
    93            of SOME (app as ({ sym = Constant const, ... }, _)) =>
    93            of SOME (app as ({ sym = Constant const, ... }, _)) =>
    94                 if is_none (const_syntax const)
    94                 if is_none (const_syntax const)
    95                 then print_case tyvars some_thm vars fxy case_expr
    95                 then print_case tyvars some_thm vars fxy case_expr
    97             | NONE => print_case tyvars some_thm vars fxy case_expr)
    97             | NONE => print_case tyvars some_thm vars fxy case_expr)
    98     and print_app_expr tyvars some_thm vars ({ sym, annotation, ... }, ts) =
    98     and print_app_expr tyvars some_thm vars ({ sym, annotation, ... }, ts) =
    99       let
    99       let
   100         val printed_const =
   100         val printed_const =
   101           case annotation of
   101           case annotation of
   102             SOME ty => brackets [(str o deresolve) sym, str "::", print_typ tyvars NOBR ty]
   102             SOME ty => brackets [(Pretty.str o deresolve) sym, Pretty.str "::", print_typ tyvars NOBR ty]
   103           | NONE => (str o deresolve) sym
   103           | NONE => (Pretty.str o deresolve) sym
   104       in 
   104       in 
   105         printed_const :: map (print_term tyvars some_thm vars BR) ts
   105         printed_const :: map (print_term tyvars some_thm vars BR) ts
   106       end
   106       end
   107     and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) const_syntax
   107     and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) const_syntax
   108     and print_bind tyvars some_thm fxy p = gen_print_bind (print_term tyvars) some_thm fxy p
   108     and print_bind tyvars some_thm fxy p = gen_print_bind (print_term tyvars) some_thm fxy p
   109     and print_case tyvars some_thm vars fxy { clauses = [], ... } =
   109     and print_case tyvars some_thm vars fxy { clauses = [], ... } =
   110           (brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""]
   110           (brackify fxy o Pretty.breaks o map Pretty.str) ["error", "\"empty case\""]
   111       | print_case tyvars some_thm vars fxy (case_expr as { clauses = [(IVar _, _)], ... }) =
   111       | print_case tyvars some_thm vars fxy (case_expr as { clauses = [(IVar _, _)], ... }) =
   112           let
   112           let
   113             val (vs, body) = Code_Thingol.unfold_let_no_pat (ICase case_expr);
   113             val (vs, body) = Code_Thingol.unfold_let_no_pat (ICase case_expr);
   114             fun print_assignment ((some_v, _), t) vars =
   114             fun print_assignment ((some_v, _), t) vars =
   115               vars
   115               vars
   116               |> print_bind tyvars some_thm BR (IVar some_v)
   116               |> print_bind tyvars some_thm BR (IVar some_v)
   117               |>> (fn p => semicolon [p, str "=", print_term tyvars some_thm vars NOBR t])
   117               |>> (fn p => semicolon [p, Pretty.str "=", print_term tyvars some_thm vars NOBR t])
   118             val (ps, vars') = fold_map print_assignment vs vars;
   118             val (ps, vars') = fold_map print_assignment vs vars;
   119           in brackify_block fxy (str "let {")
   119           in brackify_block fxy (Pretty.str "let {")
   120             ps
   120             ps
   121             (concat [str "}", str "in", print_term tyvars some_thm vars' NOBR body])
   121             (concat [Pretty.str "}", Pretty.str "in", print_term tyvars some_thm vars' NOBR body])
   122           end
   122           end
   123       | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
   123       | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
   124           let
   124           let
   125             fun print_select (pat, body) =
   125             fun print_select (pat, body) =
   126               let
   126               let
   127                 val (p, vars') = print_bind tyvars some_thm NOBR pat vars;
   127                 val (p, vars') = print_bind tyvars some_thm NOBR pat vars;
   128               in semicolon [p, str "->", print_term tyvars some_thm vars' NOBR body] end;
   128               in semicolon [p, Pretty.str "->", print_term tyvars some_thm vars' NOBR body] end;
   129           in Pretty.block_enclose
   129           in Pretty.block_enclose
   130             (concat [str "(case", print_term tyvars some_thm vars NOBR t, str "of", str "{"], str "})")
   130             (concat [Pretty.str "(case", print_term tyvars some_thm vars NOBR t, Pretty.str "of", Pretty.str "{"], Pretty.str "})")
   131             (map print_select clauses)
   131             (map print_select clauses)
   132           end;
   132           end;
   133     fun print_stmt (Constant const, Code_Thingol.Fun (((vs, ty), raw_eqs), _)) =
   133     fun print_stmt (Constant const, Code_Thingol.Fun (((vs, ty), raw_eqs), _)) =
   134           let
   134           let
   135             val tyvars = intro_vars (map fst vs) reserved;
   135             val tyvars = intro_vars (map fst vs) reserved;
   136             fun print_err n =
   136             fun print_err n =
   137               (semicolon o map str) (
   137               (semicolon o map Pretty.str) (
   138                 deresolve_const const
   138                 deresolve_const const
   139                 :: replicate n "_"
   139                 :: replicate n "_"
   140                 @ "="
   140                 @ "="
   141                 :: "error"
   141                 :: "error"
   142                 @@ print_haskell_string const
   142                 @@ print_haskell_string const
   147                   |> intro_base_names_for (is_none o const_syntax)
   147                   |> intro_base_names_for (is_none o const_syntax)
   148                       deresolve (t :: ts)
   148                       deresolve (t :: ts)
   149                   |> intro_vars (build (fold Code_Thingol.add_varnames ts));
   149                   |> intro_vars (build (fold Code_Thingol.add_varnames ts));
   150               in
   150               in
   151                 semicolon (
   151                 semicolon (
   152                   (str o deresolve_const) const
   152                   (Pretty.str o deresolve_const) const
   153                   :: map (print_term tyvars some_thm vars BR) ts
   153                   :: map (print_term tyvars some_thm vars BR) ts
   154                   @ str "="
   154                   @ Pretty.str "="
   155                   @@ print_term tyvars some_thm vars NOBR t
   155                   @@ print_term tyvars some_thm vars NOBR t
   156                 )
   156                 )
   157               end;
   157               end;
   158           in
   158           in
   159             Pretty.chunks (
   159             Pretty.chunks (
   160               semicolon [
   160               semicolon [
   161                 (str o suffix " ::" o deresolve_const) const,
   161                 (Pretty.str o suffix " ::" o deresolve_const) const,
   162                 print_typscheme tyvars (vs, ty)
   162                 print_typscheme tyvars (vs, ty)
   163               ]
   163               ]
   164               :: (case filter (snd o snd) raw_eqs
   164               :: (case filter (snd o snd) raw_eqs
   165                of [] => [print_err ((length o fst o Code_Thingol.unfold_fun) ty)]
   165                of [] => [print_err ((length o fst o Code_Thingol.unfold_fun) ty)]
   166                 | eqs => map print_eqn eqs)
   166                 | eqs => map print_eqn eqs)
   169       | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [])) =
   169       | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [])) =
   170           let
   170           let
   171             val tyvars = intro_vars vs reserved;
   171             val tyvars = intro_vars vs reserved;
   172           in
   172           in
   173             semicolon [
   173             semicolon [
   174               str "data",
   174               Pretty.str "data",
   175               print_typdecl tyvars (deresolve_tyco tyco, vs)
   175               print_typdecl tyvars (deresolve_tyco tyco, vs)
   176             ]
   176             ]
   177           end
   177           end
   178       | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [((co, _), [ty])])) =
   178       | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [((co, _), [ty])])) =
   179           let
   179           let
   180             val tyvars = intro_vars vs reserved;
   180             val tyvars = intro_vars vs reserved;
   181           in
   181           in
   182             semicolon (
   182             semicolon (
   183               str "newtype"
   183               Pretty.str "newtype"
   184               :: print_typdecl tyvars (deresolve_tyco tyco, vs)
   184               :: print_typdecl tyvars (deresolve_tyco tyco, vs)
   185               :: str "="
   185               :: Pretty.str "="
   186               :: (str o deresolve_const) co
   186               :: (Pretty.str o deresolve_const) co
   187               :: print_typ tyvars BR ty
   187               :: print_typ tyvars BR ty
   188               :: (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else [])
   188               :: (if deriving_show tyco then [Pretty.str "deriving (Prelude.Read, Prelude.Show)"] else [])
   189             )
   189             )
   190           end
   190           end
   191       | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, co :: cos)) =
   191       | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, co :: cos)) =
   192           let
   192           let
   193             val tyvars = intro_vars vs reserved;
   193             val tyvars = intro_vars vs reserved;
   194             fun print_co ((co, _), tys) =
   194             fun print_co ((co, _), tys) =
   195               concat (
   195               concat (
   196                 (str o deresolve_const) co
   196                 (Pretty.str o deresolve_const) co
   197                 :: map (print_typ tyvars BR) tys
   197                 :: map (print_typ tyvars BR) tys
   198               )
   198               )
   199           in
   199           in
   200             semicolon (
   200             semicolon (
   201               str "data"
   201               Pretty.str "data"
   202               :: print_typdecl tyvars (deresolve_tyco tyco, vs)
   202               :: print_typdecl tyvars (deresolve_tyco tyco, vs)
   203               :: str "="
   203               :: Pretty.str "="
   204               :: print_co co
   204               :: print_co co
   205               :: map ((fn p => Pretty.block [str "| ", p]) o print_co) cos
   205               :: map ((fn p => Pretty.block [Pretty.str "| ", p]) o print_co) cos
   206               @ (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else [])
   206               @ (if deriving_show tyco then [Pretty.str "deriving (Prelude.Read, Prelude.Show)"] else [])
   207             )
   207             )
   208           end
   208           end
   209       | print_stmt (Type_Class class, Code_Thingol.Class (v, (classrels, classparams))) =
   209       | print_stmt (Type_Class class, Code_Thingol.Class (v, (classrels, classparams))) =
   210           let
   210           let
   211             val tyvars = intro_vars [v] reserved;
   211             val tyvars = intro_vars [v] reserved;
   212             fun print_classparam (classparam, ty) =
   212             fun print_classparam (classparam, ty) =
   213               semicolon [
   213               semicolon [
   214                 (str o deresolve_const) classparam,
   214                 (Pretty.str o deresolve_const) classparam,
   215                 str "::",
   215                 Pretty.str "::",
   216                 print_typ tyvars NOBR ty
   216                 print_typ tyvars NOBR ty
   217               ]
   217               ]
   218           in
   218           in
   219             Pretty.block_enclose (
   219             Pretty.block_enclose (
   220               Pretty.block [
   220               Pretty.block [
   221                 str "class ",
   221                 Pretty.str "class ",
   222                 Pretty.block (print_typcontext tyvars [(v, map snd classrels)]),
   222                 Pretty.block (print_typcontext tyvars [(v, map snd classrels)]),
   223                 str (deresolve_class class ^ " " ^ lookup_var tyvars v),
   223                 Pretty.str (deresolve_class class ^ " " ^ lookup_var tyvars v),
   224                 str " where {"
   224                 Pretty.str " where {"
   225               ],
   225               ],
   226               str "};"
   226               Pretty.str "};"
   227             ) (map print_classparam classparams)
   227             ) (map print_classparam classparams)
   228           end
   228           end
   229       | print_stmt (_, Code_Thingol.Classinst { class, tyco, vs, inst_params, ... }) =
   229       | print_stmt (_, Code_Thingol.Classinst { class, tyco, vs, inst_params, ... }) =
   230           let
   230           let
   231             val tyvars = intro_vars (map fst vs) reserved;
   231             val tyvars = intro_vars (map fst vs) reserved;
   232             fun print_classparam_instance ((classparam, (const, _)), (thm, _)) =
   232             fun print_classparam_instance ((classparam, (const, _)), (thm, _)) =
   233               case const_syntax classparam of
   233               case const_syntax classparam of
   234                 NONE => semicolon [
   234                 NONE => semicolon [
   235                     (str o Long_Name.base_name o deresolve_const) classparam,
   235                     (Pretty.str o Long_Name.base_name o deresolve_const) classparam,
   236                     str "=",
   236                     Pretty.str "=",
   237                     print_app tyvars (SOME thm) reserved NOBR (const, [])
   237                     print_app tyvars (SOME thm) reserved NOBR (const, [])
   238                   ]
   238                   ]
   239               | SOME (_, Code_Printer.Plain_printer s) => semicolon [
   239               | SOME (_, Code_Printer.Plain_printer s) => semicolon [
   240                     (str o Long_Name.base_name) s,
   240                     (Pretty.str o Long_Name.base_name) s,
   241                     str "=",
   241                     Pretty.str "=",
   242                     print_app tyvars (SOME thm) reserved NOBR (const, [])
   242                     print_app tyvars (SOME thm) reserved NOBR (const, [])
   243                   ]
   243                   ]
   244               | SOME (wanted, Code_Printer.Complex_printer _) =>
   244               | SOME (wanted, Code_Printer.Complex_printer _) =>
   245                   let
   245                   let
   246                     val { sym = Constant c, dom, range, ... } = const;
   246                     val { sym = Constant c, dom, range, ... } = const;
   256                       (*dictionaries are not relevant in Haskell,
   256                       (*dictionaries are not relevant in Haskell,
   257                         and these consts never need type annotations for disambiguation *)
   257                         and these consts never need type annotations for disambiguation *)
   258                   in
   258                   in
   259                     semicolon [
   259                     semicolon [
   260                       print_term tyvars (SOME thm) vars NOBR lhs,
   260                       print_term tyvars (SOME thm) vars NOBR lhs,
   261                       str "=",
   261                       Pretty.str "=",
   262                       print_term tyvars (SOME thm) vars NOBR rhs
   262                       print_term tyvars (SOME thm) vars NOBR rhs
   263                     ]
   263                     ]
   264                   end;
   264                   end;
   265           in
   265           in
   266             Pretty.block_enclose (
   266             Pretty.block_enclose (
   267               Pretty.block [
   267               Pretty.block [
   268                 str "instance ",
   268                 Pretty.str "instance ",
   269                 Pretty.block (print_typcontext tyvars vs),
   269                 Pretty.block (print_typcontext tyvars vs),
   270                 str (class_name class ^ " "),
   270                 Pretty.str (class_name class ^ " "),
   271                 print_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
   271                 print_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
   272                 str " where {"
   272                 Pretty.str " where {"
   273               ],
   273               ],
   274               str "};"
   274               Pretty.str "};"
   275             ) (map print_classparam_instance inst_params)
   275             ) (map print_classparam_instance inst_params)
   276           end;
   276           end;
   277   in print_stmt end;
   277   in print_stmt end;
   278 
   278 
   279 fun haskell_program_of_program ctxt module_prefix module_name reserved identifiers =
   279 fun haskell_program_of_program ctxt module_prefix module_name reserved identifiers =
   366         val (xs, x) = split_last (Long_Name.explode module_name)
   366         val (xs, x) = split_last (Long_Name.explode module_name)
   367       in xs @ [x ^ ".hs"] end
   367       in xs @ [x ^ ".hs"] end
   368     fun print_module_frame module_name header exports ps =
   368     fun print_module_frame module_name header exports ps =
   369       (module_names module_name, Pretty.chunks2 (
   369       (module_names module_name, Pretty.chunks2 (
   370         header
   370         header
   371         @ concat [str "module",
   371         @ concat [Pretty.str "module",
   372           case exports of
   372           case exports of
   373             SOME ps => Pretty.block [str module_name, enclose "(" ")" (commas ps)]
   373             SOME ps => Pretty.block [Pretty.str module_name, Pretty.enclose "(" ")" (Pretty.commas ps)]
   374           | NONE => str module_name,
   374           | NONE => Pretty.str module_name,
   375           str "where {"
   375           Pretty.str "where {"
   376         ]
   376         ]
   377         :: ps
   377         :: ps
   378         @| str "}"
   378         @| Pretty.str "}"
   379       ));
   379       ));
   380     fun print_qualified_import module_name =
   380     fun print_qualified_import module_name =
   381       semicolon [str "import qualified", str module_name];
   381       semicolon [Pretty.str "import qualified", Pretty.str module_name];
   382     val import_common_ps =
   382     val import_common_ps =
   383       enclose "import Prelude (" ");" (commas (map str
   383       Pretty.enclose "import Prelude (" ");" (Pretty.commas (map Pretty.str
   384         (map (Library.enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
   384         (map (enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
   385           @ map (fn (tyco, constrs) => (enclose (tyco ^ "(") ")" o commas o map str) constrs) prelude_import_unqualified_constr))
   385           @ map (fn (tyco, constrs) => (Pretty.enclose (tyco ^ "(") ")" o Pretty.commas o map Pretty.str) constrs) prelude_import_unqualified_constr))
   386       :: enclose "import Data.Bits (" ");" (commas
   386       :: Pretty.enclose "import Data.Bits (" ");" (Pretty.commas
   387         (map (str o Library.enclose "(" ")") data_bits_import_operators))
   387         (map (Pretty.str o enclose "(" ")") data_bits_import_operators))
   388       :: print_qualified_import "Prelude"
   388       :: print_qualified_import "Prelude"
   389       :: print_qualified_import "Data.Bits"
   389       :: print_qualified_import "Data.Bits"
   390       :: map (print_qualified_import o fst) includes;
   390       :: map (print_qualified_import o fst) includes;
   391     fun print_module module_name (gr, imports) =
   391     fun print_module module_name (gr, imports) =
   392       let
   392       let
   393         val deresolve = deresolver module_name;
   393         val deresolve = deresolver module_name;
   394         val deresolve_import = SOME o str o deresolve;
   394         val deresolve_import = SOME o Pretty.str o deresolve;
   395         val deresolve_import_attached = SOME o str o suffix "(..)" o deresolve;
   395         val deresolve_import_attached = SOME o Pretty.str o suffix "(..)" o deresolve;
   396         fun print_import (sym, (_, Code_Thingol.Fun _)) = deresolve_import sym
   396         fun print_import (sym, (_, Code_Thingol.Fun _)) = deresolve_import sym
   397           | print_import (sym, (Code_Namespace.Public, Code_Thingol.Datatype _)) = deresolve_import_attached sym
   397           | print_import (sym, (Code_Namespace.Public, Code_Thingol.Datatype _)) = deresolve_import_attached sym
   398           | print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Datatype _)) = deresolve_import sym
   398           | print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Datatype _)) = deresolve_import sym
   399           | print_import (sym, (Code_Namespace.Public, Code_Thingol.Class _)) = deresolve_import_attached sym
   399           | print_import (sym, (Code_Namespace.Public, Code_Thingol.Class _)) = deresolve_import_attached sym
   400           | print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Class _)) = deresolve_import sym
   400           | print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Class _)) = deresolve_import sym
   407         val (export_ps, body_ps) = (flat o rev o Code_Symbol.Graph.strong_conn) gr
   407         val (export_ps, body_ps) = (flat o rev o Code_Symbol.Graph.strong_conn) gr
   408           |> map_filter print_stmt'
   408           |> map_filter print_stmt'
   409           |> split_list
   409           |> split_list
   410           |> apfst (map_filter I);
   410           |> apfst (map_filter I);
   411       in
   411       in
   412         print_module_frame module_name [str language_pragma] (SOME export_ps)
   412         print_module_frame module_name [Pretty.str language_pragma] (SOME export_ps)
   413           ((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)
   413           ((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)
   414       end;
   414       end;
   415 
   415 
   416   in
   416   in
   417     (Code_Target.Hierarchy (map (fn (module_name, content) => ([module_name ^ ".hs"], content)) includes
   417     (Code_Target.Hierarchy (map (fn (module_name, content) => ([module_name ^ ".hs"], content)) includes
   428 fun print_numeral typ = Library.enclose "(" (" :: " ^ typ ^ ")") o signed_string_of_int;
   428 fun print_numeral typ = Library.enclose "(" (" :: " ^ typ ^ ")") o signed_string_of_int;
   429 
   429 
   430 val literals = Literals {
   430 val literals = Literals {
   431   literal_string = print_haskell_string,
   431   literal_string = print_haskell_string,
   432   literal_numeral = print_numeral "Integer",
   432   literal_numeral = print_numeral "Integer",
   433   literal_list = enum "," "[" "]",
   433   literal_list = Pretty.enum "," "[" "]",
   434   infix_cons = (5, ":")
   434   infix_cons = (5, ":")
   435 };
   435 };
   436 
   436 
   437 
   437 
   438 (** optional monad syntax **)
   438 (** optional monad syntax **)
   453     val implode_monad = Code_Thingol.unfoldr dest_monad;
   453     val implode_monad = Code_Thingol.unfoldr dest_monad;
   454     fun print_monad print_bind print_term (NONE, t) vars =
   454     fun print_monad print_bind print_term (NONE, t) vars =
   455           (semicolon [print_term vars NOBR t], vars)
   455           (semicolon [print_term vars NOBR t], vars)
   456       | print_monad print_bind print_term (SOME ((bind, _), true), t) vars = vars
   456       | print_monad print_bind print_term (SOME ((bind, _), true), t) vars = vars
   457           |> print_bind NOBR bind
   457           |> print_bind NOBR bind
   458           |>> (fn p => semicolon [p, str "<-", print_term vars NOBR t])
   458           |>> (fn p => semicolon [p, Pretty.str "<-", print_term vars NOBR t])
   459       | print_monad print_bind print_term (SOME ((bind, _), false), t) vars = vars
   459       | print_monad print_bind print_term (SOME ((bind, _), false), t) vars = vars
   460           |> print_bind NOBR bind
   460           |> print_bind NOBR bind
   461           |>> (fn p => semicolon [str "let", str "{", p, str "=", print_term vars NOBR t, str "}"]);
   461           |>> (fn p => semicolon [Pretty.str "let", Pretty.str "{", p, Pretty.str "=", print_term vars NOBR t, Pretty.str "}"]);
   462     fun pretty _ print_term thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2
   462     fun pretty _ print_term thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2
   463      of SOME (bind, t') => let
   463      of SOME (bind, t') => let
   464           val (binds, t'') = implode_monad t'
   464           val (binds, t'') = implode_monad t'
   465           val (ps, vars') = fold_map (print_monad (gen_print_bind (K print_term) thm) print_term)
   465           val (ps, vars') = fold_map (print_monad (gen_print_bind (K print_term) thm) print_term)
   466             (bind :: binds) vars;
   466             (bind :: binds) vars;
   467         in
   467         in
   468           brackify_block fxy (str "do { ")
   468           brackify_block fxy (Pretty.str "do { ")
   469             (ps @| print_term vars' NOBR t'')
   469             (ps @| print_term vars' NOBR t'')
   470             (str " }")
   470             (Pretty.str " }")
   471         end
   471         end
   472       | NONE => brackify_infix (1, L) fxy
   472       | NONE => brackify_infix (1, L) fxy
   473           (print_term vars (INFX (1, L)) t1, str ">>=", print_term vars (INFX (1, X)) t2)
   473           (print_term vars (INFX (1, L)) t1, Pretty.str ">>=", print_term vars (INFX (1, X)) t2)
   474   in (2, pretty) end;
   474   in (2, pretty) end;
   475 
   475 
   476 fun add_monad target' raw_c_bind thy =
   476 fun add_monad target' raw_c_bind thy =
   477   let
   477   let
   478     val c_bind = Code.read_const thy raw_c_bind;
   478     val c_bind = Code.read_const thy raw_c_bind;
   497       evaluation_args = []})
   497       evaluation_args = []})
   498   #> Code_Target.set_printings (Type_Constructor ("fun",
   498   #> Code_Target.set_printings (Type_Constructor ("fun",
   499     [(target, SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
   499     [(target, SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
   500       brackify_infix (1, R) fxy (
   500       brackify_infix (1, R) fxy (
   501         print_typ (INFX (1, X)) ty1,
   501         print_typ (INFX (1, X)) ty1,
   502         str "->",
   502         Pretty.str "->",
   503         print_typ (INFX (1, R)) ty2
   503         print_typ (INFX (1, R)) ty2
   504       )))]))
   504       )))]))
   505   #> fold (Code_Target.add_reserved target) [
   505   #> fold (Code_Target.add_reserved target) [
   506       "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
   506       "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
   507       "import", "default", "forall", "let", "in", "class", "qualified", "data",
   507       "import", "default", "forall", "let", "in", "class", "qualified", "data",