author | wenzelm |
Sun, 01 Oct 2006 18:30:04 +0200 | |
changeset 20818 | cb7ec413f95d |
parent 20699 | 0cc77abb185a |
child 20845 | c55dcf606f65 |
permissions | -rw-r--r-- |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
1 |
(* Title: Pure/Tools/codegen_serializer.ML |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
2 |
ID: $Id$ |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
3 |
Author: Florian Haftmann, TU Muenchen |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
4 |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
5 |
Serializer from intermediate language ("Thin-gol") to |
20699 | 6 |
target languages (like SML or Haskell). |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
7 |
*) |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
8 |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
9 |
signature CODEGEN_SERIALIZER = |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
10 |
sig |
20456 | 11 |
include BASIC_CODEGEN_THINGOL; |
20699 | 12 |
val parse_serializer: string |
13 |
-> OuterParse.token list |
|
14 |
-> (theory -> string list option -> CodegenThingol.module -> unit) |
|
15 |
* OuterParse.token list |
|
16 |
val eval_verbose: bool ref; |
|
17 |
val eval_term: theory -> |
|
18 |
(string * 'a option ref) * CodegenThingol.iterm -> CodegenThingol.module |
|
19 |
-> 'a; |
|
19150 | 20 |
val mk_flat_ml_resolver: string list -> string -> string; |
20699 | 21 |
val ml_fun_datatype: |
22 |
theory -> (string -> string) |
|
20456 | 23 |
-> ((string * ((iterm list * iterm) list * CodegenThingol.typscheme)) list -> Pretty.T) |
24 |
* ((string * ((vname * sort) list * (string * itype list) list)) list -> Pretty.T); |
|
20699 | 25 |
val const_has_serialization: theory -> string list -> string -> bool; |
26 |
val tyco_has_serialization: theory -> string list -> string -> bool; |
|
27 |
val check_serializer: string -> unit; |
|
28 |
val add_syntax_class: |
|
29 |
string -> string -> string * (string * string) list -> theory -> theory; |
|
30 |
val add_syntax_inst: string -> (string * string) -> theory -> theory; |
|
31 |
val parse_syntax_tyco: (theory |
|
32 |
-> CodegenConsts.const list * (string * typ) list |
|
33 |
-> string |
|
34 |
-> CodegenNames.tyco |
|
35 |
-> typ list -> CodegenThingol.itype list) |
|
36 |
-> Symtab.key |
|
37 |
-> xstring |
|
38 |
-> OuterParse.token list |
|
39 |
-> (theory -> theory) * OuterParse.token list; |
|
40 |
val parse_syntax_const: (theory |
|
41 |
-> CodegenConsts.const list * (string * typ) list |
|
42 |
-> string |
|
43 |
-> CodegenNames.const |
|
44 |
-> term list -> CodegenThingol.iterm list) |
|
45 |
-> Symtab.key |
|
46 |
-> string |
|
47 |
-> OuterParse.token list |
|
48 |
-> (theory -> theory) * OuterParse.token list; |
|
49 |
val add_pretty_list: string -> string -> string -> (Pretty.T list -> Pretty.T) |
|
50 |
-> ((string -> string) * (string -> string)) option -> int * string |
|
51 |
-> theory -> theory; |
|
52 |
val add_pretty_ml_string: string -> string -> string -> string |
|
53 |
-> (string -> string) -> (string -> string) -> string -> theory -> theory; |
|
54 |
val add_undefined: string -> string -> string -> theory -> theory; |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
55 |
end; |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
56 |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
57 |
structure CodegenSerializer: CODEGEN_SERIALIZER = |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
58 |
struct |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
59 |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
60 |
open BasicCodegenThingol; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
61 |
val debug_msg = CodegenThingol.debug_msg; |
18850 | 62 |
|
20699 | 63 |
(** precedences **) |
18702 | 64 |
|
18216 | 65 |
datatype lrx = L | R | X; |
66 |
||
18516 | 67 |
datatype fixity = |
18216 | 68 |
BR |
69 |
| NOBR |
|
70 |
| INFX of (int * lrx); |
|
71 |
||
20699 | 72 |
type 'a pretty_syntax = int * (fixity -> (fixity -> 'a -> Pretty.T) |
18865 | 73 |
-> 'a list -> Pretty.T); |
18516 | 74 |
|
18216 | 75 |
fun eval_lrx L L = false |
76 |
| eval_lrx R R = false |
|
77 |
| eval_lrx _ _ = true; |
|
78 |
||
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
79 |
fun eval_fxy NOBR _ = false |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
80 |
| eval_fxy _ BR = true |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
81 |
| eval_fxy _ NOBR = false |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
82 |
| eval_fxy (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
83 |
pr < pr_ctxt |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
84 |
orelse pr = pr_ctxt |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
85 |
andalso eval_lrx lr lr_ctxt |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
86 |
| eval_fxy _ (INFX _) = false; |
18216 | 87 |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
88 |
fun gen_brackify _ [p] = p |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
89 |
| gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
90 |
| gen_brackify false (ps as _::_) = Pretty.block ps; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
91 |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
92 |
fun brackify fxy_ctxt ps = |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
93 |
gen_brackify (eval_fxy BR fxy_ctxt) (Pretty.breaks ps); |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
94 |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
95 |
fun brackify_infix infx fxy_ctxt ps = |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
96 |
gen_brackify (eval_fxy (INFX infx) fxy_ctxt) (Pretty.breaks ps); |
18216 | 97 |
|
20699 | 98 |
fun from_app mk_app from_expr const_syntax fxy (const as (c, (_, ty)), es) = |
20191 | 99 |
case const_syntax c |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
100 |
of NONE => brackify fxy (mk_app c es) |
20699 | 101 |
| SOME (i, pr) => |
102 |
let |
|
103 |
val k = if i < 0 then (length o fst o CodegenThingol.unfold_fun) ty else i |
|
104 |
in if k <= length es |
|
105 |
then case chop i es of (es1, es2) => |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
106 |
brackify fxy (pr fxy from_expr es1 :: map (from_expr BR) es2) |
20699 | 107 |
else from_expr fxy (CodegenThingol.eta_expand (const, es) i) |
108 |
end; |
|
109 |
||
110 |
||
111 |
(** user-defined syntax **) |
|
112 |
||
113 |
(* theory data *) |
|
114 |
||
115 |
type target_data = { |
|
116 |
syntax_class: ((string * (string -> string option)) * stamp) Symtab.table, |
|
117 |
syntax_inst: unit Symtab.table, |
|
118 |
syntax_tyco: (itype pretty_syntax * stamp) Symtab.table, |
|
119 |
syntax_const: (iterm pretty_syntax * stamp) Symtab.table |
|
120 |
}; |
|
121 |
||
122 |
fun merge_target_data |
|
123 |
({ syntax_class = syntax_class1, syntax_inst = syntax_inst1, |
|
124 |
syntax_tyco = syntax_tyco1, syntax_const = syntax_const1 }, |
|
125 |
{ syntax_class = syntax_class2, syntax_inst = syntax_inst2, |
|
126 |
syntax_tyco = syntax_tyco2, syntax_const = syntax_const2 }) = |
|
127 |
{ syntax_class = Symtab.merge (eq_snd (op =)) (syntax_class1, syntax_class2), |
|
128 |
syntax_inst = Symtab.merge (op =) (syntax_inst1, syntax_inst2), |
|
129 |
syntax_tyco = Symtab.merge (eq_snd (op =)) (syntax_tyco1, syntax_tyco2), |
|
130 |
syntax_const = Symtab.merge (eq_snd (op =)) (syntax_const1, syntax_const2) } : target_data; |
|
131 |
||
132 |
structure CodegenSerializerData = TheoryDataFun |
|
133 |
(struct |
|
134 |
val name = "Pure/codegen_serializer"; |
|
135 |
type T = target_data Symtab.table; |
|
136 |
val empty = |
|
137 |
Symtab.empty |
|
138 |
|> fold (fn target => |
|
139 |
Symtab.update (target, |
|
140 |
{ syntax_class = Symtab.empty, syntax_inst = Symtab.empty, |
|
141 |
syntax_tyco = Symtab.empty, syntax_const = Symtab.empty }) |
|
142 |
) ["SML", "Haskell"] : T; |
|
143 |
val copy = I; |
|
144 |
val extend = I; |
|
145 |
fun merge _ = Symtab.join (K merge_target_data); |
|
146 |
fun print _ _ = (); |
|
147 |
end); |
|
148 |
||
149 |
fun has_serialization f thy targets name = |
|
150 |
forall ( |
|
151 |
is_some o (fn tab => Symtab.lookup tab name) o f o the |
|
152 |
o (Symtab.lookup ((CodegenSerializerData.get) thy)) |
|
153 |
) targets; |
|
154 |
||
155 |
val tyco_has_serialization = has_serialization #syntax_tyco; |
|
156 |
val const_has_serialization = has_serialization #syntax_const; |
|
157 |
||
158 |
fun serialize thy seri target cs = |
|
159 |
let |
|
160 |
val data = CodegenSerializerData.get thy; |
|
161 |
val target_data = |
|
162 |
(the oo Symtab.lookup) data target; |
|
163 |
val syntax_class = #syntax_class target_data; |
|
164 |
val syntax_inst = #syntax_inst target_data; |
|
165 |
val syntax_tyco = #syntax_tyco target_data; |
|
166 |
val syntax_const = #syntax_const target_data; |
|
167 |
fun fun_of syntax = (Option.map fst oo Symtab.lookup) syntax; |
|
168 |
in |
|
169 |
seri (fun_of syntax_class, fun_of syntax_tyco, fun_of syntax_const) |
|
170 |
(Symtab.keys syntax_class @ Symtab.keys syntax_inst |
|
171 |
@ Symtab.keys syntax_tyco @ Symtab.keys syntax_const, cs) |
|
172 |
end; |
|
173 |
||
174 |
val _ = Context.add_setup CodegenSerializerData.init; |
|
175 |
||
176 |
val (atomK, infixK, infixlK, infixrK) = |
|
177 |
("target_atom", "infix", "infixl", "infixr"); |
|
178 |
val _ = OuterSyntax.add_keywords [atomK, infixK, infixlK, infixrK]; |
|
179 |
||
180 |
||
181 |
(* syntax parser *) |
|
182 |
||
183 |
val str = setmp print_mode [] Pretty.str; |
|
184 |
||
185 |
datatype 'a mixfix = |
|
186 |
Arg of fixity |
|
187 |
| Pretty of Pretty.T |
|
188 |
| Quote of 'a; |
|
18865 | 189 |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
190 |
fun fillin_mixfix fxy_this ms fxy_ctxt pr args = |
18702 | 191 |
let |
192 |
fun fillin [] [] = |
|
19008 | 193 |
[] |
18702 | 194 |
| fillin (Arg fxy :: ms) (a :: args) = |
195 |
pr fxy a :: fillin ms args |
|
196 |
| fillin (Pretty p :: ms) args = |
|
197 |
p :: fillin ms args |
|
198 |
| fillin (Quote q :: ms) args = |
|
19008 | 199 |
pr BR q :: fillin ms args |
200 |
| fillin [] _ = |
|
20389 | 201 |
error ("Inconsistent mixfix: too many arguments") |
19008 | 202 |
| fillin _ [] = |
20389 | 203 |
error ("Inconsistent mixfix: too less arguments"); |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
204 |
in gen_brackify (eval_fxy fxy_this fxy_ctxt) (fillin ms args) end; |
18702 | 205 |
|
206 |
fun parse_infix (fixity as INFX (i, x)) s = |
|
207 |
let |
|
208 |
val l = case x of L => fixity |
|
209 |
| _ => INFX (i, X); |
|
210 |
val r = case x of R => fixity |
|
211 |
| _ => INFX (i, X); |
|
212 |
in |
|
213 |
pair [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r] |
|
214 |
end; |
|
215 |
||
216 |
fun parse_mixfix reader s ctxt = |
|
18335 | 217 |
let |
18702 | 218 |
fun sym s = Scan.lift ($$ s); |
219 |
fun lift_reader ctxt s = |
|
220 |
ctxt |
|
221 |
|> reader s |
|
222 |
|-> (fn x => pair (Quote x)); |
|
223 |
val sym_any = Scan.lift (Scan.one Symbol.not_eof); |
|
224 |
val parse = Scan.repeat ( |
|
225 |
(sym "_" -- sym "_" >> K (Arg NOBR)) |
|
226 |
|| (sym "_" >> K (Arg BR)) |
|
227 |
|| (sym "/" |-- Scan.repeat (sym " ") >> (Pretty o Pretty.brk o length)) |
|
228 |
|| Scan.depend (fn ctxt => $$ "{" |-- $$ "*" |-- Scan.repeat1 |
|
229 |
( $$ "'" |-- Scan.one Symbol.not_eof |
|
230 |
|| Scan.unless ($$ "*" -- $$ "}") (Scan.one Symbol.not_eof)) --| |
|
231 |
$$ "*" --| $$ "}" >> (implode #> lift_reader ctxt #> swap)) |
|
232 |
|| (Scan.repeat1 |
|
233 |
( sym "'" |-- sym_any |
|
234 |
|| Scan.unless (sym "_" || sym "?" || sym "/" || sym "{" |-- sym "*") |
|
235 |
sym_any) >> (Pretty o str o implode))); |
|
236 |
in case Scan.finite' Symbol.stopper parse (ctxt, Symbol.explode s) |
|
237 |
of (p, (ctxt, [])) => (p, ctxt) |
|
238 |
| _ => error ("Malformed mixfix annotation: " ^ quote s) |
|
239 |
end; |
|
240 |
||
20439 | 241 |
fun parse_syntax num_args reader = |
18335 | 242 |
let |
18702 | 243 |
fun is_arg (Arg _) = true |
244 |
| is_arg _ = false; |
|
20439 | 245 |
fun parse_nonatomic s ctxt = |
246 |
case parse_mixfix reader s ctxt |
|
247 |
of ([Pretty _], _) => |
|
248 |
error ("Mixfix contains just one pretty element; either declare as " |
|
249 |
^ quote atomK ^ " or consider adding a break") |
|
250 |
| x => x; |
|
20456 | 251 |
val parse = ( |
20439 | 252 |
OuterParse.$$$ infixK |-- OuterParse.nat |
253 |
>> (fn i => (parse_infix (INFX (i, X)), INFX (i, X))) |
|
254 |
|| OuterParse.$$$ infixlK |-- OuterParse.nat |
|
255 |
>> (fn i => (parse_infix (INFX (i, L)), INFX (i, L))) |
|
256 |
|| OuterParse.$$$ infixrK |-- OuterParse.nat |
|
257 |
>> (fn i => (parse_infix (INFX (i, R)), INFX (i, R))) |
|
258 |
|| OuterParse.$$$ atomK |-- pair (parse_mixfix reader, NOBR) |
|
259 |
|| pair (parse_nonatomic, BR) |
|
20456 | 260 |
) -- OuterParse.string >> (fn ((p, fxy), s) => (p s, fxy)); |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
261 |
fun mk fixity mfx ctxt = |
18702 | 262 |
let |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
263 |
val i = (length o List.filter is_arg) mfx; |
20600 | 264 |
val _ = if i > num_args ctxt then error "Too many arguments in code syntax" else (); |
20699 | 265 |
in ((i, fillin_mixfix fixity mfx), ctxt) end; |
18702 | 266 |
in |
20439 | 267 |
parse |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
268 |
#-> (fn (mfx_reader, fixity) => |
20699 | 269 |
pair (mfx_reader #-> (fn mfx => (mk fixity mfx))) |
18702 | 270 |
) |
271 |
end; |
|
272 |
||
273 |
||
20699 | 274 |
|
275 |
(** generic abstract serializer **) |
|
18282 | 276 |
|
20699 | 277 |
val nsp_module = CodegenNames.nsp_module; |
278 |
val nsp_class = CodegenNames.nsp_class; |
|
279 |
val nsp_tyco = CodegenNames.nsp_tyco; |
|
280 |
val nsp_inst = CodegenNames.nsp_inst; |
|
281 |
val nsp_fun = CodegenNames.nsp_fun; |
|
282 |
val nsp_classop = CodegenNames.nsp_classop; |
|
283 |
val nsp_dtco = CodegenNames.nsp_dtco; |
|
284 |
val nsp_eval = CodegenNames.nsp_eval; |
|
18702 | 285 |
|
18919 | 286 |
fun abstract_serializer (target, nspgrp) name_root (from_defs, from_module, validator, postproc) |
20699 | 287 |
postprocess (class_syntax, tyco_syntax, const_syntax) |
288 |
(drop: string list, select) code = |
|
18702 | 289 |
let |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
290 |
fun from_module' resolv imps ((name_qual, name), defs) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
291 |
from_module resolv imps ((name_qual, name), defs) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
292 |
|> postprocess (resolv name_qual); |
18702 | 293 |
in |
20699 | 294 |
code |
20191 | 295 |
|> debug_msg (fn _ => "dropping shadowed defintions...") |
296 |
|> CodegenThingol.delete_garbage drop |
|
20699 | 297 |
|> debug_msg (fn _ => "projecting...") |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
298 |
|> (if is_some select then (CodegenThingol.project_module o the) select else I) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
299 |
|> debug_msg (fn _ => "serializing...") |
20428 | 300 |
|> CodegenThingol.serialize (from_defs (class_syntax, tyco_syntax, const_syntax)) |
18919 | 301 |
from_module' validator postproc nspgrp name_root |
18850 | 302 |
|> K () |
18702 | 303 |
end; |
304 |
||
305 |
fun abstract_validator keywords name = |
|
306 |
let |
|
20699 | 307 |
fun replace_invalid c = (*FIXME*) |
20203 | 308 |
if Symbol.is_ascii_letter c orelse Symbol.is_ascii_digit c orelse c = "'" |
18702 | 309 |
andalso not (NameSpace.separator = c) |
310 |
then c |
|
311 |
else "_" |
|
20191 | 312 |
fun suffix_it name= |
18702 | 313 |
name |
314 |
|> member (op =) keywords ? suffix "'" |
|
315 |
|> (fn name' => if name = name' then name else suffix_it name') |
|
316 |
in |
|
317 |
name |
|
318 |
|> translate_string replace_invalid |
|
319 |
|> suffix_it |
|
320 |
|> (fn name' => if name = name' then NONE else SOME name') |
|
321 |
end; |
|
322 |
||
18850 | 323 |
fun write_file mkdir path p = ( |
324 |
if mkdir |
|
325 |
then |
|
326 |
File.mkdir (Path.dir path) |
|
327 |
else (); |
|
328 |
File.write path (Pretty.output p ^ "\n"); |
|
329 |
p |
|
330 |
); |
|
331 |
||
332 |
fun mk_module_file postprocess_module ext path name p = |
|
333 |
let |
|
334 |
val prfx = Path.dir path; |
|
335 |
val name' = case name |
|
336 |
of "" => Path.base path |
|
337 |
| _ => (Path.ext ext o Path.unpack o implode o separate "/" o NameSpace.unpack) name; |
|
338 |
in |
|
339 |
p |
|
340 |
|> write_file true (Path.append prfx name') |
|
341 |
|> postprocess_module name |
|
342 |
end; |
|
343 |
||
18756 | 344 |
fun parse_single_file serializer = |
18850 | 345 |
OuterParse.path |
346 |
>> (fn path => serializer |
|
347 |
(fn "" => write_file false path #> K NONE |
|
348 |
| _ => SOME)); |
|
349 |
||
350 |
fun parse_multi_file postprocess_module ext serializer = |
|
351 |
OuterParse.path |
|
352 |
>> (fn path => (serializer o mk_module_file postprocess_module ext) path); |
|
18702 | 353 |
|
18756 | 354 |
fun parse_internal serializer = |
355 |
OuterParse.name |
|
18850 | 356 |
>> (fn "-" => serializer |
357 |
(fn "" => (fn p => (use_text Context.ml_output false (Pretty.output p); NONE)) |
|
358 |
| _ => SOME) |
|
18756 | 359 |
| _ => Scan.fail ()); |
18702 | 360 |
|
20105 | 361 |
fun parse_stdout serializer = |
362 |
OuterParse.name |
|
363 |
>> (fn "_" => serializer |
|
364 |
(fn "" => (fn p => (Pretty.writeln p; NONE)) |
|
365 |
| _ => SOME) |
|
366 |
| _ => Scan.fail ()); |
|
18282 | 367 |
|
20699 | 368 |
fun constructive_fun (name, (eqs, ty)) = |
369 |
let |
|
370 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
|
371 |
fun check_eq (eq as (lhs, rhs)) = |
|
372 |
if forall (CodegenThingol.is_pat is_cons) lhs |
|
373 |
then SOME eq |
|
374 |
else (warning ("In function " ^ quote name ^ ", throwing away one " |
|
375 |
^ "non-executable function clause"); NONE) |
|
376 |
in case map_filter check_eq eqs |
|
377 |
of [] => error ("In function " ^ quote name ^ ", no " |
|
378 |
^ "executable function clauses found") |
|
379 |
| eqs => (name, (eqs, ty)) |
|
380 |
end; |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
381 |
|
20699 | 382 |
fun make_ctxt names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty, |
383 |
Name.make_context names); |
|
384 |
||
385 |
fun intro_ctxt names (namemap, namectxt) = |
|
386 |
let |
|
387 |
val (names', namectxt') = Name.variants names namectxt; |
|
388 |
val namemap' = fold2 (curry Symtab.update) names names' namemap; |
|
389 |
in (namemap', namectxt') end; |
|
390 |
||
391 |
fun lookup_ctxt (namemap, _) name = case Symtab.lookup namemap name |
|
392 |
of SOME name' => name' |
|
393 |
| NONE => error ("invalid name in context: " ^ quote name); |
|
394 |
||
395 |
||
396 |
||
397 |
(** generic list and string serializers **) |
|
20401 | 398 |
|
399 |
fun implode_list c_nil c_cons e = |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
400 |
let |
19202 | 401 |
fun dest_cons (IConst (c, _) `$ e1 `$ e2) = |
20401 | 402 |
if c = c_cons |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
403 |
then SOME (e1, e2) |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
404 |
else NONE |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
405 |
| dest_cons _ = NONE; |
20401 | 406 |
val (es, e') = CodegenThingol.unfoldr dest_cons e; |
407 |
in case e' |
|
408 |
of IConst (c, _) => if c = c_nil then SOME es else NONE |
|
409 |
| _ => NONE |
|
410 |
end; |
|
411 |
||
412 |
fun implode_string mk_char mk_string es = |
|
413 |
if forall (fn IChar _ => true | _ => false) es |
|
414 |
then (SOME o str o mk_string o implode o map (fn IChar (c, _) => mk_char c)) es |
|
415 |
else NONE; |
|
416 |
||
417 |
fun pretty_ml_string c_nil c_cons mk_char mk_string target_implode = |
|
418 |
let |
|
419 |
fun pretty fxy pr [e] = |
|
420 |
case implode_list c_nil c_cons e |
|
421 |
of SOME es => (case implode_string mk_char mk_string es |
|
422 |
of SOME p => p |
|
423 |
| NONE => Pretty.block [str target_implode, Pretty.brk 1, pr BR e]) |
|
424 |
| NONE => Pretty.block [str target_implode, Pretty.brk 1, pr BR e] |
|
20699 | 425 |
in (1, pretty) end; |
20401 | 426 |
|
427 |
fun pretty_list c_nil c_cons mk_list mk_char_string (target_fxy, target_cons) = |
|
428 |
let |
|
429 |
fun default fxy pr e1 e2 = |
|
430 |
brackify_infix (target_fxy, R) fxy [ |
|
431 |
pr (INFX (target_fxy, X)) e1, |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
432 |
str target_cons, |
20401 | 433 |
pr (INFX (target_fxy, R)) e2 |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
434 |
]; |
20401 | 435 |
fun pretty fxy pr [e1, e2] = |
436 |
case Option.map (cons e1) (implode_list c_nil c_cons e2) |
|
437 |
of SOME es => |
|
438 |
(case mk_char_string |
|
439 |
of SOME (mk_char, mk_string) => |
|
440 |
(case implode_string mk_char mk_string es |
|
441 |
of SOME p => p |
|
442 |
| NONE => mk_list (map (pr NOBR) es)) |
|
443 |
| NONE => mk_list (map (pr NOBR) es)) |
|
444 |
| NONE => default fxy pr e1 e2; |
|
20699 | 445 |
in (2, pretty) end; |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
446 |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
447 |
|
18216 | 448 |
|
449 |
(** ML serializer **) |
|
450 |
||
451 |
local |
|
452 |
||
20699 | 453 |
val reserved_ml' = [ |
20401 | 454 |
"bool", "int", "list", "unit", "option", "true", "false", "not", "NONE", "SOME", |
455 |
"o", "string", "char", "String", "Term" |
|
456 |
]; |
|
19150 | 457 |
|
20699 | 458 |
fun ml_expr_seri (tyco_syntax, const_syntax) resolv = |
18216 | 459 |
let |
20699 | 460 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
461 |
fun first_upper s = |
|
462 |
implode (nth_map 0 (Symbol.to_ascii_upper) (explode s)); |
|
463 |
fun ml_from_dictvar v = |
|
464 |
first_upper v ^ "_"; |
|
18865 | 465 |
val ml_from_label = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
466 |
str o translate_string (fn "_" => "__" | "." => "_" | c => c) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
467 |
o NameSpace.base o resolv; |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
468 |
fun ml_from_tyvar (v, []) = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
469 |
str "()" |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
470 |
| ml_from_tyvar (v, sort) = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
471 |
let |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
472 |
fun mk_class class = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
473 |
str (prefix "'" v ^ " " ^ resolv class); |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
474 |
in |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
475 |
Pretty.block [ |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
476 |
str "(", |
20428 | 477 |
(str o ml_from_dictvar) v, |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
478 |
str ":", |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
479 |
case sort |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
480 |
of [class] => mk_class class |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
481 |
| _ => Pretty.enum " *" "" "" (map mk_class sort), |
20191 | 482 |
str ")" |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
483 |
] |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
484 |
end; |
20456 | 485 |
fun ml_from_insts fxy lss = |
18885 | 486 |
let |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
487 |
fun from_label l = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
488 |
Pretty.block [str "#", |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
489 |
if (is_some o Int.fromString) l then str l |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
490 |
else ml_from_label l |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
491 |
]; |
20401 | 492 |
fun from_lookup fxy [] p = |
493 |
p |
|
494 |
| from_lookup fxy [l] p = |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
495 |
brackify fxy [ |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
496 |
from_label l, |
20401 | 497 |
p |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
498 |
] |
20401 | 499 |
| from_lookup fxy ls p = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
500 |
brackify fxy [ |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
501 |
Pretty.enum " o" "(" ")" (map from_label ls), |
20401 | 502 |
p |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
503 |
]; |
20456 | 504 |
fun from_inst fxy (Instance (inst, lss)) = |
18885 | 505 |
brackify fxy ( |
506 |
(str o resolv) inst |
|
20456 | 507 |
:: map (ml_from_insts BR) lss |
18885 | 508 |
) |
20456 | 509 |
| from_inst fxy (Context (classes, (v, ~1))) = |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
510 |
from_lookup BR classes |
20428 | 511 |
((str o ml_from_dictvar) v) |
20456 | 512 |
| from_inst fxy (Context (classes, (v, i))) = |
20600 | 513 |
from_lookup BR (classes @ [string_of_int (i+1)]) |
20428 | 514 |
((str o ml_from_dictvar) v) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
515 |
in case lss |
19253 | 516 |
of [] => str "()" |
20456 | 517 |
| [ls] => from_inst fxy ls |
518 |
| lss => (Pretty.list "(" ")" o map (from_inst NOBR)) lss |
|
18885 | 519 |
end; |
18963 | 520 |
fun ml_from_tycoexpr fxy (tyco, tys) = |
521 |
let |
|
522 |
val tyco' = (str o resolv) tyco |
|
523 |
in case map (ml_from_type BR) tys |
|
524 |
of [] => tyco' |
|
525 |
| [p] => Pretty.block [p, Pretty.brk 1, tyco'] |
|
526 |
| (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco'] |
|
527 |
end |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
528 |
and ml_from_type fxy (tycoexpr as tyco `%% tys) = |
18702 | 529 |
(case tyco_syntax tyco |
18963 | 530 |
of NONE => ml_from_tycoexpr fxy (tyco, tys) |
20699 | 531 |
| SOME (i, pr) => |
532 |
if not (i = length tys) |
|
20389 | 533 |
then error ("Number of argument mismatch in customary serialization: " |
18865 | 534 |
^ (string_of_int o length) tys ^ " given, " |
20699 | 535 |
^ string_of_int i ^ " expected") |
18702 | 536 |
else pr fxy ml_from_type tys) |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
537 |
| ml_from_type fxy (t1 `-> t2) = |
18216 | 538 |
let |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
539 |
val brackify = gen_brackify |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
540 |
(case fxy |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
541 |
of BR => false |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
542 |
| _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks; |
18216 | 543 |
in |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
544 |
brackify [ |
18702 | 545 |
ml_from_type (INFX (1, X)) t1, |
546 |
str "->", |
|
547 |
ml_from_type (INFX (1, R)) t2 |
|
548 |
] |
|
18216 | 549 |
end |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
550 |
| ml_from_type fxy (ITyVar v) = |
18885 | 551 |
str ("'" ^ v); |
20699 | 552 |
fun ml_from_expr var_ctxt fxy (e as IConst x) = |
553 |
ml_from_app var_ctxt fxy (x, []) |
|
554 |
| ml_from_expr var_ctxt fxy (IVar v) = |
|
555 |
str (lookup_ctxt var_ctxt v) |
|
556 |
| ml_from_expr var_ctxt fxy (e as e1 `$ e2) = |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
557 |
(case CodegenThingol.unfold_const_app e |
20699 | 558 |
of SOME x => ml_from_app var_ctxt fxy x |
18865 | 559 |
| NONE => |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
560 |
brackify fxy [ |
20699 | 561 |
ml_from_expr var_ctxt NOBR e1, |
562 |
ml_from_expr var_ctxt BR e2 |
|
18216 | 563 |
]) |
20699 | 564 |
| ml_from_expr var_ctxt fxy (e as _ `|-> _) = |
20105 | 565 |
let |
20699 | 566 |
val (es, e') = CodegenThingol.unfold_abs e; |
567 |
val vs = fold CodegenThingol.add_varnames (map fst es) []; |
|
568 |
val var_ctxt' = intro_ctxt vs var_ctxt; |
|
20105 | 569 |
fun mk_abs (e, ty) = (Pretty.block o Pretty.breaks) [ |
570 |
str "fn", |
|
20699 | 571 |
ml_from_expr var_ctxt' NOBR e, |
20105 | 572 |
str "=>" |
573 |
]; |
|
20699 | 574 |
in brackify BR (map mk_abs es @ [ml_from_expr var_ctxt' NOBR e']) end |
575 |
| ml_from_expr var_ctxt fxy (INum (n, _)) = |
|
19213 | 576 |
brackify BR [ |
19202 | 577 |
(str o IntInf.toString) n, |
19884 | 578 |
str ":", |
579 |
str "IntInf.int" |
|
19202 | 580 |
] |
20699 | 581 |
| ml_from_expr var_ctxt _ (IChar (c, _)) = |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
582 |
(str o prefix "#" o quote) |
20203 | 583 |
(let val i = ord c |
20191 | 584 |
in if i < 32 |
20105 | 585 |
then prefix "\\" (string_of_int i) |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
586 |
else c |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
587 |
end) |
20699 | 588 |
| ml_from_expr var_ctxt fxy (e as ICase ((_, [_]), _)) = |
18216 | 589 |
let |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
590 |
val (ves, be) = CodegenThingol.unfold_let e; |
20699 | 591 |
fun mk_val ((ve, vty), se) var_ctxt = |
592 |
let |
|
593 |
val vs = CodegenThingol.add_varnames ve []; |
|
594 |
val var_ctxt' = intro_ctxt vs var_ctxt; |
|
595 |
in |
|
596 |
(Pretty.block [ |
|
597 |
(Pretty.block o Pretty.breaks) [ |
|
598 |
str "val", |
|
599 |
ml_from_expr var_ctxt' NOBR ve, |
|
600 |
str "=", |
|
601 |
ml_from_expr var_ctxt NOBR se |
|
602 |
], |
|
603 |
str ";" |
|
604 |
], var_ctxt') |
|
605 |
end |
|
606 |
val (binds, var_ctxt') = fold_map mk_val ves var_ctxt; |
|
18216 | 607 |
in Pretty.chunks [ |
20699 | 608 |
[str ("let"), Pretty.fbrk, binds |> Pretty.chunks] |> Pretty.block, |
609 |
[str ("in"), Pretty.fbrk, ml_from_expr var_ctxt' NOBR be] |> Pretty.block, |
|
18702 | 610 |
str ("end") |
18216 | 611 |
] end |
20699 | 612 |
| ml_from_expr var_ctxt fxy (ICase (((de, dty), bse::bses), _)) = |
18216 | 613 |
let |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
614 |
fun mk_clause definer (se, be) = |
20699 | 615 |
let |
616 |
val vs = CodegenThingol.add_varnames se []; |
|
617 |
val var_ctxt' = intro_ctxt vs var_ctxt; |
|
618 |
in |
|
619 |
(Pretty.block o Pretty.breaks) [ |
|
620 |
str definer, |
|
621 |
ml_from_expr var_ctxt' NOBR se, |
|
622 |
str "=>", |
|
623 |
ml_from_expr var_ctxt' NOBR be |
|
624 |
] |
|
625 |
end |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
626 |
in brackify fxy ( |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
627 |
str "(case" |
20699 | 628 |
:: ml_from_expr var_ctxt NOBR de |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
629 |
:: mk_clause "of" bse |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
630 |
:: map (mk_clause "|") bses |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
631 |
@ [str ")"] |
18216 | 632 |
) end |
20699 | 633 |
| ml_from_expr var_ctxt _ e = |
20389 | 634 |
error ("Dubious expression: " ^ (Pretty.output o CodegenThingol.pretty_iterm) e) |
20699 | 635 |
and ml_mk_app var_ctxt f es = |
19136 | 636 |
if is_cons f andalso length es > 1 then |
20699 | 637 |
[(str o resolv) f, Pretty.enum "," "(" ")" (map (ml_from_expr var_ctxt BR) es)] |
18702 | 638 |
else |
20699 | 639 |
(str o resolv) f :: map (ml_from_expr var_ctxt BR) es |
640 |
and ml_from_app var_ctxt fxy (app_expr as ((c, (lss, ty)), es)) = |
|
20600 | 641 |
case if is_cons c then [] else (map (ml_from_insts BR) o filter_out null) lss |
18885 | 642 |
of [] => |
20699 | 643 |
from_app (ml_mk_app var_ctxt) (ml_from_expr var_ctxt) const_syntax fxy app_expr |
18885 | 644 |
| lss => |
20456 | 645 |
if (is_none o const_syntax) c then |
646 |
brackify fxy ( |
|
647 |
(str o resolv) c |
|
648 |
:: (lss |
|
20699 | 649 |
@ map (ml_from_expr var_ctxt BR) es) |
20456 | 650 |
) |
20699 | 651 |
else error ("Can't apply user defined serialization for function expecting dicitonaries: " ^ quote c) |
652 |
in (ml_from_dictvar, ml_from_label, ml_from_tyvar, ml_from_insts, |
|
653 |
ml_from_tycoexpr, ml_from_type, ml_from_expr) end; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
654 |
|
20699 | 655 |
fun ml_fun_datatype init_ctxt (tyco_syntax, const_syntax) resolv = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
656 |
let |
20699 | 657 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
658 |
val (ml_from_dictvar, ml_from_label, ml_from_tyvar, ml_from_insts, ml_from_tycoexpr, ml_from_type, ml_from_expr) = |
|
659 |
ml_expr_seri (tyco_syntax, const_syntax) resolv; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
660 |
fun chunk_defs ps = |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
661 |
let |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
662 |
val (p_init, p_last) = split_last ps |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
663 |
in |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
664 |
Pretty.chunks (p_init @ [Pretty.block ([p_last, str ";"])]) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
665 |
end; |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
666 |
fun eta_expand_poly_fun (funn as (_, (_::_, _))) = |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
667 |
funn |
20456 | 668 |
| eta_expand_poly_fun (funn as (eqs, tysm as (_, ty))) = |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
669 |
let |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
670 |
fun no_eta (_::_, _) = I |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
671 |
| no_eta (_, _ `|-> _) = I |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
672 |
| no_eta ([], e) = K false; |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
673 |
fun has_tyvars (_ `%% tys) = |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
674 |
exists has_tyvars tys |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
675 |
| has_tyvars (ITyVar _) = |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
676 |
true |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
677 |
| has_tyvars (ty1 `-> ty2) = |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
678 |
has_tyvars ty1 orelse has_tyvars ty2; |
19622 | 679 |
in if (null o fst o CodegenThingol.unfold_fun) ty |
680 |
orelse (not o has_tyvars) ty |
|
681 |
orelse fold no_eta eqs true |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
682 |
then funn |
20456 | 683 |
else (map (fn ([], rhs) => ([IVar "x"], rhs `$ IVar "x")) eqs, tysm) |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
684 |
end; |
18912 | 685 |
fun ml_from_funs (defs as def::defs_tl) = |
18216 | 686 |
let |
19213 | 687 |
fun mk_definer [] [] = "val" |
20456 | 688 |
| mk_definer (_::_) _ = "fun" |
689 |
| mk_definer [] vs = if (null o filter_out (null o snd)) vs then "val" else "fun"; |
|
690 |
fun check_args (_, ((pats, _)::_, (vs, _))) NONE = |
|
691 |
SOME (mk_definer pats vs) |
|
692 |
| check_args (_, ((pats, _)::_, (vs, _))) (SOME definer) = |
|
693 |
if mk_definer pats vs = definer |
|
18216 | 694 |
then SOME definer |
20389 | 695 |
else error ("Mixing simultaneous vals and funs not implemented"); |
20456 | 696 |
fun mk_fun definer (name, (eqs as eq::eq_tl, (raw_vs, ty))) = |
18216 | 697 |
let |
20456 | 698 |
val vs = filter_out (null o snd) raw_vs; |
19136 | 699 |
val shift = if null eq_tl then I else |
700 |
map (Pretty.block o single o Pretty.block o single); |
|
18912 | 701 |
fun mk_eq definer (pats, expr) = |
20699 | 702 |
let |
703 |
val consts = map_filter |
|
704 |
(fn c => if (is_some o const_syntax) c |
|
705 |
then NONE else (SOME o NameSpace.base o resolv) c) (fold CodegenThingol.add_constnames (expr :: pats) []); |
|
706 |
val vars = fold CodegenThingol.add_unbound_varnames (expr :: pats) []; |
|
707 |
val var_ctxt = init_ctxt |
|
708 |
|> intro_ctxt consts |
|
709 |
|> intro_ctxt vars; |
|
710 |
in |
|
711 |
(Pretty.block o Pretty.breaks) ( |
|
712 |
[str definer, (str o resolv) name] |
|
713 |
@ (if null pats andalso null vs |
|
714 |
andalso not (ty = ITyVar "_")(*for evaluation*) |
|
715 |
then [str ":", ml_from_type NOBR ty] |
|
716 |
else |
|
717 |
map ml_from_tyvar vs |
|
718 |
@ map (ml_from_expr var_ctxt BR) pats) |
|
719 |
@ [str "=", ml_from_expr var_ctxt NOBR expr] |
|
720 |
) |
|
721 |
end |
|
18216 | 722 |
in |
18912 | 723 |
(Pretty.block o Pretty.fbreaks o shift) ( |
724 |
mk_eq definer eq |
|
725 |
:: map (mk_eq "|") eq_tl |
|
726 |
) |
|
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
727 |
end; |
20699 | 728 |
val def' :: defs' = map (apsnd eta_expand_poly_fun o constructive_fun) defs |
18216 | 729 |
in |
730 |
chunk_defs ( |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
731 |
(mk_fun (the (fold check_args defs NONE))) def' |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
732 |
:: map (mk_fun "and") defs' |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
733 |
) |
18216 | 734 |
end; |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
735 |
fun ml_from_datatypes (defs as (def::defs_tl)) = |
18216 | 736 |
let |
18702 | 737 |
fun mk_cons (co, []) = |
738 |
str (resolv co) |
|
739 |
| mk_cons (co, tys) = |
|
19136 | 740 |
Pretty.block [ |
741 |
str (resolv co), |
|
742 |
str " of", |
|
743 |
Pretty.brk 1, |
|
20600 | 744 |
Pretty.enum " *" "" "" (map (ml_from_type (INFX (2, L))) tys) |
19136 | 745 |
] |
19038 | 746 |
fun mk_datatype definer (t, (vs, cs)) = |
18912 | 747 |
(Pretty.block o Pretty.breaks) ( |
18702 | 748 |
str definer |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
749 |
:: ml_from_tycoexpr NOBR (t, map (ITyVar o fst) vs) |
18912 | 750 |
:: str "=" |
751 |
:: separate (str "|") (map mk_cons cs) |
|
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
752 |
) |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
753 |
in |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
754 |
chunk_defs ( |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
755 |
mk_datatype "datatype" def |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
756 |
:: map (mk_datatype "and") defs_tl |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
757 |
) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
758 |
end; |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
759 |
in (ml_from_funs, ml_from_datatypes) end; |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
760 |
|
20699 | 761 |
fun ml_from_defs init_ctxt |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19785
diff
changeset
|
762 |
(_, tyco_syntax, const_syntax) resolver prefix defs = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
763 |
let |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
764 |
val resolv = resolver prefix; |
20699 | 765 |
val (ml_from_dictvar, ml_from_label, ml_from_tyvar, ml_from_insts, ml_from_tycoexpr, ml_from_type, ml_from_expr) = |
766 |
ml_expr_seri (tyco_syntax, const_syntax) resolv; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
767 |
val (ml_from_funs, ml_from_datatypes) = |
20699 | 768 |
ml_fun_datatype init_ctxt (tyco_syntax, const_syntax) resolv; |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
769 |
val filter_datatype = |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
770 |
map_filter |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
771 |
(fn (name, CodegenThingol.Datatype info) => SOME (name, info) |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
772 |
| (name, CodegenThingol.Datatypecons _) => NONE |
20389 | 773 |
| (name, def) => error ("Datatype block containing illegal def: " |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
774 |
^ (Pretty.output o CodegenThingol.pretty_def) def)); |
20191 | 775 |
fun filter_class defs = |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
776 |
case map_filter |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
777 |
(fn (name, CodegenThingol.Class info) => SOME (name, info) |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
778 |
| (name, CodegenThingol.Classmember _) => NONE |
20389 | 779 |
| (name, def) => error ("Class block containing illegal def: " |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
780 |
^ (Pretty.output o CodegenThingol.pretty_def) def)) defs |
19136 | 781 |
of [class] => class |
20389 | 782 |
| _ => error ("Class block without class: " ^ (commas o map (quote o fst)) defs) |
19136 | 783 |
fun ml_from_class (name, (supclasses, (v, membrs))) = |
784 |
let |
|
20428 | 785 |
val w = ml_from_dictvar v; |
19136 | 786 |
fun from_supclass class = |
787 |
Pretty.block [ |
|
788 |
ml_from_label class, |
|
789 |
str ":", |
|
790 |
Pretty.brk 1, |
|
791 |
str ("'" ^ v), |
|
792 |
Pretty.brk 1, |
|
793 |
(str o resolv) class |
|
794 |
]; |
|
20456 | 795 |
fun from_membr (m, ty) = |
19136 | 796 |
Pretty.block [ |
797 |
ml_from_label m, |
|
798 |
str ":", |
|
799 |
Pretty.brk 1, |
|
800 |
ml_from_type NOBR ty |
|
801 |
]; |
|
802 |
fun from_membr_fun (m, _) = |
|
803 |
(Pretty.block o Pretty.breaks) [ |
|
804 |
str "fun", |
|
20191 | 805 |
(str o resolv) m, |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
806 |
Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ resolv name)], |
19136 | 807 |
str "=", |
808 |
Pretty.block [str "#", ml_from_label m], |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
809 |
str (w ^ ";") |
19136 | 810 |
]; |
811 |
in |
|
812 |
Pretty.chunks ( |
|
813 |
(Pretty.block o Pretty.breaks) [ |
|
814 |
str "type", |
|
815 |
str ("'" ^ v), |
|
816 |
(str o resolv) name, |
|
817 |
str "=", |
|
818 |
Pretty.enum "," "{" "};" ( |
|
819 |
map from_supclass supclasses @ map from_membr membrs |
|
820 |
) |
|
821 |
] |
|
822 |
:: map from_membr_fun membrs) |
|
823 |
end |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19785
diff
changeset
|
824 |
fun ml_from_def (name, CodegenThingol.Typesyn (vs, ty)) = |
18865 | 825 |
(map (fn (vname, []) => () | _ => |
20389 | 826 |
error "Can't serialize sort constrained type declaration to ML") vs; |
18702 | 827 |
Pretty.block [ |
828 |
str "type ", |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
829 |
ml_from_tycoexpr NOBR (name, map (ITyVar o fst) vs), |
18702 | 830 |
str " =", |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
831 |
Pretty.brk 1, |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
832 |
ml_from_type NOBR ty, |
18702 | 833 |
str ";" |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
834 |
] |
20191 | 835 |
) |> SOME |
20389 | 836 |
| ml_from_def (name, CodegenThingol.Classinst ((class, (tyco, arity)), (suparities, memdefs))) = |
18865 | 837 |
let |
838 |
val definer = if null arity then "val" else "fun" |
|
20466 | 839 |
fun from_supclass (supclass, (supinst, lss)) = |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
840 |
(Pretty.block o Pretty.breaks) [ |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
841 |
ml_from_label supclass, |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
842 |
str "=", |
20466 | 843 |
ml_from_insts NOBR [Instance (supinst, lss)] |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
844 |
]; |
20389 | 845 |
fun from_memdef (m, e) = |
20699 | 846 |
let |
847 |
val consts = map_filter |
|
848 |
(fn c => if (is_some o const_syntax) c |
|
849 |
then NONE else (SOME o NameSpace.base o resolv) c) (CodegenThingol.add_constnames e []); |
|
850 |
val var_ctxt = init_ctxt |
|
851 |
|> intro_ctxt consts; |
|
852 |
in |
|
853 |
(Pretty.block o Pretty.breaks) [ |
|
854 |
ml_from_label m, |
|
855 |
str "=", |
|
856 |
ml_from_expr var_ctxt NOBR e |
|
857 |
] |
|
858 |
end; |
|
19253 | 859 |
fun mk_corp rhs = |
860 |
(Pretty.block o Pretty.breaks) ( |
|
861 |
str definer |
|
862 |
:: (str o resolv) name |
|
863 |
:: map ml_from_tyvar arity |
|
864 |
@ [str "=", rhs] |
|
865 |
); |
|
20389 | 866 |
fun mk_memdefs supclassexprs memdefs = |
867 |
(mk_corp o Pretty.block o Pretty.breaks) [ |
|
868 |
Pretty.enum "," "{" "}" (supclassexprs @ memdefs), |
|
869 |
str ":", |
|
870 |
ml_from_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) |
|
871 |
]; |
|
18865 | 872 |
in |
20389 | 873 |
mk_memdefs (map from_supclass suparities) (map from_memdef memdefs) |> SOME |
19213 | 874 |
end |
20389 | 875 |
| ml_from_def (name, def) = error ("Illegal definition for " ^ quote name ^ ": " ^ |
20353 | 876 |
(Pretty.output o CodegenThingol.pretty_def) def); |
18850 | 877 |
in case defs |
20353 | 878 |
of (_, CodegenThingol.Fun _)::_ => ((*writeln "FUN";*) (SOME o ml_from_funs o map (fn (name, CodegenThingol.Fun info) => (name, info))) defs) |
879 |
| (_, CodegenThingol.Datatypecons _)::_ => ((*writeln "DTCO";*) (SOME o ml_from_datatypes o filter_datatype) defs) |
|
880 |
| (_, CodegenThingol.Datatype _)::_ => ((*writeln "DT";*) (SOME o ml_from_datatypes o filter_datatype) defs) |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
881 |
| (_, CodegenThingol.Class _)::_ => (SOME o ml_from_class o filter_class) defs |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
882 |
| (_, CodegenThingol.Classmember _)::_ => (SOME o ml_from_class o filter_class) defs |
18850 | 883 |
| [def] => ml_from_def def |
20389 | 884 |
| defs => error ("Illegal mutual dependencies: " ^ (commas o map fst) defs) |
18216 | 885 |
end; |
886 |
||
20699 | 887 |
fun ml_serializer root_name target nspgrp = |
18216 | 888 |
let |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
889 |
fun ml_from_module resolv _ ((_, name), ps) = |
18756 | 890 |
Pretty.chunks ([ |
18702 | 891 |
str ("structure " ^ name ^ " = "), |
892 |
str "struct", |
|
893 |
str "" |
|
894 |
] @ separate (str "") ps @ [ |
|
895 |
str "", |
|
896 |
str ("end; (* struct " ^ name ^ " *)") |
|
18756 | 897 |
]); |
20401 | 898 |
fun postproc (shallow, n) = |
899 |
let |
|
900 |
fun ch_first f = String.implode o nth_map 0 f o String.explode; |
|
20699 | 901 |
in if shallow = CodegenNames.nsp_dtco |
20401 | 902 |
then ch_first Char.toUpper n |
903 |
else n |
|
904 |
end; |
|
20191 | 905 |
in abstract_serializer (target, nspgrp) |
20699 | 906 |
root_name (ml_from_defs (make_ctxt ((ThmDatabase.ml_reserved @ reserved_ml'))), ml_from_module, |
907 |
abstract_validator (ThmDatabase.ml_reserved @ reserved_ml'), postproc) end; |
|
20191 | 908 |
|
909 |
in |
|
910 |
||
20699 | 911 |
val ml_fun_datatype = fn thy => |
20191 | 912 |
let |
20699 | 913 |
val target = "SML"; |
914 |
val data = CodegenSerializerData.get thy; |
|
915 |
val target_data = |
|
916 |
(the oo Symtab.lookup) data target; |
|
917 |
val syntax_tyco = #syntax_tyco target_data; |
|
918 |
val syntax_const = #syntax_const target_data; |
|
919 |
fun fun_of syntax = (Option.map fst oo Symtab.lookup) syntax; |
|
920 |
in ml_fun_datatype (make_ctxt (ThmDatabase.ml_reserved @ reserved_ml')) (fun_of syntax_tyco, fun_of syntax_const) end; |
|
921 |
||
922 |
fun ml_from_thingol target = |
|
923 |
let |
|
924 |
val serializer = ml_serializer "ROOT" target [[nsp_module], [nsp_class, nsp_tyco], |
|
925 |
[nsp_fun, nsp_dtco, nsp_class, nsp_classop, nsp_inst]] |
|
18850 | 926 |
val parse_multi = |
927 |
OuterParse.name |
|
20191 | 928 |
#-> (fn "dir" => |
18850 | 929 |
parse_multi_file |
930 |
(K o SOME o str o suffix ";" o prefix "val _ = use " |
|
931 |
o quote o suffix ".ML" o translate_string (fn "." => "/" | s => s)) "ML" serializer |
|
932 |
| _ => Scan.fail ()); |
|
18282 | 933 |
in |
19884 | 934 |
parse_multi |
935 |
|| parse_internal serializer |
|
20105 | 936 |
|| parse_stdout serializer |
19884 | 937 |
|| parse_single_file serializer |
18216 | 938 |
end; |
939 |
||
20389 | 940 |
val eval_verbose = ref false; |
941 |
||
20699 | 942 |
fun eval_term thy ((ref_name, reff), e) code = |
20191 | 943 |
let |
20699 | 944 |
val (val_name, code') = CodegenThingol.add_eval_def (nsp_eval, e) code; |
20191 | 945 |
val struct_name = "EVAL"; |
20401 | 946 |
fun output p = if !eval_verbose then (Pretty.writeln p; Pretty.output p) |
947 |
else Pretty.output p; |
|
20699 | 948 |
val target = "SML"; |
949 |
val data = CodegenSerializerData.get thy; |
|
950 |
val target_data = |
|
951 |
(the oo Symtab.lookup) data target; |
|
952 |
val syntax_tyco = #syntax_tyco target_data; |
|
953 |
val syntax_const = #syntax_const target_data; |
|
954 |
fun fun_of syntax = (Option.map fst oo Symtab.lookup) syntax; |
|
955 |
val serializer = ml_serializer struct_name "SML" [[nsp_module], [nsp_class, nsp_tyco], |
|
956 |
[nsp_fun, nsp_dtco, nsp_class, nsp_classop, nsp_inst], [nsp_eval]] |
|
20401 | 957 |
(fn "" => (fn p => (use_text Context.ml_output (!eval_verbose) (output p); NONE)) |
20699 | 958 |
| _ => SOME) (K NONE, fun_of syntax_tyco, fun_of syntax_const) |
959 |
((Symtab.keys (#syntax_tyco target_data) @ Symtab.keys (#syntax_const target_data)), SOME [NameSpace.pack [nsp_eval, val_name]]); |
|
960 |
val _ = serializer code'; |
|
20191 | 961 |
val val_name_struct = NameSpace.append struct_name val_name; |
20600 | 962 |
val _ = reff := NONE; |
963 |
val _ = use_text Context.ml_output (!eval_verbose) ("val _ = (" ^ ref_name ^ " := SOME (" ^ val_name_struct ^ "))"); |
|
964 |
in case !reff |
|
965 |
of NONE => error ("Could not retrieve value of ML reference " ^ quote ref_name |
|
966 |
^ " (reference probably has been shadowed)") |
|
967 |
| SOME value => value |
|
968 |
end; |
|
20191 | 969 |
|
20699 | 970 |
structure NameMangler = NameManglerFun ( |
971 |
type ctxt = string list; |
|
972 |
type src = string; |
|
973 |
val ord = string_ord; |
|
974 |
fun mk reserved_ml (name, i) = |
|
975 |
(Symbol.alphanum o NameSpace.base) name ^ replicate_string i "'"; |
|
976 |
fun is_valid (reserved_ml : string list) = not o member (op =) reserved_ml; |
|
977 |
fun maybe_unique _ _ = NONE; |
|
978 |
fun re_mangle _ dst = error ("No such definition name: " ^ quote dst); |
|
979 |
); |
|
980 |
||
19150 | 981 |
fun mk_flat_ml_resolver names = |
982 |
let |
|
983 |
val mangler = |
|
984 |
NameMangler.empty |
|
20699 | 985 |
|> fold_map (NameMangler.declare (ThmDatabase.ml_reserved @ reserved_ml')) names |
19150 | 986 |
|-> (fn _ => I) |
20699 | 987 |
in NameMangler.get (ThmDatabase.ml_reserved @ reserved_ml') mangler end; |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
988 |
|
18216 | 989 |
end; (* local *) |
990 |
||
20191 | 991 |
|
992 |
(** haskell serializer **) |
|
993 |
||
18282 | 994 |
local |
995 |
||
20699 | 996 |
fun hs_from_defs init_ctxt (class_syntax, tyco_syntax, const_syntax) |
19038 | 997 |
resolver prefix defs = |
18282 | 998 |
let |
20699 | 999 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
1000 |
val resolv = resolver ""; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
1001 |
val resolv_here = resolver prefix; |
20428 | 1002 |
fun hs_from_class cls = case class_syntax cls |
1003 |
of NONE => resolv cls |
|
1004 |
| SOME (cls, _) => cls; |
|
1005 |
fun hs_from_classop_name cls clsop = case class_syntax cls |
|
1006 |
of NONE => NameSpace.base clsop |
|
1007 |
| SOME (_, classop_syntax) => case classop_syntax clsop |
|
1008 |
of NONE => NameSpace.base clsop |
|
1009 |
| SOME clsop => clsop; |
|
20699 | 1010 |
fun hs_from_typparms tyvar_ctxt vs = |
18282 | 1011 |
let |
20456 | 1012 |
fun from_typparms [] = str "" |
1013 |
| from_typparms vs = |
|
18282 | 1014 |
vs |
20699 | 1015 |
|> map (fn (v, cls) => str |
1016 |
(hs_from_class cls ^ " " ^ lookup_ctxt tyvar_ctxt v)) |
|
18812 | 1017 |
|> Pretty.enum "," "(" ")" |
18702 | 1018 |
|> (fn p => Pretty.block [p, str " => "]) |
20191 | 1019 |
in |
18282 | 1020 |
vs |
20699 | 1021 |
|> maps (fn (v, sort) => map (pair v) sort) |
20456 | 1022 |
|> from_typparms |
18282 | 1023 |
end; |
20699 | 1024 |
fun hs_from_tycoexpr tyvar_ctxt fxy (tyco, tys) = |
1025 |
brackify fxy (str tyco :: map (hs_from_type tyvar_ctxt BR) tys) |
|
1026 |
and hs_from_type tyvar_ctxt fxy (tycoexpr as tyco `%% tys) = |
|
18865 | 1027 |
(case tyco_syntax tyco |
18335 | 1028 |
of NONE => |
20699 | 1029 |
hs_from_tycoexpr tyvar_ctxt fxy (resolv tyco, tys) |
1030 |
| SOME (i, pr) => |
|
1031 |
if not (i = length tys) |
|
20389 | 1032 |
then error ("Number of argument mismatch in customary serialization: " |
18865 | 1033 |
^ (string_of_int o length) tys ^ " given, " |
20699 | 1034 |
^ string_of_int i ^ " expected") |
1035 |
else pr fxy (hs_from_type tyvar_ctxt) tys) |
|
1036 |
| hs_from_type tyvar_ctxt fxy (t1 `-> t2) = |
|
18865 | 1037 |
brackify_infix (1, R) fxy [ |
20699 | 1038 |
hs_from_type tyvar_ctxt (INFX (1, X)) t1, |
18865 | 1039 |
str "->", |
20699 | 1040 |
hs_from_type tyvar_ctxt (INFX (1, R)) t2 |
18865 | 1041 |
] |
20699 | 1042 |
| hs_from_type tyvar_ctxt fxy (ITyVar v) = |
1043 |
str (lookup_ctxt tyvar_ctxt v); |
|
1044 |
fun hs_from_typparms_tycoexpr tyvar_ctxt (vs, tycoexpr) = |
|
1045 |
Pretty.block [hs_from_typparms tyvar_ctxt vs, hs_from_tycoexpr tyvar_ctxt NOBR tycoexpr] |
|
1046 |
fun hs_from_typparms_type tyvar_ctxt (vs, ty) = |
|
1047 |
Pretty.block [hs_from_typparms tyvar_ctxt vs, hs_from_type tyvar_ctxt NOBR ty] |
|
1048 |
fun hs_from_expr var_ctxt fxy (e as IConst x) = |
|
1049 |
hs_from_app var_ctxt fxy (x, []) |
|
1050 |
| hs_from_expr var_ctxt fxy (e as (e1 `$ e2)) = |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1051 |
(case CodegenThingol.unfold_const_app e |
20699 | 1052 |
of SOME x => hs_from_app var_ctxt fxy x |
18282 | 1053 |
| _ => |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1054 |
brackify fxy [ |
20699 | 1055 |
hs_from_expr var_ctxt NOBR e1, |
1056 |
hs_from_expr var_ctxt BR e2 |
|
18282 | 1057 |
]) |
20699 | 1058 |
| hs_from_expr var_ctxt fxy (IVar v) = |
1059 |
str (lookup_ctxt var_ctxt v) |
|
1060 |
| hs_from_expr var_ctxt fxy (e as _ `|-> _) = |
|
18282 | 1061 |
let |
20699 | 1062 |
val (es, e) = CodegenThingol.unfold_abs e; |
1063 |
val vs = fold CodegenThingol.add_varnames (map fst es) []; |
|
1064 |
val var_ctxt' = intro_ctxt vs var_ctxt; |
|
18282 | 1065 |
in |
19038 | 1066 |
brackify BR ( |
18702 | 1067 |
str "\\" |
20699 | 1068 |
:: map (hs_from_expr var_ctxt' BR o fst) es @ [ |
18702 | 1069 |
str "->", |
20699 | 1070 |
hs_from_expr var_ctxt' NOBR e |
18282 | 1071 |
]) |
1072 |
end |
|
20699 | 1073 |
| hs_from_expr var_ctxt fxy (INum (n, _)) = |
20600 | 1074 |
if n > 0 then |
1075 |
(str o IntInf.toString) n |
|
1076 |
else |
|
1077 |
brackify BR [(str o Library.prefix "-" o IntInf.toString o IntInf.~) n] |
|
20699 | 1078 |
| hs_from_expr var_ctxt fxy (IChar (c, _)) = |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
1079 |
(str o enclose "'" "'") |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
1080 |
(let val i = (Char.ord o the o Char.fromString) c |
20191 | 1081 |
in if i < 32 |
20105 | 1082 |
then Library.prefix "\\" (string_of_int i) |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
1083 |
else c |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
1084 |
end) |
20699 | 1085 |
| hs_from_expr var_ctxt fxy (e as ICase ((_, [_]), _)) = |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1086 |
let |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1087 |
val (ps, body) = CodegenThingol.unfold_let e; |
20699 | 1088 |
fun mk_bind ((p, _), e) var_ctxt = |
1089 |
let |
|
1090 |
val vs = CodegenThingol.add_varnames p []; |
|
1091 |
val var_ctxt' = intro_ctxt vs var_ctxt; |
|
1092 |
in |
|
1093 |
((Pretty.block o Pretty.breaks) [ |
|
1094 |
hs_from_expr var_ctxt' BR p, |
|
1095 |
str "=", |
|
1096 |
hs_from_expr var_ctxt NOBR e |
|
1097 |
], var_ctxt') |
|
1098 |
end; |
|
1099 |
val (binds, var_ctxt') = fold_map mk_bind ps var_ctxt; |
|
18282 | 1100 |
in Pretty.chunks [ |
20699 | 1101 |
[str ("let"), Pretty.fbrk, binds |> Pretty.chunks] |> Pretty.block, |
1102 |
[str ("in "), hs_from_expr var_ctxt' NOBR body] |> Pretty.block |
|
18282 | 1103 |
] end |
20699 | 1104 |
| hs_from_expr var_ctxt fxy (ICase (((de, _), bses), _)) = |
18282 | 1105 |
let |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1106 |
fun mk_clause (se, be) = |
20699 | 1107 |
let |
1108 |
val vs = CodegenThingol.add_varnames se []; |
|
1109 |
val var_ctxt' = intro_ctxt vs var_ctxt; |
|
1110 |
in |
|
1111 |
(Pretty.block o Pretty.breaks) [ |
|
1112 |
hs_from_expr var_ctxt' NOBR se, |
|
1113 |
str "->", |
|
1114 |
hs_from_expr var_ctxt' NOBR be |
|
1115 |
] |
|
1116 |
end |
|
18850 | 1117 |
in Pretty.block [ |
1118 |
str "case", |
|
1119 |
Pretty.brk 1, |
|
20699 | 1120 |
hs_from_expr var_ctxt NOBR de, |
18850 | 1121 |
Pretty.brk 1, |
1122 |
str "of", |
|
1123 |
Pretty.fbrk, |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1124 |
(Pretty.chunks o map mk_clause) bses |
18850 | 1125 |
] end |
20699 | 1126 |
and hs_mk_app var_ctxt c es = |
1127 |
(str o resolv) c :: map (hs_from_expr var_ctxt BR) es |
|
1128 |
and hs_from_app var_ctxt fxy = |
|
1129 |
from_app (hs_mk_app var_ctxt) (hs_from_expr var_ctxt) const_syntax fxy |
|
19202 | 1130 |
fun hs_from_funeqs (def as (name, _)) = |
18865 | 1131 |
let |
19202 | 1132 |
fun from_eq (args, rhs) = |
20699 | 1133 |
let |
1134 |
val consts = map_filter |
|
1135 |
(fn c => if (is_some o const_syntax) c |
|
1136 |
then NONE else (SOME o NameSpace.base o resolv) c) (fold CodegenThingol.add_constnames (rhs :: args) []); |
|
1137 |
val vars = fold CodegenThingol.add_unbound_varnames (rhs :: args) []; |
|
1138 |
val var_ctxt = init_ctxt |
|
1139 |
|> intro_ctxt consts |
|
1140 |
|> intro_ctxt vars; |
|
1141 |
in |
|
1142 |
Pretty.block [ |
|
1143 |
(str o resolv_here) name, |
|
1144 |
Pretty.block (map (fn p => Pretty.block [Pretty.brk 1, hs_from_expr var_ctxt BR p]) args), |
|
1145 |
Pretty.brk 1, |
|
1146 |
str ("="), |
|
1147 |
Pretty.brk 1, |
|
1148 |
hs_from_expr var_ctxt NOBR rhs |
|
1149 |
] |
|
1150 |
end |
|
1151 |
in Pretty.chunks ((map from_eq o fst o snd o constructive_fun) def) end; |
|
20456 | 1152 |
fun hs_from_def (name, CodegenThingol.Fun (def as (_, (vs, ty)))) = |
18963 | 1153 |
let |
20699 | 1154 |
val tyvar_ctxt = intro_ctxt (map fst vs) init_ctxt; |
19202 | 1155 |
val body = hs_from_funeqs (name, def); |
20699 | 1156 |
in |
18963 | 1157 |
Pretty.chunks [ |
1158 |
Pretty.block [ |
|
19038 | 1159 |
(str o suffix " ::" o resolv_here) name, |
18963 | 1160 |
Pretty.brk 1, |
20699 | 1161 |
hs_from_typparms_type tyvar_ctxt (vs, ty) |
18963 | 1162 |
], |
1163 |
body |
|
1164 |
] |> SOME |
|
20699 | 1165 |
end |
20456 | 1166 |
| hs_from_def (name, CodegenThingol.Typesyn (vs, ty)) = |
20699 | 1167 |
let |
1168 |
val tyvar_ctxt = intro_ctxt (map fst vs) init_ctxt; |
|
1169 |
in |
|
1170 |
(Pretty.block o Pretty.breaks) [ |
|
1171 |
str "type", |
|
1172 |
hs_from_typparms_tycoexpr tyvar_ctxt (vs, (resolv_here name, map (ITyVar o fst) vs)), |
|
1173 |
str "=", |
|
1174 |
hs_from_typparms_type tyvar_ctxt ([], ty) |
|
1175 |
] |> SOME |
|
1176 |
end |
|
20456 | 1177 |
| hs_from_def (name, CodegenThingol.Datatype (vs, [(co, [ty])])) = |
20699 | 1178 |
let |
1179 |
val tyvar_ctxt = intro_ctxt (map fst vs) init_ctxt; |
|
1180 |
in |
|
1181 |
(Pretty.block o Pretty.breaks) [ |
|
1182 |
str "newtype", |
|
1183 |
hs_from_typparms_tycoexpr tyvar_ctxt (vs, (resolv_here name, map (ITyVar o fst) vs)), |
|
1184 |
str "=", |
|
1185 |
(str o resolv_here) co, |
|
1186 |
hs_from_type tyvar_ctxt BR ty |
|
1187 |
] |> SOME |
|
1188 |
end |
|
20456 | 1189 |
| hs_from_def (name, CodegenThingol.Datatype (vs, constrs)) = |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1190 |
let |
20699 | 1191 |
val tyvar_ctxt = intro_ctxt (map fst vs) init_ctxt; |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1192 |
fun mk_cons (co, tys) = |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1193 |
(Pretty.block o Pretty.breaks) ( |
19038 | 1194 |
(str o resolv_here) co |
20699 | 1195 |
:: map (hs_from_type tyvar_ctxt BR) tys |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1196 |
) |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1197 |
in |
19785 | 1198 |
(Pretty.block o Pretty.breaks) [ |
1199 |
str "data", |
|
20699 | 1200 |
hs_from_typparms_tycoexpr tyvar_ctxt (vs, (resolv_here name, map (ITyVar o fst) vs)), |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19785
diff
changeset
|
1201 |
str "=", |
19785 | 1202 |
Pretty.block (separate (Pretty.block [Pretty.brk 1, str "| "]) (map mk_cons constrs)) |
1203 |
] |
|
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1204 |
end |> SOME |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1205 |
| hs_from_def (_, CodegenThingol.Datatypecons _) = |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1206 |
NONE |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1207 |
| hs_from_def (name, CodegenThingol.Class (supclasss, (v, membrs))) = |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1208 |
let |
20699 | 1209 |
val tyvar_ctxt = intro_ctxt [v] init_ctxt; |
20456 | 1210 |
fun mk_member (m, ty) = |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1211 |
Pretty.block [ |
19038 | 1212 |
str (resolv_here m ^ " ::"), |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1213 |
Pretty.brk 1, |
20699 | 1214 |
hs_from_type tyvar_ctxt NOBR ty |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1215 |
] |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1216 |
in |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1217 |
Pretty.block [ |
18702 | 1218 |
str "class ", |
20699 | 1219 |
hs_from_typparms tyvar_ctxt [(v, supclasss)], |
19038 | 1220 |
str (resolv_here name ^ " " ^ v), |
18702 | 1221 |
str " where", |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1222 |
Pretty.fbrk, |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1223 |
Pretty.chunks (map mk_member membrs) |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1224 |
] |> SOME |
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1225 |
end |
19213 | 1226 |
| hs_from_def (_, CodegenThingol.Classmember _) = |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1227 |
NONE |
20699 | 1228 |
| hs_from_def (_, CodegenThingol.Classinst ((clsname, (tyco, vs)), (_, memdefs))) = |
1229 |
let |
|
1230 |
val tyvar_ctxt = intro_ctxt (map fst vs) init_ctxt; |
|
1231 |
in |
|
1232 |
Pretty.block [ |
|
1233 |
str "instance ", |
|
1234 |
hs_from_typparms tyvar_ctxt vs, |
|
1235 |
str (hs_from_class clsname ^ " "), |
|
1236 |
hs_from_type tyvar_ctxt BR (tyco `%% map (ITyVar o fst) vs), |
|
1237 |
str " where", |
|
1238 |
Pretty.fbrk, |
|
1239 |
Pretty.chunks (map (fn (m, e) => |
|
1240 |
let |
|
1241 |
val consts = map_filter |
|
1242 |
(fn c => if (is_some o const_syntax) c |
|
1243 |
then NONE else (SOME o NameSpace.base o resolv) c) (CodegenThingol.add_constnames e []); |
|
1244 |
val var_ctxt = init_ctxt |
|
1245 |
|> intro_ctxt consts; |
|
1246 |
in |
|
1247 |
(Pretty.block o Pretty.breaks) [ |
|
1248 |
(str o hs_from_classop_name clsname) m, |
|
1249 |
str "=", |
|
1250 |
hs_from_expr var_ctxt NOBR e |
|
1251 |
] |
|
1252 |
end |
|
1253 |
) memdefs) |
|
1254 |
] |> SOME |
|
1255 |
end |
|
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1256 |
in |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
1257 |
case map_filter (fn (name, def) => hs_from_def (name, def)) defs |
18380
9668764224a7
substantial improvements for class code generation
haftmann
parents:
18361
diff
changeset
|
1258 |
of [] => NONE |
18702 | 1259 |
| l => (SOME o Pretty.chunks o separate (str "")) l |
18282 | 1260 |
end; |
1261 |
||
1262 |
in |
|
1263 |
||
20699 | 1264 |
fun hs_from_thingol target = |
18282 | 1265 |
let |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1266 |
val reserved_hs = [ |
18702 | 1267 |
"hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", |
1268 |
"import", "default", "forall", "let", "in", "class", "qualified", "data", |
|
1269 |
"newtype", "instance", "if", "then", "else", "type", "as", "do", "module" |
|
1270 |
] @ [ |
|
20401 | 1271 |
"Bool", "Integer", "Maybe", "True", "False", "Nothing", "Just", "not", "negate", |
1272 |
"String", "Char" |
|
18702 | 1273 |
]; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
1274 |
fun hs_from_module resolv imps ((_, name), ps) = |
19038 | 1275 |
(Pretty.chunks) ( |
1276 |
str ("module " ^ name ^ " where") |
|
1277 |
:: map (str o prefix "import qualified ") imps @ ( |
|
1278 |
str "" |
|
1279 |
:: separate (str "") ps |
|
1280 |
)); |
|
18919 | 1281 |
fun postproc (shallow, n) = |
1282 |
let |
|
1283 |
fun ch_first f = String.implode o nth_map 0 f o String.explode; |
|
20699 | 1284 |
in if member (op =) [nsp_module, nsp_class, nsp_tyco, nsp_dtco] shallow |
18919 | 1285 |
then ch_first Char.toUpper n |
1286 |
else ch_first Char.toLower n |
|
1287 |
end; |
|
20699 | 1288 |
val serializer = abstract_serializer (target, [[nsp_module], |
1289 |
[nsp_class], [nsp_tyco], [nsp_fun, nsp_classop], [nsp_dtco], [nsp_inst]]) |
|
1290 |
"Main" (hs_from_defs (make_ctxt reserved_hs), hs_from_module, abstract_validator reserved_hs, postproc); |
|
18282 | 1291 |
in |
20699 | 1292 |
parse_multi_file ((K o K) NONE) "hs" serializer |
18282 | 1293 |
end; |
1294 |
||
1295 |
end; (* local *) |
|
1296 |
||
18702 | 1297 |
|
20699 | 1298 |
|
1299 |
(** lookup table **) |
|
18702 | 1300 |
|
1301 |
val serializers = |
|
20699 | 1302 |
Symtab.empty |
1303 |
|> fold (fn (name, f) => Symtab.update_new (name, f name)) |
|
1304 |
[("SML", ml_from_thingol), ("Haskell", hs_from_thingol)]; |
|
1305 |
||
1306 |
fun check_serializer target = |
|
1307 |
case Symtab.lookup serializers target |
|
1308 |
of SOME seri => () |
|
1309 |
| NONE => error ("Unknown code target language: " ^ quote target); |
|
1310 |
||
1311 |
fun parse_serializer target = |
|
1312 |
case Symtab.lookup serializers target |
|
1313 |
of SOME seri => seri >> (fn seri' => fn thy => serialize thy seri' target) |
|
1314 |
| NONE => Scan.fail_with (fn _ => "Unknown code target language: " ^ quote target) (); |
|
1315 |
||
1316 |
fun map_target_data target f = |
|
18702 | 1317 |
let |
20699 | 1318 |
val _ = check_serializer target; |
1319 |
in |
|
1320 |
CodegenSerializerData.map ( |
|
1321 |
(Symtab.map_entry target (fn { syntax_class, syntax_inst, syntax_tyco, syntax_const } => |
|
1322 |
let |
|
1323 |
val (syntax_class, syntax_inst, syntax_tyco, syntax_const) = |
|
1324 |
f (syntax_class, syntax_inst, syntax_tyco, syntax_const) |
|
1325 |
in { |
|
1326 |
syntax_class = syntax_class, |
|
1327 |
syntax_inst = syntax_inst, |
|
1328 |
syntax_tyco = syntax_tyco, |
|
1329 |
syntax_const = syntax_const } : target_data |
|
1330 |
end |
|
1331 |
)) |
|
1332 |
) |
|
1333 |
end; |
|
1334 |
||
1335 |
||
1336 |
(** target syntax interfaces **) |
|
1337 |
||
1338 |
local |
|
1339 |
||
1340 |
fun gen_add_syntax_class prep_class prep_const target raw_class (syntax, raw_ops) thy = |
|
1341 |
let |
|
1342 |
val cls = prep_class thy raw_class |
|
1343 |
val class = CodegenNames.class thy cls; |
|
1344 |
fun mk_classop (c, _) = case AxClass.class_of_param thy c |
|
1345 |
of SOME class' => if cls = class' then c |
|
1346 |
else error ("Not a class operation for class " ^ quote class ^ ": " ^ quote c) |
|
1347 |
| NONE => error ("Not a class operation: " ^ quote c) |
|
1348 |
val ops = (map o apfst) (mk_classop o prep_const thy) raw_ops; |
|
1349 |
val syntax_ops = AList.lookup (op =) ops; |
|
1350 |
in |
|
1351 |
thy |
|
1352 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
|
1353 |
(syntax_class |> Symtab.update (class, |
|
1354 |
((syntax, syntax_ops), stamp ())), |
|
1355 |
syntax_inst, syntax_tyco, syntax_const)) |
|
1356 |
end; |
|
1357 |
||
1358 |
fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) thy = |
|
1359 |
let |
|
1360 |
val inst = CodegenNames.instance thy (prep_class thy raw_class, prep_tyco thy raw_tyco); |
|
1361 |
in |
|
1362 |
thy |
|
1363 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
|
1364 |
(syntax_class, syntax_inst |> Symtab.update (inst, ()), |
|
1365 |
syntax_tyco, syntax_const)) |
|
1366 |
end; |
|
1367 |
||
1368 |
fun gen_add_syntax_tyco prep_tyco raw_tyco target syntax thy = |
|
1369 |
let |
|
1370 |
val tyco = (CodegenNames.tyco thy o prep_tyco thy) raw_tyco; |
|
1371 |
in |
|
1372 |
thy |
|
1373 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
|
1374 |
(syntax_class, syntax_inst, syntax_tyco |
|
1375 |
|> Symtab.update (tyco, (syntax, stamp ())), syntax_const)) |
|
1376 |
end; |
|
1377 |
||
1378 |
fun gen_add_syntax_const prep_const raw_c target syntax thy = |
|
1379 |
let |
|
1380 |
val c' = prep_const thy raw_c; |
|
1381 |
val c'' = CodegenNames.const thy c'; |
|
1382 |
in |
|
1383 |
thy |
|
1384 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
|
1385 |
(syntax_class, syntax_inst, syntax_tyco, syntax_const |
|
1386 |
|> Symtab.update (c'', (syntax, stamp ())))) |
|
1387 |
end; |
|
1388 |
||
1389 |
fun read_type thy raw_tyco = |
|
1390 |
let |
|
1391 |
val tyco = Sign.intern_type thy raw_tyco; |
|
1392 |
val _ = if Sign.declared_tyname thy tyco then () |
|
1393 |
else error ("No such type constructor: " ^ quote raw_tyco); |
|
1394 |
in tyco end; |
|
1395 |
||
1396 |
fun idfs_of_const_names thy cs = |
|
1397 |
let |
|
1398 |
val cs' = AList.make (fn c => Sign.the_const_type thy c) cs; |
|
1399 |
val cs'' = map (CodegenConsts.norm_of_typ thy) cs'; |
|
1400 |
in AList.make (CodegenNames.const thy) cs'' end; |
|
1401 |
||
1402 |
fun read_quote reader consts_of target get_init gen raw_it thy = |
|
1403 |
let |
|
1404 |
val it = reader thy raw_it; |
|
1405 |
val cs = consts_of thy it; |
|
1406 |
in |
|
1407 |
gen thy cs target (get_init thy) [it] |
|
1408 |
|> (fn [it'] => (it', thy)) |
|
1409 |
end; |
|
1410 |
||
1411 |
fun parse_quote num_of reader consts_of target get_init gen adder = |
|
1412 |
parse_syntax num_of |
|
1413 |
(read_quote reader consts_of target get_init gen) |
|
1414 |
#-> (fn modifier => pair (modifier #-> adder target)); |
|
1415 |
||
1416 |
in |
|
1417 |
||
1418 |
val add_syntax_class = gen_add_syntax_class ClassPackage.read_class CodegenConsts.read_const; |
|
1419 |
val add_syntax_inst = gen_add_syntax_inst ClassPackage.read_class read_type; |
|
1420 |
||
1421 |
fun parse_syntax_tyco generate target raw_tyco = |
|
1422 |
let |
|
1423 |
fun intern thy = read_type thy raw_tyco; |
|
1424 |
fun num_of thy = Sign.arity_number thy (intern thy); |
|
1425 |
fun idf_of thy = CodegenNames.tyco thy (intern thy); |
|
1426 |
fun read_typ thy = |
|
1427 |
Sign.read_typ (thy, K NONE); |
|
1428 |
in |
|
1429 |
parse_quote num_of read_typ ((K o K) ([], [])) target idf_of generate |
|
1430 |
(gen_add_syntax_tyco read_type raw_tyco) |
|
1431 |
end; |
|
1432 |
||
1433 |
fun parse_syntax_const generate target raw_const = |
|
1434 |
let |
|
1435 |
fun intern thy = CodegenConsts.read_const thy raw_const; |
|
1436 |
fun num_of thy = (length o fst o strip_type o Sign.the_const_type thy o fst o intern) thy; |
|
1437 |
fun idf_of thy = (CodegenNames.const thy o intern) thy; |
|
1438 |
in |
|
1439 |
parse_quote num_of Sign.read_term CodegenConsts.consts_of target idf_of generate |
|
1440 |
(gen_add_syntax_const CodegenConsts.read_const raw_const) |
|
1441 |
end; |
|
1442 |
||
1443 |
fun add_pretty_list target nill cons mk_list mk_char_string target_cons thy = |
|
1444 |
let |
|
1445 |
val [(_, nil''), (cons', cons'')] = idfs_of_const_names thy [nill, cons]; |
|
1446 |
val pr = pretty_list nil'' cons'' mk_list mk_char_string target_cons; |
|
1447 |
in |
|
1448 |
thy |
|
1449 |
|> gen_add_syntax_const (K I) cons' target pr |
|
1450 |
end; |
|
1451 |
||
1452 |
fun add_pretty_ml_string target nill cons str mk_char mk_string target_implode thy = |
|
1453 |
let |
|
1454 |
val [(_, nil''), (_, cons''), (str', _)] = idfs_of_const_names thy [nill, cons, str]; |
|
1455 |
val pr = pretty_ml_string nil'' cons'' mk_char mk_string target_implode; |
|
1456 |
in |
|
1457 |
thy |
|
1458 |
|> gen_add_syntax_const (K I) str' target pr |
|
1459 |
end; |
|
1460 |
||
1461 |
fun add_undefined target undef target_undefined thy = |
|
1462 |
let |
|
1463 |
val [(undef', _)] = idfs_of_const_names thy [undef]; |
|
1464 |
fun pretty _ _ _ = str target_undefined; |
|
1465 |
in |
|
1466 |
thy |
|
1467 |
|> gen_add_syntax_const (K I) undef' target (~1, pretty) |
|
1468 |
end; |
|
1469 |
||
1470 |
end; (*local*) |
|
18702 | 1471 |
|
18216 | 1472 |
end; (* struct *) |