src/Tools/Code/code_haskell.ML
author wenzelm
Tue, 01 Apr 2014 22:25:01 +0200
changeset 56355 1a9f569b5b7e
parent 55683 5732a55b9232
child 56812 baef1c110f12
permissions -rw-r--r--
some rephrasing to ensure that this becomes cheap "foreach" and not expensive "map" (cf. 0fc032898b05);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37745
6315b6426200 checking generated code for various target languages
haftmann
parents: 37669
diff changeset
     1
(*  Title:      Tools/Code/code_haskell.ML
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     3
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     4
Serializer for Haskell.
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     5
*)
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     6
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     7
signature CODE_HASKELL =
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
     8
sig
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
     9
  val language_params: string
37745
6315b6426200 checking generated code for various target languages
haftmann
parents: 37669
diff changeset
    10
  val target: string
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    11
  val setup: theory -> theory
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    12
end;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    13
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    14
structure Code_Haskell : CODE_HASKELL =
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    15
struct
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    16
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    17
val target = "Haskell";
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    18
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
    19
val language_extensions =
44940
56fd289398a2 tune indenting
noschinl
parents: 44926
diff changeset
    20
  ["EmptyDataDecls", "RankNTypes", "ScopedTypeVariables"];
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
    21
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
    22
val language_pragma =
44940
56fd289398a2 tune indenting
noschinl
parents: 44926
diff changeset
    23
  "{-# LANGUAGE " ^ commas language_extensions ^ " #-}";
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
    24
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
    25
val language_params =
44940
56fd289398a2 tune indenting
noschinl
parents: 44926
diff changeset
    26
  space_implode " " (map (prefix "-X") language_extensions);
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
    27
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
    28
open Basic_Code_Symbol;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    29
open Basic_Code_Thingol;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    30
open Code_Printer;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    31
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    32
infixr 5 @@;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    33
infixr 5 @|;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    34
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    35
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    36
(** Haskell serializer **)
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    37
47609
b3dab1892cda dropped dead code
haftmann
parents: 46961
diff changeset
    38
fun print_haskell_stmt class_syntax tyco_syntax const_syntax
44793
fddb09e6f84d removing previous crude approximation to add type annotations to disambiguate types
bulwahn
parents: 44792
diff changeset
    39
    reserved deresolve deriving_show =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    40
  let
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
    41
    val deresolve_const = deresolve o Constant;
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
    42
    val deresolve_tyco = deresolve o Type_Constructor;
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
    43
    val deresolve_class = deresolve o Type_Class;
38923
79d7f2b4cf71 more coherent naming of syntax data structures
haftmann
parents: 38916
diff changeset
    44
    fun class_name class = case class_syntax class
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    45
     of NONE => deresolve_class class
28687
150a8a1eae1a simplified syntax for class parameters
haftmann
parents: 28663
diff changeset
    46
      | SOME class => class;
33991
haftmann
parents: 33421
diff changeset
    47
    fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    48
     of [] => []
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 37242
diff changeset
    49
      | constraints => enum "," "(" ")" (
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    50
          map (fn (v, class) =>
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 37242
diff changeset
    51
            str (class_name class ^ " " ^ lookup_var tyvars v)) constraints)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    52
          @@ str " => ";
33991
haftmann
parents: 33421
diff changeset
    53
    fun print_typforall tyvars vs = case map fst vs
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    54
     of [] => []
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    55
      | vnames => str "forall " :: Pretty.breaks
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
    56
          (map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
33991
haftmann
parents: 33421
diff changeset
    57
    fun print_tyco_expr tyvars fxy (tyco, tys) =
haftmann
parents: 33421
diff changeset
    58
      brackify fxy (str tyco :: map (print_typ tyvars BR) tys)
47609
b3dab1892cda dropped dead code
haftmann
parents: 46961
diff changeset
    59
    and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    60
         of NONE => print_tyco_expr tyvars fxy (deresolve_tyco tyco, tys)
47609
b3dab1892cda dropped dead code
haftmann
parents: 46961
diff changeset
    61
          | SOME (_, print) => print (print_typ tyvars) fxy tys)
33991
haftmann
parents: 33421
diff changeset
    62
      | print_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v;
48003
1d11af40b106 dropped sort constraints on datatype specifications
haftmann
parents: 47609
diff changeset
    63
    fun print_typdecl tyvars (tyco, vs) =
1d11af40b106 dropped sort constraints on datatype specifications
haftmann
parents: 47609
diff changeset
    64
      print_tyco_expr tyvars NOBR (tyco, map ITyVar vs);
33991
haftmann
parents: 33421
diff changeset
    65
    fun print_typscheme tyvars (vs, ty) =
haftmann
parents: 33421
diff changeset
    66
      Pretty.block (print_typforall tyvars vs @ print_typcontext tyvars vs @| print_typ tyvars NOBR ty);
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    67
    fun print_term tyvars some_thm vars fxy (IConst const) =
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    68
          print_app tyvars some_thm vars fxy (const, [])
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    69
      | print_term tyvars some_thm vars fxy (t as (t1 `$ t2)) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    70
          (case Code_Thingol.unfold_const_app t
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    71
           of SOME app => print_app tyvars some_thm vars fxy app
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    72
            | _ =>
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    73
                brackify fxy [
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    74
                  print_term tyvars some_thm vars NOBR t1,
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    75
                  print_term tyvars some_thm vars BR t2
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    76
                ])
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    77
      | print_term tyvars some_thm vars fxy (IVar NONE) =
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
    78
          str "_"
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    79
      | print_term tyvars some_thm vars fxy (IVar (SOME v)) =
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
    80
          (str o lookup_var vars) v
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    81
      | print_term tyvars some_thm vars fxy (t as _ `|=> _) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
    82
          let
31874
f172346ba805 simplified binding concept
haftmann
parents: 31775
diff changeset
    83
            val (binds, t') = Code_Thingol.unfold_pat_abs t;
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    84
            val (ps, vars') = fold_map (print_bind tyvars some_thm BR o fst) binds vars;
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    85
          in brackets (str "\\" :: ps @ str "->" @@ print_term tyvars some_thm vars' NOBR t') end
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    86
      | print_term tyvars some_thm vars fxy (ICase case_expr) =
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    87
          (case Code_Thingol.unfold_const_app (#primitive case_expr)
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
    88
           of SOME (app as ({ sym = Constant const, ... }, _)) =>
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    89
                if is_none (const_syntax const)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    90
                then print_case tyvars some_thm vars fxy case_expr
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    91
                else print_app tyvars some_thm vars fxy app
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    92
            | NONE => print_case tyvars some_thm vars fxy case_expr)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    93
    and print_app_expr tyvars some_thm vars ({ sym, dom, range, annotate, ... }, ts) =
44792
26b19918e670 adding minimalistic implementation for printing the type annotations
bulwahn
parents: 44789
diff changeset
    94
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    95
        val ty = Library.foldr (op `->) (dom, range)
45009
99e1965f9c21 syntactic improvements and tuning names in the code generator due to Florian's code review
bulwahn
parents: 44940
diff changeset
    96
        val printed_const =
99e1965f9c21 syntactic improvements and tuning names in the code generator due to Florian's code review
bulwahn
parents: 44940
diff changeset
    97
          if annotate then
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    98
            brackets [(str o deresolve) sym, str "::", print_typ tyvars NOBR ty]
45009
99e1965f9c21 syntactic improvements and tuning names in the code generator due to Florian's code review
bulwahn
parents: 44940
diff changeset
    99
          else
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   100
            (str o deresolve) sym
44792
26b19918e670 adding minimalistic implementation for printing the type annotations
bulwahn
parents: 44789
diff changeset
   101
      in 
45009
99e1965f9c21 syntactic improvements and tuning names in the code generator due to Florian's code review
bulwahn
parents: 44940
diff changeset
   102
        printed_const :: map (print_term tyvars some_thm vars BR) ts
44792
26b19918e670 adding minimalistic implementation for printing the type annotations
bulwahn
parents: 44789
diff changeset
   103
      end
38923
79d7f2b4cf71 more coherent naming of syntax data structures
haftmann
parents: 38916
diff changeset
   104
    and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) const_syntax
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   105
    and print_bind tyvars some_thm fxy p = gen_print_bind (print_term tyvars) some_thm fxy p
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   106
    and print_case tyvars some_thm vars fxy { clauses = [], ... } =
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   107
          (brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""]
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   108
      | print_case tyvars some_thm vars fxy (case_expr as { clauses = [_], ... }) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   109
          let
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   110
            val (binds, body) = Code_Thingol.unfold_let (ICase case_expr);
47609
b3dab1892cda dropped dead code
haftmann
parents: 46961
diff changeset
   111
            fun print_match ((pat, _), t) vars =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   112
              vars
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   113
              |> print_bind tyvars some_thm BR pat
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   114
              |>> (fn p => semicolon [p, str "=", print_term tyvars some_thm vars NOBR t])
33991
haftmann
parents: 33421
diff changeset
   115
            val (ps, vars') = fold_map print_match binds vars;
31665
a1f4d3b3f6c8 tuned brackets for let expressions etc.
haftmann
parents: 31376
diff changeset
   116
          in brackify_block fxy (str "let {")
a1f4d3b3f6c8 tuned brackets for let expressions etc.
haftmann
parents: 31376
diff changeset
   117
            ps
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   118
            (concat [str "}", str "in", print_term tyvars some_thm vars' NOBR body])
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   119
          end
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   120
      | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   121
          let
33991
haftmann
parents: 33421
diff changeset
   122
            fun print_select (pat, body) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   123
              let
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   124
                val (p, vars') = print_bind tyvars some_thm NOBR pat vars;
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   125
              in semicolon [p, str "->", print_term tyvars some_thm vars' NOBR body] end;
36576
736994daaf08 enclose case expression in brackets
haftmann
parents: 36535
diff changeset
   126
          in Pretty.block_enclose
736994daaf08 enclose case expression in brackets
haftmann
parents: 36535
diff changeset
   127
            (concat [str "(case", print_term tyvars some_thm vars NOBR t, str "of", str "{"], str "})")
33991
haftmann
parents: 33421
diff changeset
   128
            (map print_select clauses)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   129
          end;
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   130
    fun print_stmt (Constant const, Code_Thingol.Fun (((vs, ty), raw_eqs), _)) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   131
          let
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
   132
            val tyvars = intro_vars (map fst vs) reserved;
34269
b5c99df2e4b1 more correct handling of empty functions
haftmann
parents: 34178
diff changeset
   133
            fun print_err n =
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   134
              (semicolon o map str) (
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   135
                deresolve_const const
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   136
                :: replicate n "_"
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   137
                @ "="
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   138
                :: "error"
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   139
                @@ ML_Syntax.print_string const
34269
b5c99df2e4b1 more correct handling of empty functions
haftmann
parents: 34178
diff changeset
   140
              );
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   141
            fun print_eqn ((ts, t), (some_thm, _)) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   142
              let
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
   143
                val vars = reserved
55145
2bb3cd36bcf7 more abstract declaration of unqualified constant names in code printing context
haftmann
parents: 52801
diff changeset
   144
                  |> intro_base_names_for (is_none o const_syntax)
2bb3cd36bcf7 more abstract declaration of unqualified constant names in code printing context
haftmann
parents: 52801
diff changeset
   145
                      deresolve (t :: ts)
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
   146
                  |> intro_vars ((fold o Code_Thingol.fold_varnames)
32925
980f5aa0d2d7 tuned whitespace
haftmann
parents: 32924
diff changeset
   147
                      (insert (op =)) ts []);
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   148
              in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   149
                semicolon (
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   150
                  (str o deresolve_const) const
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   151
                  :: map (print_term tyvars some_thm vars BR) ts
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   152
                  @ str "="
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   153
                  @@ print_term tyvars some_thm vars NOBR t
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   154
                )
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   155
              end;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   156
          in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   157
            Pretty.chunks (
33991
haftmann
parents: 33421
diff changeset
   158
              semicolon [
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   159
                (str o suffix " ::" o deresolve_const) const,
33991
haftmann
parents: 33421
diff changeset
   160
                print_typscheme tyvars (vs, ty)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   161
              ]
34269
b5c99df2e4b1 more correct handling of empty functions
haftmann
parents: 34178
diff changeset
   162
              :: (case filter (snd o snd) raw_eqs
b5c99df2e4b1 more correct handling of empty functions
haftmann
parents: 34178
diff changeset
   163
               of [] => [print_err ((length o fst o Code_Thingol.unfold_fun) ty)]
b5c99df2e4b1 more correct handling of empty functions
haftmann
parents: 34178
diff changeset
   164
                | eqs => map print_eqn eqs)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   165
            )
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   166
          end
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   167
      | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [])) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   168
          let
48003
1d11af40b106 dropped sort constraints on datatype specifications
haftmann
parents: 47609
diff changeset
   169
            val tyvars = intro_vars vs reserved;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   170
          in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   171
            semicolon [
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   172
              str "data",
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   173
              print_typdecl tyvars (deresolve_tyco tyco, vs)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   174
            ]
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   175
          end
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   176
      | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [((co, _), [ty])])) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   177
          let
48003
1d11af40b106 dropped sort constraints on datatype specifications
haftmann
parents: 47609
diff changeset
   178
            val tyvars = intro_vars vs reserved;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   179
          in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   180
            semicolon (
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   181
              str "newtype"
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   182
              :: print_typdecl tyvars (deresolve_tyco tyco, vs)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   183
              :: str "="
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   184
              :: (str o deresolve_const) co
33991
haftmann
parents: 33421
diff changeset
   185
              :: print_typ tyvars BR ty
55374
636a8523876f restoring ancient string_classes option
haftmann
parents: 55373
diff changeset
   186
              :: (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else [])
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   187
            )
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   188
          end
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   189
      | print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, co :: cos)) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   190
          let
48003
1d11af40b106 dropped sort constraints on datatype specifications
haftmann
parents: 47609
diff changeset
   191
            val tyvars = intro_vars vs reserved;
37449
034ebe92f090 more precise code
haftmann
parents: 37447
diff changeset
   192
            fun print_co ((co, _), tys) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   193
              concat (
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   194
                (str o deresolve_const) co
33991
haftmann
parents: 33421
diff changeset
   195
                :: map (print_typ tyvars BR) tys
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   196
              )
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   197
          in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   198
            semicolon (
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   199
              str "data"
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   200
              :: print_typdecl tyvars (deresolve_tyco tyco, vs)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   201
              :: str "="
33991
haftmann
parents: 33421
diff changeset
   202
              :: print_co co
haftmann
parents: 33421
diff changeset
   203
              :: map ((fn p => Pretty.block [str "| ", p]) o print_co) cos
55374
636a8523876f restoring ancient string_classes option
haftmann
parents: 55373
diff changeset
   204
              @ (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else [])
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   205
            )
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   206
          end
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   207
      | print_stmt (Type_Class class, Code_Thingol.Class (v, (classrels, classparams))) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   208
          let
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
   209
            val tyvars = intro_vars [v] reserved;
33991
haftmann
parents: 33421
diff changeset
   210
            fun print_classparam (classparam, ty) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   211
              semicolon [
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   212
                (str o deresolve_const) classparam,
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   213
                str "::",
33991
haftmann
parents: 33421
diff changeset
   214
                print_typ tyvars NOBR ty
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   215
              ]
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   216
          in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   217
            Pretty.block_enclose (
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   218
              Pretty.block [
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   219
                str "class ",
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   220
                Pretty.block (print_typcontext tyvars [(v, map snd classrels)]),
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   221
                str (deresolve_class class ^ " " ^ lookup_var tyvars v),
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   222
                str " where {"
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   223
              ],
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   224
              str "};"
33991
haftmann
parents: 33421
diff changeset
   225
            ) (map print_classparam classparams)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   226
          end
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   227
      | print_stmt (_, Code_Thingol.Classinst { class, tyco, vs, inst_params, ... }) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   228
          let
32924
d2e9b2dab760 dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents: 32913
diff changeset
   229
            val tyvars = intro_vars (map fst vs) reserved;
52519
598addf65209 explicit hint for domain of class parameters in instance statements
haftmann
parents: 52435
diff changeset
   230
            fun print_classparam_instance ((classparam, (const, _)), (thm, _)) =
55373
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   231
              case const_syntax classparam of
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   232
                NONE => semicolon [
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   233
                    (str o Long_Name.base_name o deresolve_const) classparam,
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   234
                    str "=",
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   235
                    print_app tyvars (SOME thm) reserved NOBR (const, [])
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   236
                  ]
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   237
              | SOME (_, Code_Printer.Plain_printer s) => semicolon [
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   238
                    (str o Long_Name.base_name) s,
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   239
                    str "=",
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   240
                    print_app tyvars (SOME thm) reserved NOBR (const, [])
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   241
                  ]
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   242
              | SOME (k, Code_Printer.Complex_printer _) =>
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   243
                  let
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   244
                    val { sym = Constant c, dom, range, ... } = const;
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   245
                    val (vs, rhs) = (apfst o map) fst
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   246
                      (Code_Thingol.unfold_abs (Code_Thingol.eta_expand k (const, [])));
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   247
                    val s = if (is_some o const_syntax) c
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   248
                      then NONE else (SOME o Long_Name.base_name o deresolve_const) c;
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   249
                    val vars = reserved
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   250
                      |> intro_vars (map_filter I (s :: vs));
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   251
                    val lhs = IConst { sym = Constant classparam, typargs = [],
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   252
                      dicts = [], dom = dom, range = range, annotate = false } `$$ map IVar vs;
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   253
                      (*dictionaries are not relevant at this late stage,
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   254
                        and these consts never need type annotations for disambiguation *)
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   255
                  in
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   256
                    semicolon [
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   257
                      print_term tyvars (SOME thm) vars NOBR lhs,
28687
150a8a1eae1a simplified syntax for class parameters
haftmann
parents: 28663
diff changeset
   258
                      str "=",
55373
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   259
                      print_term tyvars (SOME thm) vars NOBR rhs
28687
150a8a1eae1a simplified syntax for class parameters
haftmann
parents: 28663
diff changeset
   260
                    ]
55373
2b4204cb7904 method names in instance declarations are always unqualified
haftmann
parents: 55153
diff changeset
   261
                  end;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   262
          in
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   263
            Pretty.block_enclose (
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   264
              Pretty.block [
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   265
                str "instance ",
33991
haftmann
parents: 33421
diff changeset
   266
                Pretty.block (print_typcontext tyvars vs),
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   267
                str (class_name class ^ " "),
33991
haftmann
parents: 33421
diff changeset
   268
                print_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   269
                str " where {"
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   270
              ],
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   271
              str "};"
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   272
            ) (map print_classparam_instance inst_params)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   273
          end;
33991
haftmann
parents: 33421
diff changeset
   274
  in print_stmt end;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   275
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   276
fun haskell_program_of_program ctxt module_prefix module_name reserved identifiers =
39202
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   277
  let
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   278
    fun namify_fun upper base (nsp_fun, nsp_typ) =
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   279
      let
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42003
diff changeset
   280
        val (base', nsp_fun') =
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42003
diff changeset
   281
          Name.variant (if upper then first_upper base else base) nsp_fun;
39202
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   282
      in (base', (nsp_fun', nsp_typ)) end;
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   283
    fun namify_typ base (nsp_fun, nsp_typ) =
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   284
      let
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42003
diff changeset
   285
        val (base', nsp_typ') = Name.variant (first_upper base) nsp_typ;
39202
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   286
      in (base', (nsp_fun, nsp_typ')) end;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   287
    fun namify_stmt (Code_Thingol.Fun (_, SOME _)) = pair
39202
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   288
      | namify_stmt (Code_Thingol.Fun _) = namify_fun false
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   289
      | namify_stmt (Code_Thingol.Datatype _) = namify_typ
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   290
      | namify_stmt (Code_Thingol.Datatypecons _) = namify_fun true
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   291
      | namify_stmt (Code_Thingol.Class _) = namify_typ
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   292
      | namify_stmt (Code_Thingol.Classrel _) = pair
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   293
      | namify_stmt (Code_Thingol.Classparam _) = namify_fun false
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   294
      | namify_stmt (Code_Thingol.Classinst _) = pair;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   295
    fun select_stmt (Code_Thingol.Fun (_, SOME _)) = false
39202
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   296
      | select_stmt (Code_Thingol.Fun _) = true
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   297
      | select_stmt (Code_Thingol.Datatype _) = true
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   298
      | select_stmt (Code_Thingol.Datatypecons _) = false
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   299
      | select_stmt (Code_Thingol.Class _) = true
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   300
      | select_stmt (Code_Thingol.Classrel _) = false
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   301
      | select_stmt (Code_Thingol.Classparam _) = false
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   302
      | select_stmt (Code_Thingol.Classinst _) = true;
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   303
  in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   304
    Code_Namespace.flat_program ctxt
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51143
diff changeset
   305
      { module_prefix = module_prefix, module_name = module_name, reserved = reserved,
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51143
diff changeset
   306
        identifiers = identifiers, empty_nsp = (reserved, reserved), namify_stmt = namify_stmt,
39202
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   307
        modify_stmt = fn stmt => if select_stmt stmt then SOME stmt else NONE }
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   308
  end;
dd0660d93c31 added generic flat_program procedure
haftmann
parents: 39142
diff changeset
   309
48431
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   310
val prelude_import_operators = [
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   311
  "==", "/=", "<", "<=", ">=", ">", "+", "-", "*", "/", "**", ">>=", ">>", "=<<", "&&", "||", "^", "^^", ".", "$", "$!", "++", "!!"
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   312
];
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   313
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   314
val prelude_import_unqualified = [
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   315
  "Eq",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   316
  "error",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   317
  "id",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   318
  "return",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   319
  "not",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   320
  "fst", "snd",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   321
  "map", "filter", "concat", "concatMap", "reverse", "zip", "null", "takeWhile", "dropWhile", "all", "any",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   322
  "Integer", "negate", "abs", "divMod",
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   323
  "String"
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   324
];
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   325
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   326
val prelude_import_unqualified_constr = [
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   327
  ("Bool", ["True", "False"]),
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   328
  ("Maybe", ["Nothing", "Just"])
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   329
];
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   330
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   331
fun serialize_haskell module_prefix string_classes ctxt { module_name,
55683
5732a55b9232 explicit option for "open" code generation
haftmann
parents: 55681
diff changeset
   332
    reserved_syms, identifiers, includes, class_syntax, tyco_syntax, const_syntax } exports program =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   333
  let
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   334
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   335
    (* build program *)
38926
24f82786cc57 record argument for serializers
haftmann
parents: 38924
diff changeset
   336
    val reserved = fold (insert (op =) o fst) includes reserved_syms;
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   337
    val { deresolver, flat_program = haskell_program } = haskell_program_of_program
55683
5732a55b9232 explicit option for "open" code generation
haftmann
parents: 55681
diff changeset
   338
      ctxt module_prefix module_name (Name.make_context reserved) identifiers exports program;
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   339
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   340
    (* print statements *)
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   341
    fun deriving_show tyco =
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   342
      let
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   343
        fun deriv _ "fun" = false
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   344
          | deriv tycos tyco = member (op =) tycos tyco
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   345
              orelse case try (Code_Symbol.Graph.get_node program) (Type_Constructor tyco)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   346
                of SOME (Code_Thingol.Datatype (_, cs)) => forall (deriv' (tyco :: tycos))
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   347
                    (maps snd cs)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   348
                 | NONE => true
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   349
        and deriv' tycos (tyco `%% tys) = deriv tycos tyco
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   350
              andalso forall (deriv' tycos) tys
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   351
          | deriv' _ (ITyVar _) = true
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   352
      in deriv [] tyco end;
47609
b3dab1892cda dropped dead code
haftmann
parents: 46961
diff changeset
   353
    fun print_stmt deresolve = print_haskell_stmt
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   354
      class_syntax tyco_syntax const_syntax (make_vars reserved)
45009
99e1965f9c21 syntactic improvements and tuning names in the code generator due to Florian's code review
bulwahn
parents: 44940
diff changeset
   355
      deresolve (if string_classes then deriving_show else K false);
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   356
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   357
    (* print modules *)
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   358
    fun print_module_frame module_name exports ps =
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   359
      (module_name, Pretty.chunks2 (
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   360
        concat [str "module",
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   361
          case exports of
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   362
            SOME ps => Pretty.block [str module_name, enclose "(" ")" (commas ps)]
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   363
          | NONE => str module_name,
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   364
          str "where {"
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   365
        ]
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   366
        :: ps
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   367
        @| str "}"
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   368
      ));
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   369
    fun print_qualified_import module_name =
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   370
      semicolon [str "import qualified", str module_name];
48431
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   371
    val import_common_ps =
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   372
      enclose "import Prelude (" ");" (commas (map str
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   373
        (map (Library.enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   374
          @ map (fn (tyco, constrs) => (enclose (tyco ^ "(") ")" o commas o map str) constrs) prelude_import_unqualified_constr))
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   375
      :: print_qualified_import "Prelude"
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   376
      :: map (print_qualified_import o fst) includes;
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   377
    fun print_module module_name (gr, imports) =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   378
      let
48431
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   379
        val deresolve = deresolver module_name;
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   380
        val deresolve_import = SOME o str o deresolve;
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   381
        val deresolve_import_attached = SOME o str o suffix "(..)" o deresolve;
55681
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   382
        fun print_import (sym, (_, Code_Thingol.Fun _)) = deresolve_import sym
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   383
          | print_import (sym, (Code_Namespace.Public, Code_Thingol.Datatype _)) = deresolve_import_attached sym
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   384
          | print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Datatype _)) = deresolve_import sym
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   385
          | print_import (sym, (Code_Namespace.Public, Code_Thingol.Class _)) = deresolve_import_attached sym
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   386
          | print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Class _)) = deresolve_import sym
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   387
          | print_import (sym, (_, Code_Thingol.Classinst _)) = NONE;
48431
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   388
        val import_ps = import_common_ps @ map (print_qualified_import o fst) imports;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   389
        fun print_stmt' sym = case Code_Symbol.Graph.get_node gr sym
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   390
         of (_, NONE) => NONE
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   391
          | (_, SOME (export, stmt)) =>
55681
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   392
              SOME (if Code_Namespace.not_private export then print_import (sym, (export, stmt)) else NONE, markup_stmt sym (print_stmt deresolve (sym, stmt)));
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   393
        val (export_ps, body_ps) = (flat o rev o Code_Symbol.Graph.strong_conn) gr
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   394
          |> map_filter print_stmt'
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   395
          |> split_list
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   396
          |> apfst (map_filter I);
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   397
      in
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   398
        print_module_frame module_name (SOME export_ps)
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   399
          ((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   400
      end;
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   401
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   402
    (*serialization*)
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   403
    fun write_module width (SOME destination) (module_name, content) =
38915
haftmann
parents: 38913
diff changeset
   404
          let
42003
6e45dc518ebb replaced File.check by specific File.check_file, File.check_dir;
wenzelm
parents: 41952
diff changeset
   405
            val _ = File.check_dir destination;
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   406
            val filepath = (Path.append destination o Path.ext "hs" o Path.explode o implode
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   407
              o separate "/" o Long_Name.explode) module_name;
41246
e1da70df68c1 allocate intermediate directories in module hierarchy
haftmann
parents: 40804
diff changeset
   408
            val _ = Isabelle_System.mkdirs (Path.dir filepath);
39209
1ca9055ba1f7 only write ghc pragma when writing to a file
haftmann
parents: 39207
diff changeset
   409
          in
1ca9055ba1f7 only write ghc pragma when writing to a file
haftmann
parents: 39207
diff changeset
   410
            (File.write filepath o format [] width o Pretty.chunks2)
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
   411
              [str language_pragma, content]
39209
1ca9055ba1f7 only write ghc pragma when writing to a file
haftmann
parents: 39207
diff changeset
   412
          end
39056
fa197571676b formal markup of generated code for statements
haftmann
parents: 39055
diff changeset
   413
      | write_module width NONE (_, content) = writeln (format [] width content);
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   414
  in
38916
haftmann
parents: 38915
diff changeset
   415
    Code_Target.serialization
38915
haftmann
parents: 38913
diff changeset
   416
      (fn width => fn destination => K () o map (write_module width destination))
40705
03f1266a066e toplevel deresolving for flat module name space
haftmann
parents: 39209
diff changeset
   417
      (fn present => fn width => rpair (try (deresolver ""))
48568
084cd758a8ab evaluation: allow multiple code modules
haftmann
parents: 48431
diff changeset
   418
        o (map o apsnd) (format present width))
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55374
diff changeset
   419
      (map (fn (module_name, content) => print_module_frame module_name NONE [content]) includes
39204
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   420
        @ map (fn module_name => print_module module_name (Graph.get_node haskell_program module_name))
3d30f501b7c2 Haskell uses generic flat_program combinator
haftmann
parents: 39202
diff changeset
   421
          ((flat o rev o Graph.strong_conn) haskell_program))
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   422
  end;
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   423
38966
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   424
val serializer : Code_Target.serializer =
39206
303b63be1a9d moved flat_program to code_namespace
haftmann
parents: 39204
diff changeset
   425
  Code_Target.parse_args (Scan.optional (Args.$$$ "root" -- Args.colon |-- Args.name) ""
38966
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   426
    -- Scan.optional (Args.$$$ "string_classes" >> K true) false
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   427
    >> (fn (module_prefix, string_classes) =>
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   428
      serialize_haskell module_prefix string_classes));
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   429
28064
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   430
val literals = let
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   431
  fun char_haskell c =
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   432
    let
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   433
      val s = ML_Syntax.print_char c;
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   434
    in if s = "'" then "\\'" else s end;
34944
970e1466028d code literals: distinguish numeral classes by different entries
haftmann
parents: 34269
diff changeset
   435
  fun numeral_haskell k = if k >= 0 then string_of_int k
970e1466028d code literals: distinguish numeral classes by different entries
haftmann
parents: 34269
diff changeset
   436
    else Library.enclose "(" ")" (signed_string_of_int k);
28064
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   437
in Literals {
34178
a78b8d5b91cb take care for destructive print mode properly using dedicated pretty builders
haftmann
parents: 34085
diff changeset
   438
  literal_char = Library.enclose "'" "'" o char_haskell,
28064
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   439
  literal_string = quote o translate_string char_haskell,
34944
970e1466028d code literals: distinguish numeral classes by different entries
haftmann
parents: 34269
diff changeset
   440
  literal_numeral = numeral_haskell,
34178
a78b8d5b91cb take care for destructive print mode properly using dedicated pretty builders
haftmann
parents: 34085
diff changeset
   441
  literal_list = enum "," "[" "]",
28064
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   442
  infix_cons = (5, ":")
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   443
} end;
d4a6460c53d1 restructured code generation of literals
haftmann
parents: 28054
diff changeset
   444
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   445
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   446
(** optional monad syntax **)
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   447
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   448
fun pretty_haskell_monad c_bind =
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   449
  let
31874
f172346ba805 simplified binding concept
haftmann
parents: 31775
diff changeset
   450
    fun dest_bind t1 t2 = case Code_Thingol.split_pat_abs t2
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   451
     of SOME ((pat, ty), t') =>
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   452
          SOME ((SOME ((pat, ty), true), t1), t')
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   453
      | NONE => NONE;
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   454
    fun dest_monad (IConst { sym = Constant c, ... } `$ t1 `$ t2) =
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   455
          if c = c_bind then dest_bind t1 t2
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   456
          else NONE
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   457
      | dest_monad t = case Code_Thingol.split_let t
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   458
         of SOME (((pat, ty), tbind), t') =>
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   459
              SOME ((SOME ((pat, ty), false), tbind), t')
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   460
          | NONE => NONE;
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   461
    val implode_monad = Code_Thingol.unfoldr dest_monad;
33991
haftmann
parents: 33421
diff changeset
   462
    fun print_monad print_bind print_term (NONE, t) vars =
haftmann
parents: 33421
diff changeset
   463
          (semicolon [print_term vars NOBR t], vars)
haftmann
parents: 33421
diff changeset
   464
      | print_monad print_bind print_term (SOME ((bind, _), true), t) vars = vars
haftmann
parents: 33421
diff changeset
   465
          |> print_bind NOBR bind
haftmann
parents: 33421
diff changeset
   466
          |>> (fn p => semicolon [p, str "<-", print_term vars NOBR t])
haftmann
parents: 33421
diff changeset
   467
      | print_monad print_bind print_term (SOME ((bind, _), false), t) vars = vars
haftmann
parents: 33421
diff changeset
   468
          |> print_bind NOBR bind
37832
f8fcfc678280 braced needed in layout-insensitive syntax
haftmann
parents: 37822
diff changeset
   469
          |>> (fn p => semicolon [str "let", str "{", p, str "=", print_term vars NOBR t, str "}"]);
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   470
    fun pretty _ print_term thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   471
     of SOME (bind, t') => let
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   472
          val (binds, t'') = implode_monad t'
33991
haftmann
parents: 33421
diff changeset
   473
          val (ps, vars') = fold_map (print_monad (gen_print_bind (K print_term) thm) print_term)
haftmann
parents: 33421
diff changeset
   474
            (bind :: binds) vars;
haftmann
parents: 33421
diff changeset
   475
        in
37833
1381665d9550 more consistent spacing in generated monadic code
haftmann
parents: 37832
diff changeset
   476
          (brackify fxy o single o enclose "do { " " }" o Pretty.breaks)
33991
haftmann
parents: 33421
diff changeset
   477
            (ps @| print_term vars' NOBR t'')
haftmann
parents: 33421
diff changeset
   478
        end
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   479
      | NONE => brackify_infix (1, L) fxy
37242
97097e589715 brackify_infix etc.: no break before infix operator -- eases survival in Scala
haftmann
parents: 36960
diff changeset
   480
          (print_term vars (INFX (1, L)) t1, str ">>=", print_term vars (INFX (1, X)) t2)
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   481
  in (2, pretty) end;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   482
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28064
diff changeset
   483
fun add_monad target' raw_c_bind thy =
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   484
  let
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31054
diff changeset
   485
    val c_bind = Code.read_const thy raw_c_bind;
52801
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   486
  in
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   487
    if target = target' then
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   488
      thy
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   489
      |> Code_Target.set_printings (Constant (c_bind, [(target,
52801
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   490
        SOME (Code_Printer.complex_const_syntax (pretty_haskell_monad c_bind)))]))
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   491
    else error "Only Haskell target allows for monad syntax"
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   492
  end;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   493
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   494
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   495
(** Isar setup **)
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   496
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   497
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45009
diff changeset
   498
  Outer_Syntax.command @{command_spec "code_monad"} "define code syntax for monads"
52801
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   499
    (Parse.term -- Parse.name >> (fn (raw_bind, target) =>
6f88e379aa3e proper PIDE markup for codegen arguments;
wenzelm
parents: 52519
diff changeset
   500
      Toplevel.theory (add_monad target raw_bind)));
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   501
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   502
val setup =
37821
3cbb22cec751 formal slot for code checker
haftmann
parents: 37819
diff changeset
   503
  Code_Target.add_target
38966
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   504
    (target, { serializer = serializer, literals = literals,
41952
c7297638599b cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
wenzelm
parents: 41940
diff changeset
   505
      check = { env_var = "ISABELLE_GHC", make_destination = I,
41940
a3b68a7a0e15 allow spaces in executable names;
wenzelm
parents: 41343
diff changeset
   506
        make_command = fn module_name =>
44926
de3ed037c9a5 create central list for language extensions used by the haskell code generator
noschinl
parents: 44852
diff changeset
   507
          "\"$ISABELLE_GHC\" " ^ language_params  ^ " -odir build -hidir build -stubdir build -e \"\" " ^
41940
a3b68a7a0e15 allow spaces in executable names;
wenzelm
parents: 41343
diff changeset
   508
            module_name ^ ".hs" } })
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55148
diff changeset
   509
  #> Code_Target.set_printings (Type_Constructor ("fun",
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52138
diff changeset
   510
    [(target, SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
37242
97097e589715 brackify_infix etc.: no break before infix operator -- eases survival in Scala
haftmann
parents: 36960
diff changeset
   511
      brackify_infix (1, R) fxy (
33991
haftmann
parents: 33421
diff changeset
   512
        print_typ (INFX (1, X)) ty1,
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   513
        str "->",
33991
haftmann
parents: 33421
diff changeset
   514
        print_typ (INFX (1, R)) ty2
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52138
diff changeset
   515
      )))]))
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   516
  #> fold (Code_Target.add_reserved target) [
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   517
      "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   518
      "import", "default", "forall", "let", "in", "class", "qualified", "data",
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   519
      "newtype", "instance", "if", "then", "else", "type", "as", "do", "module"
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   520
    ]
48431
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   521
  #> fold (Code_Target.add_reserved target) prelude_import_unqualified
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   522
  #> fold (Code_Target.add_reserved target o fst) prelude_import_unqualified_constr
6efff142bb54 restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents: 48072
diff changeset
   523
  #> fold (fold (Code_Target.add_reserved target) o snd) prelude_import_unqualified_constr;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   524
2b84d34c5d02 restructured and split code serializer module
haftmann
parents:
diff changeset
   525
end; (*struct*)