author | haftmann |
Tue, 10 Nov 2009 16:11:46 +0100 | |
changeset 33596 | 27c5023ee818 |
parent 33421 | 3789fe962a08 |
child 33955 | fff6f11b1f09 |
child 33991 | 45d94947426a |
permissions | -rw-r--r-- |
28054 | 1 |
(* Title: Tools/code/code_haskell.ML |
2 |
Author: Florian Haftmann, TU Muenchen |
|
3 |
||
4 |
Serializer for Haskell. |
|
5 |
*) |
|
6 |
||
7 |
signature CODE_HASKELL = |
|
8 |
sig |
|
9 |
val setup: theory -> theory |
|
10 |
end; |
|
11 |
||
12 |
structure Code_Haskell : CODE_HASKELL = |
|
13 |
struct |
|
14 |
||
15 |
val target = "Haskell"; |
|
16 |
||
17 |
open Basic_Code_Thingol; |
|
18 |
open Code_Printer; |
|
19 |
||
20 |
infixr 5 @@; |
|
21 |
infixr 5 @|; |
|
22 |
||
23 |
||
24 |
(** Haskell serializer **) |
|
25 |
||
31054 | 26 |
fun pr_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
27 |
reserved deresolve contr_classparam_typs deriving_show = |
28054 | 28 |
let |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
29 |
val deresolve_base = Long_Name.base_name o deresolve; |
28054 | 30 |
fun class_name class = case syntax_class class |
31 |
of NONE => deresolve class |
|
28687 | 32 |
| SOME class => class; |
28054 | 33 |
fun pr_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs |
34 |
of [] => [] |
|
35 |
| classbinds => Pretty.enum "," "(" ")" ( |
|
36 |
map (fn (v, class) => |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
37 |
str (class_name class ^ " " ^ lookup_var tyvars v)) classbinds) |
28054 | 38 |
@@ str " => "; |
39 |
fun pr_typforall tyvars vs = case map fst vs |
|
40 |
of [] => [] |
|
41 |
| vnames => str "forall " :: Pretty.breaks |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
42 |
(map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1; |
28054 | 43 |
fun pr_tycoexpr tyvars fxy (tyco, tys) = |
44 |
brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) |
|
45 |
and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco |
|
46 |
of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys) |
|
47 |
| SOME (i, pr) => pr (pr_typ tyvars) fxy tys) |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
48 |
| pr_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v; |
28054 | 49 |
fun pr_typdecl tyvars (vs, tycoexpr) = |
50 |
Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr); |
|
51 |
fun pr_typscheme tyvars (vs, ty) = |
|
52 |
Pretty.block (pr_typforall tyvars vs @ pr_typcontext tyvars vs @| pr_typ tyvars NOBR ty); |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
53 |
fun pr_term tyvars thm vars fxy (IConst c) = |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
54 |
pr_app tyvars thm vars fxy (c, []) |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
55 |
| pr_term tyvars thm vars fxy (t as (t1 `$ t2)) = |
28054 | 56 |
(case Code_Thingol.unfold_const_app t |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
57 |
of SOME app => pr_app tyvars thm vars fxy app |
28054 | 58 |
| _ => |
59 |
brackify fxy [ |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
60 |
pr_term tyvars thm vars NOBR t1, |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
61 |
pr_term tyvars thm vars BR t2 |
28054 | 62 |
]) |
31889 | 63 |
| pr_term tyvars thm vars fxy (IVar NONE) = |
64 |
str "_" |
|
65 |
| pr_term tyvars 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
|
66 |
(str o lookup_var vars) v |
31724 | 67 |
| pr_term tyvars thm vars fxy (t as _ `|=> _) = |
28054 | 68 |
let |
31874 | 69 |
val (binds, t') = Code_Thingol.unfold_pat_abs t; |
31889 | 70 |
val (ps, vars') = fold_map (pr_bind tyvars thm BR o fst) binds vars; |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
71 |
in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm vars' NOBR t') end |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
72 |
| pr_term tyvars thm vars fxy (ICase (cases as (_, t0))) = |
28054 | 73 |
(case Code_Thingol.unfold_const_app t0 |
74 |
of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) |
|
75 |
then pr_case tyvars thm vars fxy cases |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
76 |
else pr_app tyvars thm vars fxy c_ts |
28054 | 77 |
| NONE => pr_case tyvars thm vars fxy cases) |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
78 |
and pr_app' tyvars thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
79 |
of [] => (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts |
28054 | 80 |
| fingerprint => let |
81 |
val ts_fingerprint = ts ~~ curry Library.take (length ts) fingerprint; |
|
82 |
val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) => |
|
83 |
(not o Code_Thingol.locally_monomorphic) t) ts_fingerprint; |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
84 |
fun pr_term_anno (t, NONE) _ = pr_term tyvars thm vars BR t |
28054 | 85 |
| pr_term_anno (t, SOME _) ty = |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
86 |
brackets [pr_term tyvars thm vars NOBR t, str "::", pr_typ tyvars NOBR ty]; |
28054 | 87 |
in |
88 |
if needs_annotation then |
|
89 |
(str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry Library.take (length ts) tys) |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
90 |
else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts |
28054 | 91 |
end |
31054 | 92 |
and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const |
31889 | 93 |
and pr_bind tyvars thm fxy p = gen_pr_bind (pr_term tyvars) thm fxy p |
28054 | 94 |
and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) = |
95 |
let |
|
29952
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
haftmann
parents:
29832
diff
changeset
|
96 |
val (binds, body) = Code_Thingol.unfold_let (ICase cases); |
28054 | 97 |
fun pr ((pat, ty), t) vars = |
98 |
vars |
|
31889 | 99 |
|> pr_bind tyvars thm BR pat |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
100 |
|>> (fn p => semicolon [p, str "=", pr_term tyvars thm vars NOBR t]) |
28054 | 101 |
val (ps, vars') = fold_map pr binds vars; |
31665 | 102 |
in brackify_block fxy (str "let {") |
103 |
ps |
|
104 |
(concat [str "}", str "in", pr_term tyvars thm vars' NOBR body]) |
|
28054 | 105 |
end |
29952
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
haftmann
parents:
29832
diff
changeset
|
106 |
| pr_case tyvars thm vars fxy (((t, ty), clauses as _ :: _), _) = |
28054 | 107 |
let |
29952
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
haftmann
parents:
29832
diff
changeset
|
108 |
fun pr (pat, body) = |
28054 | 109 |
let |
31889 | 110 |
val (p, vars') = pr_bind tyvars thm NOBR pat vars; |
29952
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
haftmann
parents:
29832
diff
changeset
|
111 |
in semicolon [p, str "->", pr_term tyvars thm vars' NOBR body] end; |
31665 | 112 |
in brackify_block fxy |
113 |
(concat [str "case", pr_term tyvars thm vars NOBR t, str "of", str "{"]) |
|
114 |
(map pr clauses) |
|
115 |
(str "}") |
|
28054 | 116 |
end |
31376 | 117 |
| pr_case tyvars thm vars fxy ((_, []), _) = |
118 |
(brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""]; |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
119 |
fun pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), []))) = |
28054 | 120 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
121 |
val tyvars = intro_vars (map fst vs) reserved; |
28054 | 122 |
val n = (length o fst o Code_Thingol.unfold_fun) ty; |
123 |
in |
|
124 |
Pretty.chunks [ |
|
125 |
Pretty.block [ |
|
126 |
(str o suffix " ::" o deresolve_base) name, |
|
127 |
Pretty.brk 1, |
|
128 |
pr_typscheme tyvars (vs, ty), |
|
129 |
str ";" |
|
130 |
], |
|
131 |
concat ( |
|
132 |
(str o deresolve_base) name |
|
133 |
:: map str (replicate n "_") |
|
134 |
@ str "=" |
|
135 |
:: str "error" |
|
136 |
@@ (str o (fn s => s ^ ";") o ML_Syntax.print_string |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
137 |
o Long_Name.base_name o Long_Name.qualifier) name |
28054 | 138 |
) |
139 |
] |
|
140 |
end |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
141 |
| pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) = |
28054 | 142 |
let |
28350 | 143 |
val eqs = filter (snd o snd) raw_eqs; |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
144 |
val tyvars = intro_vars (map fst vs) reserved; |
28350 | 145 |
fun pr_eq ((ts, t), (thm, _)) = |
28054 | 146 |
let |
32913 | 147 |
val consts = fold Code_Thingol.add_constnames (t :: ts) []; |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
148 |
val vars = reserved |
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
149 |
|> intro_base_names |
32913 | 150 |
(is_none o syntax_const) deresolve consts |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
151 |
|> intro_vars ((fold o Code_Thingol.fold_varnames) |
32925 | 152 |
(insert (op =)) ts []); |
28054 | 153 |
in |
154 |
semicolon ( |
|
155 |
(str o deresolve_base) name |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
156 |
:: map (pr_term tyvars thm vars BR) ts |
28054 | 157 |
@ str "=" |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
158 |
@@ pr_term tyvars thm vars NOBR t |
28054 | 159 |
) |
160 |
end; |
|
161 |
in |
|
162 |
Pretty.chunks ( |
|
163 |
Pretty.block [ |
|
164 |
(str o suffix " ::" o deresolve_base) name, |
|
165 |
Pretty.brk 1, |
|
166 |
pr_typscheme tyvars (vs, ty), |
|
167 |
str ";" |
|
168 |
] |
|
169 |
:: map pr_eq eqs |
|
170 |
) |
|
171 |
end |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
172 |
| pr_stmt (name, Code_Thingol.Datatype (_, (vs, []))) = |
28054 | 173 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
174 |
val tyvars = intro_vars (map fst vs) reserved; |
28054 | 175 |
in |
176 |
semicolon [ |
|
177 |
str "data", |
|
178 |
pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) |
|
179 |
] |
|
180 |
end |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
181 |
| pr_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) = |
28054 | 182 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
183 |
val tyvars = intro_vars (map fst vs) reserved; |
28054 | 184 |
in |
185 |
semicolon ( |
|
186 |
str "newtype" |
|
187 |
:: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) |
|
188 |
:: str "=" |
|
189 |
:: (str o deresolve_base) co |
|
190 |
:: pr_typ tyvars BR ty |
|
191 |
:: (if deriving_show name then [str "deriving (Read, Show)"] else []) |
|
192 |
) |
|
193 |
end |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
194 |
| pr_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) = |
28054 | 195 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
196 |
val tyvars = intro_vars (map fst vs) reserved; |
28054 | 197 |
fun pr_co (co, tys) = |
198 |
concat ( |
|
199 |
(str o deresolve_base) co |
|
200 |
:: map (pr_typ tyvars BR) tys |
|
201 |
) |
|
202 |
in |
|
203 |
semicolon ( |
|
204 |
str "data" |
|
205 |
:: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) |
|
206 |
:: str "=" |
|
207 |
:: pr_co co |
|
208 |
:: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos |
|
209 |
@ (if deriving_show name then [str "deriving (Read, Show)"] else []) |
|
210 |
) |
|
211 |
end |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
212 |
| pr_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) = |
28054 | 213 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
214 |
val tyvars = intro_vars [v] reserved; |
28054 | 215 |
fun pr_classparam (classparam, ty) = |
216 |
semicolon [ |
|
28687 | 217 |
(str o deresolve_base) classparam, |
28054 | 218 |
str "::", |
219 |
pr_typ tyvars NOBR ty |
|
220 |
] |
|
221 |
in |
|
222 |
Pretty.block_enclose ( |
|
223 |
Pretty.block [ |
|
224 |
str "class ", |
|
225 |
Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]), |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
226 |
str (deresolve_base name ^ " " ^ lookup_var tyvars v), |
28054 | 227 |
str " where {" |
228 |
], |
|
229 |
str "};" |
|
230 |
) (map pr_classparam classparams) |
|
231 |
end |
|
232 |
| pr_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) = |
|
233 |
let |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
234 |
val tyvars = intro_vars (map fst vs) reserved; |
28687 | 235 |
fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam |
236 |
of NONE => semicolon [ |
|
237 |
(str o deresolve_base) classparam, |
|
238 |
str "=", |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
239 |
pr_app tyvars thm reserved NOBR (c_inst, []) |
28687 | 240 |
] |
241 |
| SOME (k, pr) => |
|
242 |
let |
|
243 |
val (c_inst_name, (_, tys)) = c_inst; |
|
244 |
val const = if (is_some o syntax_const) c_inst_name |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
245 |
then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name; |
28687 | 246 |
val proto_rhs = Code_Thingol.eta_expand k (c_inst, []); |
31874 | 247 |
val (vs, rhs) = (apfst o map) fst (Code_Thingol.unfold_abs proto_rhs); |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
248 |
val vars = reserved |
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
249 |
|> intro_vars (the_list const) |
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
250 |
|> intro_vars (map_filter I vs); |
31889 | 251 |
val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs; |
28687 | 252 |
(*dictionaries are not relevant at this late stage*) |
253 |
in |
|
254 |
semicolon [ |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
255 |
pr_term tyvars thm vars NOBR lhs, |
28687 | 256 |
str "=", |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28687
diff
changeset
|
257 |
pr_term tyvars thm vars NOBR rhs |
28687 | 258 |
] |
259 |
end; |
|
28054 | 260 |
in |
261 |
Pretty.block_enclose ( |
|
262 |
Pretty.block [ |
|
263 |
str "instance ", |
|
264 |
Pretty.block (pr_typcontext tyvars vs), |
|
265 |
str (class_name class ^ " "), |
|
266 |
pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), |
|
267 |
str " where {" |
|
268 |
], |
|
269 |
str "};" |
|
270 |
) (map pr_instdef classparam_insts) |
|
271 |
end; |
|
272 |
in pr_stmt end; |
|
273 |
||
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
274 |
fun haskell_program_of_program labelled_name module_name module_prefix reserved raw_module_alias program = |
28054 | 275 |
let |
276 |
val module_alias = if is_some module_name then K module_name else raw_module_alias; |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
277 |
val reserved = Name.make_context reserved; |
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
278 |
val mk_name_module = mk_name_module reserved module_prefix module_alias program; |
28054 | 279 |
fun add_stmt (name, (stmt, deps)) = |
280 |
let |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
281 |
val (module_name, base) = dest_name name; |
28054 | 282 |
val module_name' = mk_name_module module_name; |
283 |
val mk_name_stmt = yield_singleton Name.variants; |
|
284 |
fun add_fun upper (nsp_fun, nsp_typ) = |
|
285 |
let |
|
286 |
val (base', nsp_fun') = |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
287 |
mk_name_stmt (if upper then first_upper base else base) nsp_fun |
28054 | 288 |
in (base', (nsp_fun', nsp_typ)) end; |
289 |
fun add_typ (nsp_fun, nsp_typ) = |
|
290 |
let |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
291 |
val (base', nsp_typ') = mk_name_stmt (first_upper base) nsp_typ |
28054 | 292 |
in (base', (nsp_fun, nsp_typ')) end; |
293 |
val add_name = case stmt |
|
294 |
of Code_Thingol.Fun _ => add_fun false |
|
295 |
| Code_Thingol.Datatype _ => add_typ |
|
296 |
| Code_Thingol.Datatypecons _ => add_fun true |
|
297 |
| Code_Thingol.Class _ => add_typ |
|
298 |
| Code_Thingol.Classrel _ => pair base |
|
299 |
| Code_Thingol.Classparam _ => add_fun false |
|
300 |
| Code_Thingol.Classinst _ => pair base; |
|
301 |
fun add_stmt' base' = case stmt |
|
302 |
of Code_Thingol.Datatypecons _ => |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
303 |
cons (name, (Long_Name.append module_name' base', NONE)) |
28054 | 304 |
| Code_Thingol.Classrel _ => I |
305 |
| Code_Thingol.Classparam _ => |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
306 |
cons (name, (Long_Name.append module_name' base', NONE)) |
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
307 |
| _ => cons (name, (Long_Name.append module_name' base', SOME stmt)); |
28054 | 308 |
in |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
309 |
Symtab.map_default (module_name', ([], ([], (reserved, reserved)))) |
28054 | 310 |
(apfst (fold (insert (op = : string * string -> bool)) deps)) |
311 |
#> `(fn program => add_name ((snd o snd o the o Symtab.lookup program) module_name')) |
|
312 |
#-> (fn (base', names) => |
|
313 |
(Symtab.map_entry module_name' o apsnd) (fn (stmts, _) => |
|
314 |
(add_stmt' base' stmts, names))) |
|
315 |
end; |
|
316 |
val hs_program = fold add_stmt (AList.make (fn name => |
|
317 |
(Graph.get_node program name, Graph.imm_succs program name)) |
|
318 |
(Graph.strong_conn program |> flat)) Symtab.empty; |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
319 |
fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
320 |
o Symtab.lookup hs_program) ((mk_name_module o fst o dest_name) name))) name |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
321 |
handle Option => error ("Unknown statement name: " ^ labelled_name name); |
28054 | 322 |
in (deresolver, hs_program) end; |
323 |
||
324 |
fun serialize_haskell module_prefix raw_module_name string_classes labelled_name |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
325 |
raw_reserved includes raw_module_alias |
31054 | 326 |
syntax_class syntax_tyco syntax_const program cs destination = |
28054 | 327 |
let |
328 |
val stmt_names = Code_Target.stmt_names_of_destination destination; |
|
329 |
val module_name = if null stmt_names then raw_module_name else SOME "Code"; |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
330 |
val reserved = fold (insert (op =) o fst) includes raw_reserved; |
28054 | 331 |
val (deresolver, hs_program) = haskell_program_of_program labelled_name |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
332 |
module_name module_prefix reserved raw_module_alias program; |
28054 | 333 |
val contr_classparam_typs = Code_Thingol.contr_classparam_typs program; |
334 |
fun deriving_show tyco = |
|
335 |
let |
|
336 |
fun deriv _ "fun" = false |
|
337 |
| deriv tycos tyco = member (op =) tycos tyco orelse |
|
338 |
case try (Graph.get_node program) tyco |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
339 |
of SOME (Code_Thingol.Datatype (_, (_, cs))) => forall (deriv' (tyco :: tycos)) |
28054 | 340 |
(maps snd cs) |
341 |
| NONE => true |
|
342 |
and deriv' tycos (tyco `%% tys) = deriv tycos tyco |
|
343 |
andalso forall (deriv' tycos) tys |
|
344 |
| deriv' _ (ITyVar _) = true |
|
345 |
in deriv [] tyco end; |
|
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
346 |
val reserved = make_vars reserved; |
31054 | 347 |
fun pr_stmt qualified = pr_haskell_stmt labelled_name |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
348 |
syntax_class syntax_tyco syntax_const reserved |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
349 |
(if qualified then deresolver else Long_Name.base_name o deresolver) |
32903 | 350 |
contr_classparam_typs |
28054 | 351 |
(if string_classes then deriving_show else K false); |
352 |
fun pr_module name content = |
|
353 |
(name, Pretty.chunks [ |
|
354 |
str ("module " ^ name ^ " where {"), |
|
355 |
str "", |
|
356 |
content, |
|
357 |
str "", |
|
358 |
str "}" |
|
359 |
]); |
|
360 |
fun serialize_module1 (module_name', (deps, (stmts, _))) = |
|
361 |
let |
|
362 |
val stmt_names = map fst stmts; |
|
33421
3789fe962a08
always be qualified -- suspected smartness in fact never worked as expected
haftmann
parents:
32925
diff
changeset
|
363 |
val qualified = is_none module_name; |
3789fe962a08
always be qualified -- suspected smartness in fact never worked as expected
haftmann
parents:
32925
diff
changeset
|
364 |
val imports = subtract (op =) stmt_names deps |
28054 | 365 |
|> distinct (op =) |
33421
3789fe962a08
always be qualified -- suspected smartness in fact never worked as expected
haftmann
parents:
32925
diff
changeset
|
366 |
|> map_filter (try deresolver) |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
367 |
|> map Long_Name.qualifier |
28054 | 368 |
|> distinct (op =); |
29794 | 369 |
fun pr_import_include (name, _) = str ("import qualified " ^ name ^ ";"); |
28054 | 370 |
val pr_import_module = str o (if qualified |
371 |
then prefix "import qualified " |
|
372 |
else prefix "import ") o suffix ";"; |
|
373 |
val content = Pretty.chunks ( |
|
374 |
map pr_import_include includes |
|
375 |
@ map pr_import_module imports |
|
376 |
@ str "" |
|
377 |
:: separate (str "") (map_filter |
|
378 |
(fn (name, (_, SOME stmt)) => SOME (pr_stmt qualified (name, stmt)) |
|
379 |
| (_, (_, NONE)) => NONE) stmts) |
|
380 |
) |
|
381 |
in pr_module module_name' content end; |
|
382 |
fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks ( |
|
383 |
separate (str "") (map_filter |
|
384 |
(fn (name, (_, SOME stmt)) => if null stmt_names |
|
385 |
orelse member (op =) stmt_names name |
|
386 |
then SOME (pr_stmt false (name, stmt)) |
|
387 |
else NONE |
|
388 |
| (_, (_, NONE)) => NONE) stmts)); |
|
389 |
val serialize_module = |
|
390 |
if null stmt_names then serialize_module1 else pair "" o serialize_module2; |
|
29832
b4919260eaec
check for destination directory, do not allocate it gratuitously
haftmann
parents:
29794
diff
changeset
|
391 |
fun check_destination destination = |
b4919260eaec
check for destination directory, do not allocate it gratuitously
haftmann
parents:
29794
diff
changeset
|
392 |
(File.check destination; destination); |
28054 | 393 |
fun write_module destination (modlname, content) = |
394 |
let |
|
395 |
val filename = case modlname |
|
396 |
of "" => Path.explode "Main.hs" |
|
397 |
| _ => (Path.ext "hs" o Path.explode o implode o separate "/" |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
398 |
o Long_Name.explode) modlname; |
28054 | 399 |
val pathname = Path.append destination filename; |
400 |
val _ = File.mkdir (Path.dir pathname); |
|
29135 | 401 |
in File.write pathname |
402 |
("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n" |
|
403 |
^ Code_Target.code_of_pretty content) |
|
404 |
end |
|
28054 | 405 |
in |
406 |
Code_Target.mk_serialization target NONE |
|
29832
b4919260eaec
check for destination directory, do not allocate it gratuitously
haftmann
parents:
29794
diff
changeset
|
407 |
(fn NONE => K () o map (Code_Target.code_writeln o snd) | SOME file => K () o map |
b4919260eaec
check for destination directory, do not allocate it gratuitously
haftmann
parents:
29794
diff
changeset
|
408 |
(write_module (check_destination file))) |
28054 | 409 |
(rpair [] o cat_lines o map (Code_Target.code_of_pretty o snd)) |
410 |
(map (uncurry pr_module) includes |
|
411 |
@ map serialize_module (Symtab.dest hs_program)) |
|
412 |
destination |
|
413 |
end; |
|
414 |
||
28064 | 415 |
val literals = let |
416 |
fun char_haskell c = |
|
417 |
let |
|
418 |
val s = ML_Syntax.print_char c; |
|
419 |
in if s = "'" then "\\'" else s end; |
|
420 |
in Literals { |
|
421 |
literal_char = enclose "'" "'" o char_haskell, |
|
422 |
literal_string = quote o translate_string char_haskell, |
|
423 |
literal_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k |
|
424 |
else enclose "(" ")" (signed_string_of_int k), |
|
425 |
literal_list = Pretty.enum "," "[" "]", |
|
426 |
infix_cons = (5, ":") |
|
427 |
} end; |
|
428 |
||
28054 | 429 |
|
430 |
(** optional monad syntax **) |
|
431 |
||
432 |
fun pretty_haskell_monad c_bind = |
|
433 |
let |
|
31874 | 434 |
fun dest_bind t1 t2 = case Code_Thingol.split_pat_abs t2 |
31889 | 435 |
of SOME ((pat, ty), t') => |
436 |
SOME ((SOME ((pat, ty), true), t1), t') |
|
28145 | 437 |
| NONE => NONE; |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
438 |
fun dest_monad c_bind_name (IConst (c, _) `$ t1 `$ t2) = |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
439 |
if c = c_bind_name then dest_bind t1 t2 |
28145 | 440 |
else NONE |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
441 |
| dest_monad _ t = case Code_Thingol.split_let t |
28145 | 442 |
of SOME (((pat, ty), tbind), t') => |
31889 | 443 |
SOME ((SOME ((pat, ty), false), tbind), t') |
28145 | 444 |
| NONE => NONE; |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
445 |
fun implode_monad c_bind_name = Code_Thingol.unfoldr (dest_monad c_bind_name); |
28145 | 446 |
fun pr_monad pr_bind pr (NONE, t) vars = |
447 |
(semicolon [pr vars NOBR t], vars) |
|
31889 | 448 |
| pr_monad pr_bind pr (SOME ((bind, _), true), t) vars = vars |
28145 | 449 |
|> pr_bind NOBR bind |
450 |
|>> (fn p => semicolon [p, str "<-", pr vars NOBR t]) |
|
31889 | 451 |
| pr_monad pr_bind pr (SOME ((bind, _), false), t) vars = vars |
28145 | 452 |
|> pr_bind NOBR bind |
453 |
|>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]); |
|
31054 | 454 |
fun pretty _ [c_bind'] pr thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2 |
28145 | 455 |
of SOME (bind, t') => let |
31054 | 456 |
val (binds, t'') = implode_monad c_bind' t' |
31889 | 457 |
val (ps, vars') = fold_map (pr_monad (gen_pr_bind (K pr) thm) pr) (bind :: binds) vars; |
28145 | 458 |
in (brackify fxy o single o Pretty.enclose "do {" "}" o Pretty.breaks) (ps @| pr vars' NOBR t'') end |
459 |
| NONE => brackify_infix (1, L) fxy |
|
460 |
[pr vars (INFX (1, L)) t1, str ">>=", pr vars (INFX (1, X)) t2] |
|
31054 | 461 |
in (2, ([c_bind], pretty)) end; |
28054 | 462 |
|
28145 | 463 |
fun add_monad target' raw_c_bind thy = |
28054 | 464 |
let |
31156 | 465 |
val c_bind = Code.read_const thy raw_c_bind; |
28054 | 466 |
in if target = target' then |
467 |
thy |
|
28145 | 468 |
|> Code_Target.add_syntax_const target c_bind |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28350
diff
changeset
|
469 |
(SOME (pretty_haskell_monad c_bind)) |
28054 | 470 |
else error "Only Haskell target allows for monad syntax" end; |
471 |
||
472 |
||
473 |
(** Isar setup **) |
|
474 |
||
475 |
fun isar_seri_haskell module = |
|
476 |
Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name) |
|
477 |
-- Scan.optional (Args.$$$ "string_classes" >> K true) false |
|
478 |
>> (fn (module_prefix, string_classes) => |
|
479 |
serialize_haskell module_prefix module string_classes)); |
|
480 |
||
481 |
val _ = |
|
482 |
OuterSyntax.command "code_monad" "define code syntax for monads" OuterKeyword.thy_decl ( |
|
28145 | 483 |
OuterParse.term_group -- OuterParse.name >> (fn (raw_bind, target) => |
484 |
Toplevel.theory (add_monad target raw_bind)) |
|
28054 | 485 |
); |
486 |
||
487 |
val setup = |
|
28064 | 488 |
Code_Target.add_target (target, (isar_seri_haskell, literals)) |
28054 | 489 |
#> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => |
490 |
brackify_infix (1, R) fxy [ |
|
491 |
pr_typ (INFX (1, X)) ty1, |
|
492 |
str "->", |
|
493 |
pr_typ (INFX (1, R)) ty2 |
|
494 |
])) |
|
495 |
#> fold (Code_Target.add_reserved target) [ |
|
496 |
"hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", |
|
497 |
"import", "default", "forall", "let", "in", "class", "qualified", "data", |
|
498 |
"newtype", "instance", "if", "then", "else", "type", "as", "do", "module" |
|
499 |
] |
|
500 |
#> fold (Code_Target.add_reserved target) [ |
|
501 |
"Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int", |
|
502 |
"Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded", |
|
503 |
"Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor", |
|
504 |
"AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException", |
|
505 |
"BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException", |
|
506 |
"Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure", |
|
507 |
"ExitSuccess", "False", "GT", "HeapOverflow", |
|
508 |
"IOError", "IOException", "IllegalOperation", |
|
509 |
"IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError", |
|
510 |
"NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow", |
|
511 |
"PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError", |
|
512 |
"RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow", |
|
513 |
"ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow", |
|
514 |
"UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all", |
|
515 |
"and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan", |
|
516 |
"atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen", |
|
517 |
"boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break", |
|
518 |
"catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const", |
|
519 |
"cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod", |
|
520 |
"doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem", |
|
521 |
"emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo", |
|
522 |
"enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip", |
|
523 |
"floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational", |
|
524 |
"floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble", |
|
525 |
"fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj", |
|
526 |
"fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head", |
|
527 |
"id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha", |
|
528 |
"isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite", |
|
529 |
"isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'", |
|
530 |
"last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log", |
|
531 |
"logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum", |
|
532 |
"maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem", |
|
533 |
"null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo", |
|
534 |
"numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred", |
|
535 |
"print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn", |
|
536 |
"quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat", |
|
537 |
"rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile", |
|
538 |
"readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn", |
|
539 |
"readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational", |
|
540 |
"realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse", |
|
541 |
"round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence", |
|
542 |
"sequence_", "show", "showChar", "showException", "showField", "showList", |
|
543 |
"showLitChar", "showParen", "showString", "shows", "showsPrec", "significand", |
|
544 |
"signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract", |
|
545 |
"succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult", |
|
546 |
"throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry", |
|
547 |
"undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords", |
|
548 |
"unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3" |
|
549 |
] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*); |
|
550 |
||
551 |
end; (*struct*) |