author | wenzelm |
Tue, 10 Sep 2024 20:06:51 +0200 | |
changeset 80847 | 93c2058896a4 |
parent 77703 | 0262155d2743 |
child 80910 | 406a85a25189 |
permissions | -rw-r--r-- |
37745
6315b6426200
checking generated code for various target languages
haftmann
parents:
37669
diff
changeset
|
1 |
(* Title: Tools/Code/code_haskell.ML |
28054 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
4 |
Serializer for Haskell. |
|
5 |
*) |
|
6 |
||
7 |
signature CODE_HASKELL = |
|
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 | 12 |
end; |
13 |
||
14 |
structure Code_Haskell : CODE_HASKELL = |
|
15 |
struct |
|
16 |
||
17 |
val target = "Haskell"; |
|
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 | 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 | 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 | 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 | 28 |
open Basic_Code_Symbol; |
28054 | 29 |
open Basic_Code_Thingol; |
30 |
open Code_Printer; |
|
31 |
||
32 |
infixr 5 @@; |
|
33 |
infixr 5 @|; |
|
34 |
||
35 |
||
36 |
(** Haskell serializer **) |
|
37 |
||
68028 | 38 |
val print_haskell_string = |
69210 | 39 |
quote o translate_string (fn c => |
40 |
if Symbol.is_ascii c then GHC.print_codepoint (ord c) |
|
41 |
else error "non-ASCII byte in Haskell string literal"); |
|
68028 | 42 |
|
47609 | 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 | 45 |
let |
55150 | 46 |
val deresolve_const = deresolve o Constant; |
47 |
val deresolve_tyco = deresolve o Type_Constructor; |
|
48 |
val deresolve_class = deresolve o Type_Class; |
|
38923 | 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 | 51 |
| SOME class => class; |
33991 | 52 |
fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs |
28054 | 53 |
of [] => [] |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37242
diff
changeset
|
54 |
| constraints => enum "," "(" ")" ( |
28054 | 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 | 57 |
@@ str " => "; |
33991 | 58 |
fun print_typforall tyvars vs = case map fst vs |
28054 | 59 |
of [] => [] |
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 | 62 |
fun print_tyco_expr tyvars fxy (tyco, tys) = |
63 |
brackify fxy (str tyco :: map (print_typ tyvars BR) tys) |
|
47609 | 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 | 66 |
| SOME (_, print) => print (print_typ tyvars) fxy tys) |
33991 | 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 | 70 |
fun print_typscheme tyvars (vs, ty) = |
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 | 74 |
| print_term tyvars some_thm vars fxy (t as (t1 `$ t2)) = |
28054 | 75 |
(case Code_Thingol.unfold_const_app t |
35228 | 76 |
of SOME app => print_app tyvars some_thm vars fxy app |
28054 | 77 |
| _ => |
78 |
brackify fxy [ |
|
35228 | 79 |
print_term tyvars some_thm vars NOBR t1, |
80 |
print_term tyvars some_thm vars BR t2 |
|
28054 | 81 |
]) |
35228 | 82 |
| print_term tyvars some_thm vars fxy (IVar NONE) = |
31889 | 83 |
str "_" |
35228 | 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 | 86 |
| print_term tyvars some_thm vars fxy (t as _ `|=> _) = |
28054 | 87 |
let |
31874 | 88 |
val (binds, t') = Code_Thingol.unfold_pat_abs t; |
35228 | 89 |
val (ps, vars') = fold_map (print_bind tyvars some_thm BR o fst) binds vars; |
90 |
in brackets (str "\\" :: ps @ str "->" @@ print_term tyvars some_thm vars' NOBR t') end |
|
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 | 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 | 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 | 101 |
case annotation of |
102 |
SOME ty => brackets [(str o deresolve) sym, str "::", print_typ tyvars NOBR ty] |
|
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 | 107 |
and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) const_syntax |
35228 | 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 | 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 | 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 | 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 | 119 |
in brackify_block fxy (str "let {") |
120 |
ps |
|
35228 | 121 |
(concat [str "}", str "in", print_term tyvars some_thm vars' NOBR body]) |
28054 | 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 | 124 |
let |
33991 | 125 |
fun print_select (pat, body) = |
28054 | 126 |
let |
35228 | 127 |
val (p, vars') = print_bind tyvars some_thm NOBR pat vars; |
128 |
in semicolon [p, str "->", print_term tyvars some_thm vars' NOBR body] end; |
|
36576 | 129 |
in Pretty.block_enclose |
130 |
(concat [str "(case", print_term tyvars some_thm vars NOBR t, str "of", str "{"], str "})") |
|
33991 | 131 |
(map print_select clauses) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
132 |
end; |
55150 | 133 |
fun print_stmt (Constant const, Code_Thingol.Fun (((vs, ty), raw_eqs), _)) = |
28054 | 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 | 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 | 142 |
@@ print_haskell_string const |
34269 | 143 |
); |
35228 | 144 |
fun print_eqn ((ts, t), (some_thm, _)) = |
28054 | 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) |
75356 | 149 |
|> intro_vars (build (fold Code_Thingol.add_varnames ts)); |
28054 | 150 |
in |
151 |
semicolon ( |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
152 |
(str o deresolve_const) const |
35228 | 153 |
:: map (print_term tyvars some_thm vars BR) ts |
28054 | 154 |
@ str "=" |
35228 | 155 |
@@ print_term tyvars some_thm vars NOBR t |
28054 | 156 |
) |
157 |
end; |
|
158 |
in |
|
159 |
Pretty.chunks ( |
|
33991 | 160 |
semicolon [ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
161 |
(str o suffix " ::" o deresolve_const) const, |
33991 | 162 |
print_typscheme tyvars (vs, ty) |
28054 | 163 |
] |
34269 | 164 |
:: (case filter (snd o snd) raw_eqs |
165 |
of [] => [print_err ((length o fst o Code_Thingol.unfold_fun) ty)] |
|
166 |
| eqs => map print_eqn eqs) |
|
28054 | 167 |
) |
168 |
end |
|
55150 | 169 |
| print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [])) = |
28054 | 170 |
let |
48003
1d11af40b106
dropped sort constraints on datatype specifications
haftmann
parents:
47609
diff
changeset
|
171 |
val tyvars = intro_vars vs reserved; |
28054 | 172 |
in |
173 |
semicolon [ |
|
174 |
str "data", |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
175 |
print_typdecl tyvars (deresolve_tyco tyco, vs) |
28054 | 176 |
] |
177 |
end |
|
55150 | 178 |
| print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [((co, _), [ty])])) = |
28054 | 179 |
let |
48003
1d11af40b106
dropped sort constraints on datatype specifications
haftmann
parents:
47609
diff
changeset
|
180 |
val tyvars = intro_vars vs reserved; |
28054 | 181 |
in |
182 |
semicolon ( |
|
183 |
str "newtype" |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
184 |
:: print_typdecl tyvars (deresolve_tyco tyco, vs) |
28054 | 185 |
:: str "=" |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
186 |
:: (str o deresolve_const) co |
33991 | 187 |
:: print_typ tyvars BR ty |
55374 | 188 |
:: (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else []) |
28054 | 189 |
) |
190 |
end |
|
55150 | 191 |
| print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, co :: cos)) = |
28054 | 192 |
let |
48003
1d11af40b106
dropped sort constraints on datatype specifications
haftmann
parents:
47609
diff
changeset
|
193 |
val tyvars = intro_vars vs reserved; |
37449 | 194 |
fun print_co ((co, _), tys) = |
28054 | 195 |
concat ( |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
196 |
(str o deresolve_const) co |
33991 | 197 |
:: map (print_typ tyvars BR) tys |
28054 | 198 |
) |
199 |
in |
|
200 |
semicolon ( |
|
201 |
str "data" |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
202 |
:: print_typdecl tyvars (deresolve_tyco tyco, vs) |
28054 | 203 |
:: str "=" |
33991 | 204 |
:: print_co co |
205 |
:: map ((fn p => Pretty.block [str "| ", p]) o print_co) cos |
|
55374 | 206 |
@ (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else []) |
28054 | 207 |
) |
208 |
end |
|
55150 | 209 |
| print_stmt (Type_Class class, Code_Thingol.Class (v, (classrels, classparams))) = |
28054 | 210 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
211 |
val tyvars = intro_vars [v] reserved; |
33991 | 212 |
fun print_classparam (classparam, ty) = |
28054 | 213 |
semicolon [ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
214 |
(str o deresolve_const) classparam, |
28054 | 215 |
str "::", |
33991 | 216 |
print_typ tyvars NOBR ty |
28054 | 217 |
] |
218 |
in |
|
219 |
Pretty.block_enclose ( |
|
220 |
Pretty.block [ |
|
221 |
str "class ", |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
222 |
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
|
223 |
str (deresolve_class class ^ " " ^ lookup_var tyvars v), |
28054 | 224 |
str " where {" |
225 |
], |
|
226 |
str "};" |
|
33991 | 227 |
) (map print_classparam classparams) |
28054 | 228 |
end |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
229 |
| print_stmt (_, Code_Thingol.Classinst { class, tyco, vs, inst_params, ... }) = |
28054 | 230 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
231 |
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
|
232 |
fun print_classparam_instance ((classparam, (const, _)), (thm, _)) = |
55373
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
233 |
case const_syntax classparam of |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
234 |
NONE => semicolon [ |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
235 |
(str o Long_Name.base_name o deresolve_const) classparam, |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
236 |
str "=", |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
237 |
print_app tyvars (SOME thm) reserved NOBR (const, []) |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
238 |
] |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
239 |
| SOME (_, Code_Printer.Plain_printer s) => semicolon [ |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
240 |
(str o Long_Name.base_name) s, |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
241 |
str "=", |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
242 |
print_app tyvars (SOME thm) reserved NOBR (const, []) |
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
243 |
] |
77232 | 244 |
| SOME (wanted, Code_Printer.Complex_printer _) => |
55373
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
245 |
let |
77233
6bdd125d932b
explicit range types in abstractions
stuebinm <stuebinm@disroot.org>
parents:
77232
diff
changeset
|
246 |
val { sym = Constant c, dom, range, ... } = const; |
77703
0262155d2743
more uniform approach towards satisfied applications
haftmann
parents:
77233
diff
changeset
|
247 |
val ((vs_tys, (ts, _)), _) = Code_Thingol.satisfied_application wanted (const, []); |
0262155d2743
more uniform approach towards satisfied applications
haftmann
parents:
77233
diff
changeset
|
248 |
val vs = map fst vs_tys; |
0262155d2743
more uniform approach towards satisfied applications
haftmann
parents:
77233
diff
changeset
|
249 |
val rhs = IConst const `$$ ts; |
55373
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 = [], |
77233
6bdd125d932b
explicit range types in abstractions
stuebinm <stuebinm@disroot.org>
parents:
77232
diff
changeset
|
255 |
dicts = [], dom = dom, range = range, annotation = NONE } `$$ map IVar vs; |
77231 | 256 |
(*dictionaries are not relevant in Haskell, |
55373
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 | 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 | 263 |
] |
55373
2b4204cb7904
method names in instance declarations are always unqualified
haftmann
parents:
55153
diff
changeset
|
264 |
end; |
28054 | 265 |
in |
266 |
Pretty.block_enclose ( |
|
267 |
Pretty.block [ |
|
268 |
str "instance ", |
|
33991 | 269 |
Pretty.block (print_typcontext tyvars vs), |
28054 | 270 |
str (class_name class ^ " "), |
33991 | 271 |
print_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), |
28054 | 272 |
str " where {" |
273 |
], |
|
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 | 276 |
end; |
33991 | 277 |
in print_stmt end; |
28054 | 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 | 280 |
let |
281 |
fun namify_fun upper base (nsp_fun, nsp_typ) = |
|
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 | 284 |
in (base', (nsp_fun', nsp_typ)) end; |
285 |
fun namify_typ base (nsp_fun, nsp_typ) = |
|
286 |
let |
|
56812 | 287 |
val (base', nsp_typ') = Name.variant (Name.enforce_case true base) nsp_typ; |
39202 | 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 | 290 |
| namify_stmt (Code_Thingol.Fun _) = namify_fun false |
291 |
| namify_stmt (Code_Thingol.Datatype _) = namify_typ |
|
292 |
| namify_stmt (Code_Thingol.Datatypecons _) = namify_fun true |
|
293 |
| namify_stmt (Code_Thingol.Class _) = namify_typ |
|
294 |
| namify_stmt (Code_Thingol.Classrel _) = pair |
|
295 |
| namify_stmt (Code_Thingol.Classparam _) = namify_fun false |
|
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 | 298 |
| select_stmt (Code_Thingol.Fun _) = true |
299 |
| select_stmt (Code_Thingol.Datatype _) = true |
|
300 |
| select_stmt (Code_Thingol.Datatypecons _) = false |
|
301 |
| select_stmt (Code_Thingol.Class _) = true |
|
302 |
| select_stmt (Code_Thingol.Classrel _) = false |
|
303 |
| select_stmt (Code_Thingol.Classparam _) = false |
|
304 |
| select_stmt (Code_Thingol.Classinst _) = true; |
|
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 | 309 |
modify_stmt = fn stmt => if select_stmt stmt then SOME stmt else NONE } |
310 |
end; |
|
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, |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
334 |
reserved_syms, identifiers, includes, class_syntax, tyco_syntax, const_syntax } program exports = |
28054 | 335 |
let |
39204 | 336 |
|
337 |
(* build program *) |
|
38926 | 338 |
val reserved = fold (insert (op =) o fst) includes reserved_syms; |
39204 | 339 |
val { deresolver, flat_program = haskell_program } = haskell_program_of_program |
55683 | 340 |
ctxt module_prefix module_name (Name.make_context reserved) identifiers exports program; |
39204 | 341 |
|
342 |
(* print statements *) |
|
28054 | 343 |
fun deriving_show tyco = |
344 |
let |
|
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 | 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 | 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 | 351 |
and deriv' tycos (tyco `%% tys) = deriv tycos tyco |
352 |
andalso forall (deriv' tycos) tys |
|
353 |
| deriv' _ (ITyVar _) = true |
|
354 |
in deriv [] tyco end; |
|
47609 | 355 |
fun print_stmt deresolve = print_haskell_stmt |
39204 | 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 | 358 |
|
359 |
(* print modules *) |
|
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
360 |
fun module_names module_name = |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
361 |
let |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
362 |
val (xs, x) = split_last (Long_Name.explode module_name) |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
363 |
in xs @ [x ^ ".hs"] end |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
364 |
fun print_module_frame module_name header exports ps = |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
365 |
(module_names module_name, Pretty.chunks2 ( |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
366 |
header |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
367 |
@ concat [str "module", |
55679 | 368 |
case exports of |
369 |
SOME ps => Pretty.block [str module_name, enclose "(" ")" (commas ps)] |
|
370 |
| NONE => str module_name, |
|
371 |
str "where {" |
|
372 |
] |
|
39204 | 373 |
:: ps |
374 |
@| str "}" |
|
375 |
)); |
|
55679 | 376 |
fun print_qualified_import module_name = |
377 |
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
|
378 |
val import_common_ps = |
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents:
48072
diff
changeset
|
379 |
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
|
380 |
(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
|
381 |
@ 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
|
382 |
:: print_qualified_import "Prelude" |
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents:
48072
diff
changeset
|
383 |
:: map (print_qualified_import o fst) includes; |
39204 | 384 |
fun print_module module_name (gr, imports) = |
28054 | 385 |
let |
48431
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents:
48072
diff
changeset
|
386 |
val deresolve = deresolver module_name; |
55679 | 387 |
val deresolve_import = SOME o str o deresolve; |
388 |
val deresolve_import_attached = SOME o str o suffix "(..)" o deresolve; |
|
55681 | 389 |
fun print_import (sym, (_, Code_Thingol.Fun _)) = deresolve_import sym |
390 |
| print_import (sym, (Code_Namespace.Public, Code_Thingol.Datatype _)) = deresolve_import_attached sym |
|
391 |
| print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Datatype _)) = deresolve_import sym |
|
392 |
| print_import (sym, (Code_Namespace.Public, Code_Thingol.Class _)) = deresolve_import_attached sym |
|
393 |
| print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Class _)) = deresolve_import sym |
|
394 |
| 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
|
395 |
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
|
396 |
fun print_stmt' sym = case Code_Symbol.Graph.get_node gr sym |
39204 | 397 |
of (_, NONE) => NONE |
55679 | 398 |
| (_, SOME (export, stmt)) => |
55681 | 399 |
SOME (if Code_Namespace.not_private export then print_import (sym, (export, stmt)) else NONE, markup_stmt sym (print_stmt deresolve (sym, stmt))); |
55679 | 400 |
val (export_ps, body_ps) = (flat o rev o Code_Symbol.Graph.strong_conn) gr |
401 |
|> map_filter print_stmt' |
|
402 |
|> split_list |
|
403 |
|> apfst (map_filter I); |
|
39204 | 404 |
in |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
405 |
print_module_frame module_name [str language_pragma] (SOME export_ps) |
39204 | 406 |
((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps) |
407 |
end; |
|
408 |
||
28054 | 409 |
in |
69690 | 410 |
(Code_Target.Hierarchy (map (fn (module_name, content) => ([module_name ^ ".hs"], content)) includes |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
411 |
@ map (fn module_name => print_module module_name (Graph.get_node haskell_program module_name)) |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69593
diff
changeset
|
412 |
((flat o rev o Graph.strong_conn) haskell_program)), try (deresolver "")) |
28054 | 413 |
end; |
414 |
||
38966 | 415 |
val serializer : Code_Target.serializer = |
39206 | 416 |
Code_Target.parse_args (Scan.optional (Args.$$$ "root" -- Args.colon |-- Args.name) "" |
38966 | 417 |
-- Scan.optional (Args.$$$ "string_classes" >> K true) false |
418 |
>> (fn (module_prefix, string_classes) => |
|
419 |
serialize_haskell module_prefix string_classes)); |
|
420 |
||
58400
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
haftmann
parents:
58397
diff
changeset
|
421 |
fun print_numeral typ = Library.enclose "(" (" :: " ^ typ ^ ")") o signed_string_of_int; |
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
haftmann
parents:
58397
diff
changeset
|
422 |
|
68028 | 423 |
val literals = Literals { |
424 |
literal_string = print_haskell_string, |
|
58400
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
haftmann
parents:
58397
diff
changeset
|
425 |
literal_numeral = print_numeral "Integer", |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34085
diff
changeset
|
426 |
literal_list = enum "," "[" "]", |
28064 | 427 |
infix_cons = (5, ":") |
68028 | 428 |
}; |
28064 | 429 |
|
28054 | 430 |
|
431 |
(** optional monad syntax **) |
|
432 |
||
433 |
fun pretty_haskell_monad c_bind = |
|
434 |
let |
|
31874 | 435 |
fun dest_bind t1 t2 = case Code_Thingol.split_pat_abs t2 |
31889 | 436 |
of SOME ((pat, ty), t') => |
437 |
SOME ((SOME ((pat, ty), true), t1), t') |
|
28145 | 438 |
| NONE => NONE; |
55150 | 439 |
fun dest_monad (IConst { sym = Constant c, ... } `$ t1 `$ t2) = |
55148
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
haftmann
parents:
55147
diff
changeset
|
440 |
if c = c_bind then dest_bind t1 t2 |
28145 | 441 |
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
|
442 |
| 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
|
443 |
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
|
444 |
SOME ((SOME ((IVar some_v, ty), false), tbind), t') |
28145 | 445 |
| NONE => NONE; |
55148
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
haftmann
parents:
55147
diff
changeset
|
446 |
val implode_monad = Code_Thingol.unfoldr dest_monad; |
33991 | 447 |
fun print_monad print_bind print_term (NONE, t) vars = |
448 |
(semicolon [print_term vars NOBR t], vars) |
|
449 |
| print_monad print_bind print_term (SOME ((bind, _), true), t) vars = vars |
|
450 |
|> print_bind NOBR bind |
|
451 |
|>> (fn p => semicolon [p, str "<-", print_term vars NOBR t]) |
|
452 |
| print_monad print_bind print_term (SOME ((bind, _), false), t) vars = vars |
|
453 |
|> print_bind NOBR bind |
|
37832 | 454 |
|>> (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
|
455 |
fun pretty _ print_term thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2 |
28145 | 456 |
of SOME (bind, t') => let |
55148
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
haftmann
parents:
55147
diff
changeset
|
457 |
val (binds, t'') = implode_monad t' |
33991 | 458 |
val (ps, vars') = fold_map (print_monad (gen_print_bind (K print_term) thm) print_term) |
459 |
(bind :: binds) vars; |
|
460 |
in |
|
58454 | 461 |
brackify_block fxy (str "do { ") |
33991 | 462 |
(ps @| print_term vars' NOBR t'') |
58454 | 463 |
(str " }") |
33991 | 464 |
end |
28145 | 465 |
| 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
|
466 |
(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
|
467 |
in (2, pretty) end; |
28054 | 468 |
|
28145 | 469 |
fun add_monad target' raw_c_bind thy = |
28054 | 470 |
let |
31156 | 471 |
val c_bind = Code.read_const thy raw_c_bind; |
52801 | 472 |
in |
473 |
if target = target' then |
|
474 |
thy |
|
55150 | 475 |
|> Code_Target.set_printings (Constant (c_bind, [(target, |
52801 | 476 |
SOME (Code_Printer.complex_const_syntax (pretty_haskell_monad c_bind)))])) |
477 |
else error "Only Haskell target allows for monad syntax" |
|
478 |
end; |
|
28054 | 479 |
|
480 |
||
481 |
(** Isar setup **) |
|
482 |
||
59323 | 483 |
val _ = Theory.setup |
484 |
(Code_Target.add_language |
|
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
65483
diff
changeset
|
485 |
(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
|
486 |
check = { env_var = "ISABELLE_GHC", make_destination = I, |
41940 | 487 |
make_command = fn module_name => |
44926
de3ed037c9a5
create central list for language extensions used by the haskell code generator
noschinl
parents:
44852
diff
changeset
|
488 |
"\"$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
|
489 |
module_name ^ ".hs"}, |
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
65483
diff
changeset
|
490 |
evaluation_args = []}) |
55150 | 491 |
#> 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
|
492 |
[(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
|
493 |
brackify_infix (1, R) fxy ( |
33991 | 494 |
print_typ (INFX (1, X)) ty1, |
28054 | 495 |
str "->", |
33991 | 496 |
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
|
497 |
)))])) |
28054 | 498 |
#> fold (Code_Target.add_reserved target) [ |
499 |
"hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", |
|
500 |
"import", "default", "forall", "let", "in", "class", "qualified", "data", |
|
501 |
"newtype", "instance", "if", "then", "else", "type", "as", "do", "module" |
|
502 |
] |
|
48431
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
haftmann
parents:
48072
diff
changeset
|
503 |
#> 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
|
504 |
#> fold (Code_Target.add_reserved target o fst) prelude_import_unqualified_constr |
59323 | 505 |
#> fold (fold (Code_Target.add_reserved target) o snd) prelude_import_unqualified_constr); |
506 |
||
507 |
val _ = |
|
69593 | 508 |
Outer_Syntax.command \<^command_keyword>\<open>code_monad\<close> "define code syntax for monads" |
59323 | 509 |
(Parse.term -- Parse.name >> (fn (raw_bind, target) => |
510 |
Toplevel.theory (add_monad target raw_bind))); |
|
28054 | 511 |
|
512 |
end; (*struct*) |