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