author | haftmann |
Fri, 24 Mar 2023 18:30:17 +0000 | |
changeset 77707 | a6a81f848135 |
parent 77703 | 0262155d2743 |
permissions | -rw-r--r-- |
37745
6315b6426200
checking generated code for various target languages
haftmann
parents:
37449
diff
changeset
|
1 |
(* Title: Tools/Code/code_ml.ML |
28054 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
4 |
Serializer for SML and OCaml. |
|
5 |
*) |
|
6 |
||
7 |
signature CODE_ML = |
|
8 |
sig |
|
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
9 |
val target_SML: string |
37745
6315b6426200
checking generated code for various target languages
haftmann
parents:
37449
diff
changeset
|
10 |
val target_OCaml: string |
28054 | 11 |
end; |
12 |
||
13 |
structure Code_ML : CODE_ML = |
|
14 |
struct |
|
15 |
||
55150 | 16 |
open Basic_Code_Symbol; |
28054 | 17 |
open Basic_Code_Thingol; |
18 |
open Code_Printer; |
|
19 |
||
20 |
infixr 5 @@; |
|
21 |
infixr 5 @|; |
|
22 |
||
33992 | 23 |
|
24 |
(** generic **) |
|
25 |
||
28054 | 26 |
val target_SML = "SML"; |
27 |
val target_OCaml = "OCaml"; |
|
28 |
||
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
29 |
datatype ml_binding = |
35228 | 30 |
ML_Function of string * (typscheme * ((iterm list * iterm) * (thm option * bool)) list) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
31 |
| ML_Instance of (string * class) * { class: class, tyco: string, vs: (vname * sort) list, |
77707
a6a81f848135
More explicit type information in dictionary arguments.
haftmann
parents:
77703
diff
changeset
|
32 |
superinsts: (class * (itype * dict list) list) list, |
52519
598addf65209
explicit hint for domain of class parameters in instance statements
haftmann
parents:
52435
diff
changeset
|
33 |
inst_params: ((string * (const * int)) * (thm * bool)) list, |
598addf65209
explicit hint for domain of class parameters in instance statements
haftmann
parents:
52435
diff
changeset
|
34 |
superinst_params: ((string * (const * int)) * (thm * bool)) list }; |
28054 | 35 |
|
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
36 |
datatype ml_stmt = |
33992 | 37 |
ML_Exc of string * (typscheme * int) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
38 |
| ML_Val of ml_binding |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
39 |
| ML_Funs of (Code_Namespace.export * ml_binding) list * Code_Symbol.T list |
48003
1d11af40b106
dropped sort constraints on datatype specifications
haftmann
parents:
47609
diff
changeset
|
40 |
| ML_Datas of (string * (vname list * ((string * vname list) * itype list) list)) list |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
41 |
| ML_Class of string * (vname * ((class * class) list * (string * itype) list)); |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
42 |
|
33992 | 43 |
fun print_product _ [] = NONE |
44 |
| print_product print [x] = SOME (print x) |
|
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
45 |
| print_product print xs = (SOME o enum " *" "" "") (map print xs); |
28054 | 46 |
|
38922 | 47 |
fun tuplify _ _ [] = NONE |
48 |
| tuplify print fxy [x] = SOME (print fxy x) |
|
49 |
| tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs)); |
|
28054 | 50 |
|
33992 | 51 |
|
52 |
(** SML serializer **) |
|
53 |
||
68028 | 54 |
fun print_sml_char c = |
55 |
if c = "\\" |
|
56 |
then "\\092" (*produce strings suitable for both SML as well as Isabelle/ML*) |
|
57 |
else if Symbol.is_ascii c |
|
69207 | 58 |
then ML_Syntax.print_symbol_char c |
68028 | 59 |
else error "non-ASCII byte in SML string literal"; |
59456
180555df34ea
string printing conformant to both (S)ML and Isabelle/ML
haftmann
parents:
59323
diff
changeset
|
60 |
|
68028 | 61 |
val print_sml_string = quote o translate_string print_sml_char; |
59456
180555df34ea
string printing conformant to both (S)ML and Isabelle/ML
haftmann
parents:
59323
diff
changeset
|
62 |
|
50625 | 63 |
fun print_sml_stmt tyco_syntax const_syntax reserved is_constr deresolve = |
28054 | 64 |
let |
55150 | 65 |
val deresolve_const = deresolve o Constant; |
66 |
val deresolve_classrel = deresolve o Class_Relation; |
|
67 |
val deresolve_inst = deresolve o Class_Instance; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
68 |
fun print_tyco_expr (sym, []) = (str o deresolve) sym |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
69 |
| print_tyco_expr (sym, [ty]) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
70 |
concat [print_typ BR ty, (str o deresolve) sym] |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
71 |
| print_tyco_expr (sym, tys) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
72 |
concat [enum "," "(" ")" (map (print_typ BR) tys), (str o deresolve) sym] |
38923 | 73 |
and print_typ fxy (tyco `%% tys) = (case tyco_syntax tyco |
55150 | 74 |
of NONE => print_tyco_expr (Type_Constructor tyco, tys) |
47609 | 75 |
| SOME (_, print) => print print_typ fxy tys) |
33992 | 76 |
| print_typ fxy (ITyVar v) = str ("'" ^ v); |
55150 | 77 |
fun print_dicttyp (class, ty) = print_tyco_expr (Type_Class class, [ty]); |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
78 |
fun print_typscheme_prefix (vs, p) = enum " ->" "" "" |
33992 | 79 |
(map_filter (fn (v, sort) => |
80 |
(print_product (fn class => print_dicttyp (class, ITyVar v)) sort)) vs @| p); |
|
81 |
fun print_typscheme (vs, ty) = print_typscheme_prefix (vs, print_typ NOBR ty); |
|
82 |
fun print_dicttypscheme (vs, class_ty) = print_typscheme_prefix (vs, print_dicttyp class_ty); |
|
41100
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
haftmann
parents:
40627
diff
changeset
|
83 |
fun print_classrels fxy [] ps = brackify fxy ps |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
84 |
| print_classrels fxy [classrel] ps = brackify fxy [(str o deresolve_classrel) classrel, brackify BR ps] |
41100
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
haftmann
parents:
40627
diff
changeset
|
85 |
| print_classrels fxy classrels ps = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
86 |
brackify fxy [enum " o" "(" ")" (map (str o deresolve_classrel) classrels), brackify BR ps] |
41118
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
87 |
fun print_dict is_pseudo_fun fxy (Dict (classrels, x)) = |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
88 |
print_classrels fxy classrels (print_plain_dict is_pseudo_fun fxy x) |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
89 |
and print_plain_dict is_pseudo_fun fxy (Dict_Const (inst, dss)) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
90 |
((str o deresolve_inst) inst :: |
55150 | 91 |
(if is_pseudo_fun (Class_Instance inst) then [str "()"] |
77707
a6a81f848135
More explicit type information in dictionary arguments.
haftmann
parents:
77703
diff
changeset
|
92 |
else map_filter (print_dicts is_pseudo_fun BR o snd) dss)) |
63303 | 93 |
| print_plain_dict is_pseudo_fun fxy (Dict_Var { var, index, length, ... }) = |
62539 | 94 |
[str (if length = 1 then Name.enforce_case true var ^ "_" |
95 |
else Name.enforce_case true var ^ string_of_int (index + 1) ^ "_")] |
|
38922 | 96 |
and print_dicts is_pseudo_fun = tuplify (print_dict is_pseudo_fun); |
33992 | 97 |
val print_dict_args = map_filter (fn (v, sort) => print_dicts (K false) BR |
63303 | 98 |
(map_index (fn (i, _) => Dict ([], |
99 |
Dict_Var { var = v, index = i, length = length sort, class = nth sort i, unique = true })) sort)); |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
100 |
fun print_term is_pseudo_fun some_thm vars fxy (IConst const) = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
101 |
print_app is_pseudo_fun some_thm vars fxy (const, []) |
35228 | 102 |
| print_term is_pseudo_fun some_thm vars fxy (IVar NONE) = |
31889 | 103 |
str "_" |
35228 | 104 |
| print_term is_pseudo_fun some_thm vars fxy (IVar (SOME v)) = |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
105 |
str (lookup_var vars v) |
35228 | 106 |
| print_term is_pseudo_fun some_thm vars fxy (t as t1 `$ t2) = |
28054 | 107 |
(case Code_Thingol.unfold_const_app t |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
108 |
of SOME app => print_app is_pseudo_fun some_thm vars fxy app |
35228 | 109 |
| NONE => brackify fxy [print_term is_pseudo_fun some_thm vars NOBR t1, |
110 |
print_term is_pseudo_fun some_thm vars BR t2]) |
|
111 |
| print_term is_pseudo_fun some_thm vars fxy (t as _ `|=> _) = |
|
28054 | 112 |
let |
31874 | 113 |
val (binds, t') = Code_Thingol.unfold_pat_abs t; |
33992 | 114 |
fun print_abs (pat, ty) = |
35228 | 115 |
print_bind is_pseudo_fun some_thm NOBR pat |
28054 | 116 |
#>> (fn p => concat [str "fn", p, str "=>"]); |
33992 | 117 |
val (ps, vars') = fold_map print_abs binds vars; |
35228 | 118 |
in brackets (ps @ [print_term is_pseudo_fun some_thm vars' NOBR t']) end |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
119 |
| print_term is_pseudo_fun some_thm vars fxy (ICase case_expr) = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
120 |
(case Code_Thingol.unfold_const_app (#primitive case_expr) |
55150 | 121 |
of SOME (app as ({ sym = Constant const, ... }, _)) => |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
122 |
if is_none (const_syntax const) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
123 |
then print_case is_pseudo_fun some_thm vars fxy case_expr |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
124 |
else print_app is_pseudo_fun some_thm vars fxy app |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
125 |
| NONE => print_case is_pseudo_fun some_thm vars fxy case_expr) |
77232 | 126 |
and print_app_expr is_pseudo_fun some_thm vars (app as ({ sym, dicts, dom, ... }, ts)) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
127 |
if is_constr sym then |
77232 | 128 |
let val wanted = length dom in |
129 |
if wanted < 2 orelse length ts = wanted |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
130 |
then (str o deresolve) sym |
38922 | 131 |
:: the_list (tuplify (print_term is_pseudo_fun some_thm vars) BR ts) |
77703
0262155d2743
more uniform approach towards satisfied applications
haftmann
parents:
77232
diff
changeset
|
132 |
else [print_term is_pseudo_fun some_thm vars BR (Code_Thingol.saturated_application wanted app)] |
33992 | 133 |
end |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
134 |
else if is_pseudo_fun sym |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
135 |
then (str o deresolve) sym @@ str "()" |
77232 | 136 |
else (str o deresolve) sym :: map_filter (print_dicts is_pseudo_fun BR) dicts |
35228 | 137 |
@ map (print_term is_pseudo_fun some_thm vars BR) ts |
138 |
and print_app is_pseudo_fun some_thm vars = gen_print_app (print_app_expr is_pseudo_fun) |
|
38923 | 139 |
(print_term is_pseudo_fun) const_syntax some_thm vars |
33992 | 140 |
and print_bind is_pseudo_fun = gen_print_bind (print_term is_pseudo_fun) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
141 |
and print_case is_pseudo_fun some_thm vars fxy { clauses = [], ... } = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
142 |
(concat o map str) ["raise", "Fail", "\"empty case\""] |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
143 |
| print_case is_pseudo_fun some_thm vars fxy (case_expr as { clauses = [_], ... }) = |
28054 | 144 |
let |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
145 |
val (binds, body) = Code_Thingol.unfold_let (ICase case_expr); |
47609 | 146 |
fun print_match ((pat, _), t) vars = |
28054 | 147 |
vars |
35228 | 148 |
|> print_bind is_pseudo_fun some_thm NOBR pat |
33992 | 149 |
|>> (fn p => semicolon [str "val", p, str "=", |
35228 | 150 |
print_term is_pseudo_fun some_thm vars NOBR t]) |
33992 | 151 |
val (ps, vars') = fold_map print_match binds vars; |
28054 | 152 |
in |
153 |
Pretty.chunks [ |
|
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
154 |
Pretty.block [str "let", Pretty.fbrk, Pretty.chunks ps], |
35228 | 155 |
Pretty.block [str "in", Pretty.fbrk, print_term is_pseudo_fun some_thm vars' NOBR body], |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
156 |
str "end" |
28054 | 157 |
] |
158 |
end |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
159 |
| print_case is_pseudo_fun some_thm vars fxy { term = t, typ = ty, clauses = clause :: clauses, ... } = |
28054 | 160 |
let |
33992 | 161 |
fun print_select delim (pat, body) = |
28054 | 162 |
let |
35228 | 163 |
val (p, vars') = print_bind is_pseudo_fun some_thm NOBR pat vars; |
28054 | 164 |
in |
35228 | 165 |
concat [str delim, p, str "=>", print_term is_pseudo_fun some_thm vars' NOBR body] |
28054 | 166 |
end; |
167 |
in |
|
31665 | 168 |
brackets ( |
28054 | 169 |
str "case" |
35228 | 170 |
:: print_term is_pseudo_fun some_thm vars NOBR t |
33992 | 171 |
:: print_select "of" clause |
172 |
:: map (print_select "|") clauses |
|
28054 | 173 |
) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
174 |
end; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
175 |
fun print_val_decl print_typscheme (sym, typscheme) = concat |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
176 |
[str "val", str (deresolve sym), str ":", print_typscheme typscheme]; |
33992 | 177 |
fun print_datatype_decl definer (tyco, (vs, cos)) = |
178 |
let |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
179 |
fun print_co ((co, _), []) = str (deresolve_const co) |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
180 |
| print_co ((co, _), tys) = concat [str (deresolve_const co), str "of", |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
181 |
enum " *" "" "" (map (print_typ (INFX (2, X))) tys)]; |
33992 | 182 |
in |
183 |
concat ( |
|
184 |
str definer |
|
55150 | 185 |
:: print_tyco_expr (Type_Constructor tyco, map ITyVar vs) |
33992 | 186 |
:: str "=" |
187 |
:: separate (str "|") (map print_co cos) |
|
188 |
) |
|
189 |
end; |
|
190 |
fun print_def is_pseudo_fun needs_typ definer |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
191 |
(ML_Function (const, (vs_ty as (vs, ty), eq :: eqs))) = |
29189 | 192 |
let |
35228 | 193 |
fun print_eqn definer ((ts, t), (some_thm, _)) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
194 |
let |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
195 |
val vars = reserved |
55145
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
196 |
|> intro_base_names_for (is_none o const_syntax) |
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
197 |
deresolve (t :: ts) |
75356 | 198 |
|> intro_vars (build (fold Code_Thingol.add_varnames ts)); |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
199 |
val prolog = if needs_typ then |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
200 |
concat [str definer, (str o deresolve_const) const, str ":", print_typ NOBR ty] |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
201 |
else (concat o map str) [definer, deresolve_const const]; |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
202 |
in |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
203 |
concat ( |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
204 |
prolog |
55150 | 205 |
:: (if is_pseudo_fun (Constant const) then [str "()"] |
33992 | 206 |
else print_dict_args vs |
35228 | 207 |
@ map (print_term is_pseudo_fun some_thm vars BR) ts) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
208 |
@ str "=" |
35228 | 209 |
@@ print_term is_pseudo_fun some_thm vars NOBR t |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
210 |
) |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
211 |
end |
33992 | 212 |
val shift = if null eqs then I else |
213 |
map (Pretty.block o single o Pretty.block o single); |
|
214 |
in pair |
|
55150 | 215 |
(print_val_decl print_typscheme (Constant const, vs_ty)) |
33992 | 216 |
((Pretty.block o Pretty.fbreaks o shift) ( |
217 |
print_eqn definer eq |
|
218 |
:: map (print_eqn "|") eqs |
|
219 |
)) |
|
29189 | 220 |
end |
33992 | 221 |
| print_def is_pseudo_fun _ definer |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
222 |
(ML_Instance (inst as (tyco, class), { vs, superinsts, inst_params, ... })) = |
29189 | 223 |
let |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
224 |
fun print_super_instance (super_class, x) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
225 |
concat [ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
226 |
(str o Long_Name.base_name o deresolve_classrel) (class, super_class), |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
227 |
str "=", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
228 |
print_dict is_pseudo_fun NOBR (Dict ([], Dict_Const ((tyco, super_class), x))) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
229 |
]; |
52519
598addf65209
explicit hint for domain of class parameters in instance statements
haftmann
parents:
52435
diff
changeset
|
230 |
fun print_classparam_instance ((classparam, (const, _)), (thm, _)) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
231 |
concat [ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
232 |
(str o Long_Name.base_name o deresolve_const) classparam, |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
233 |
str "=", |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37242
diff
changeset
|
234 |
print_app (K false) (SOME thm) reserved NOBR (const, []) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
235 |
]; |
33992 | 236 |
in pair |
237 |
(print_val_decl print_dicttypscheme |
|
55150 | 238 |
(Class_Instance inst, (vs, (class, tyco `%% map (ITyVar o fst) vs)))) |
33992 | 239 |
(concat ( |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
240 |
str definer |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
241 |
:: (str o deresolve_inst) inst |
55150 | 242 |
:: (if is_pseudo_fun (Class_Instance inst) then [str "()"] |
33992 | 243 |
else print_dict_args vs) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
244 |
@ str "=" |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
245 |
:: enum "," "{" "}" |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
246 |
(map print_super_instance superinsts |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
247 |
@ map print_classparam_instance inst_params) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
248 |
:: str ":" |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
249 |
@@ print_dicttyp (class, tyco `%% map (ITyVar o fst) vs) |
33992 | 250 |
)) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
251 |
end; |
55681 | 252 |
fun print_stmt _ (ML_Exc (const, (vs_ty, n))) = pair |
55150 | 253 |
[print_val_decl print_typscheme (Constant const, vs_ty)] |
33992 | 254 |
((semicolon o map str) ( |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
255 |
(if n = 0 then "val" else "fun") |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
256 |
:: deresolve_const const |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
257 |
:: replicate n "_" |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
258 |
@ "=" |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
259 |
:: "raise" |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
260 |
:: "Fail" |
68028 | 261 |
@@ print_sml_string const |
33992 | 262 |
)) |
55681 | 263 |
| print_stmt _ (ML_Val binding) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
264 |
let |
33992 | 265 |
val (sig_p, p) = print_def (K false) true "val" binding |
266 |
in pair |
|
267 |
[sig_p] |
|
268 |
(semicolon [p]) |
|
269 |
end |
|
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
270 |
| print_stmt _ (ML_Funs ((export, binding) :: exports_bindings, pseudo_funs)) = |
33992 | 271 |
let |
272 |
val print_def' = print_def (member (op =) pseudo_funs) false; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
273 |
fun print_pseudo_fun sym = concat [ |
29189 | 274 |
str "val", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
275 |
(str o deresolve) sym, |
29189 | 276 |
str "=", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
277 |
(str o deresolve) sym, |
29189 | 278 |
str "();" |
279 |
]; |
|
33992 | 280 |
val (sig_ps, (ps, p)) = (apsnd split_last o split_list) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
281 |
(print_def' "fun" binding :: map (print_def' "and" o snd) exports_bindings); |
33992 | 282 |
val pseudo_ps = map print_pseudo_fun pseudo_funs; |
283 |
in pair |
|
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
284 |
(map_filter (fn (export, p) => if Code_Namespace.not_private export then SOME p else NONE) |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
285 |
((export :: map fst exports_bindings) ~~ sig_ps)) |
33992 | 286 |
(Pretty.chunks (ps @ semicolon [p] :: pseudo_ps)) |
287 |
end |
|
55681 | 288 |
| print_stmt _ (ML_Datas [(tyco, (vs, []))]) = |
33992 | 289 |
let |
55150 | 290 |
val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs); |
33992 | 291 |
in |
292 |
pair |
|
293 |
[concat [str "type", ty_p]] |
|
55681 | 294 |
(semicolon [str "datatype", ty_p, str "=", str "EMPTY__"]) |
33992 | 295 |
end |
55681 | 296 |
| print_stmt export (ML_Datas (data :: datas)) = |
28054 | 297 |
let |
55681 | 298 |
val decl_ps = print_datatype_decl "datatype" data |
33992 | 299 |
:: map (print_datatype_decl "and") datas; |
55681 | 300 |
val (ps, p) = split_last decl_ps; |
33992 | 301 |
in pair |
55681 | 302 |
(if Code_Namespace.is_public export |
303 |
then decl_ps |
|
304 |
else map (fn (tyco, (vs, _)) => |
|
305 |
concat [str "type", print_tyco_expr (Type_Constructor tyco, map ITyVar vs)]) |
|
306 |
(data :: datas)) |
|
33992 | 307 |
(Pretty.chunks (ps @| semicolon [p])) |
308 |
end |
|
55681 | 309 |
| print_stmt export (ML_Class (class, (v, (classrels, classparams)))) = |
28054 | 310 |
let |
33992 | 311 |
fun print_field s p = concat [str s, str ":", p]; |
312 |
fun print_proj s p = semicolon |
|
313 |
(map str ["val", s, "=", "#" ^ s, ":"] @| p); |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
314 |
fun print_super_class_decl (classrel as (_, super_class)) = |
33992 | 315 |
print_val_decl print_dicttypscheme |
55150 | 316 |
(Class_Relation classrel, ([(v, [class])], (super_class, ITyVar v))); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
317 |
fun print_super_class_field (classrel as (_, super_class)) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
318 |
print_field (deresolve_classrel classrel) (print_dicttyp (super_class, ITyVar v)); |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
319 |
fun print_super_class_proj (classrel as (_, super_class)) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
320 |
print_proj (deresolve_classrel classrel) |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37242
diff
changeset
|
321 |
(print_dicttypscheme ([(v, [class])], (super_class, ITyVar v))); |
33992 | 322 |
fun print_classparam_decl (classparam, ty) = |
323 |
print_val_decl print_typscheme |
|
55150 | 324 |
(Constant classparam, ([(v, [class])], ty)); |
33992 | 325 |
fun print_classparam_field (classparam, ty) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
326 |
print_field (deresolve_const classparam) (print_typ NOBR ty); |
33992 | 327 |
fun print_classparam_proj (classparam, ty) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
328 |
print_proj (deresolve_const classparam) |
33992 | 329 |
(print_typscheme ([(v, [class])], ty)); |
330 |
in pair |
|
331 |
(concat [str "type", print_dicttyp (class, ITyVar v)] |
|
55681 | 332 |
:: (if Code_Namespace.is_public export |
333 |
then map print_super_class_decl classrels |
|
334 |
@ map print_classparam_decl classparams |
|
335 |
else [])) |
|
33992 | 336 |
(Pretty.chunks ( |
28054 | 337 |
concat [ |
55681 | 338 |
str "type", |
339 |
print_dicttyp (class, ITyVar v), |
|
28054 | 340 |
str "=", |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
341 |
enum "," "{" "};" ( |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
342 |
map print_super_class_field classrels |
33992 | 343 |
@ map print_classparam_field classparams |
28054 | 344 |
) |
345 |
] |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
346 |
:: map print_super_class_proj classrels |
33992 | 347 |
@ map print_classparam_proj classparams |
348 |
)) |
|
28054 | 349 |
end; |
33992 | 350 |
in print_stmt end; |
28054 | 351 |
|
55677 | 352 |
fun print_sml_module name decls body = |
38933
bd77e092f67c
avoid strange special treatment of empty module names
haftmann
parents:
38928
diff
changeset
|
353 |
Pretty.chunks2 ( |
55677 | 354 |
Pretty.chunks [ |
355 |
str ("structure " ^ name ^ " : sig"), |
|
356 |
(indent 2 o Pretty.chunks) decls, |
|
357 |
str "end = struct" |
|
358 |
] |
|
33992 | 359 |
:: body |
360 |
@| str ("end; (*struct " ^ name ^ "*)") |
|
28054 | 361 |
); |
362 |
||
28064 | 363 |
val literals_sml = Literals { |
68028 | 364 |
literal_string = print_sml_string, |
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34246
diff
changeset
|
365 |
literal_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)", |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
366 |
literal_list = enum "," "[" "]", |
28064 | 367 |
infix_cons = (7, "::") |
368 |
}; |
|
369 |
||
28054 | 370 |
|
371 |
(** OCaml serializer **) |
|
372 |
||
68028 | 373 |
val print_ocaml_string = |
374 |
let |
|
375 |
fun chr i = |
|
376 |
let |
|
377 |
val xs = string_of_int i; |
|
378 |
val ys = replicate_string (3 - length (raw_explode xs)) "0"; |
|
379 |
in "\\" ^ ys ^ xs end; |
|
380 |
fun char c = |
|
381 |
let |
|
382 |
val i = ord c; |
|
383 |
val s = |
|
384 |
if i >= 128 then error "non-ASCII byte in OCaml string literal" |
|
385 |
else if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126 |
|
386 |
then chr i else c |
|
387 |
in s end; |
|
388 |
in quote o translate_string char end; |
|
389 |
||
50625 | 390 |
fun print_ocaml_stmt tyco_syntax const_syntax reserved is_constr deresolve = |
28054 | 391 |
let |
55150 | 392 |
val deresolve_const = deresolve o Constant; |
393 |
val deresolve_classrel = deresolve o Class_Relation; |
|
394 |
val deresolve_inst = deresolve o Class_Instance; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
395 |
fun print_tyco_expr (sym, []) = (str o deresolve) sym |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
396 |
| print_tyco_expr (sym, [ty]) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
397 |
concat [print_typ BR ty, (str o deresolve) sym] |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
398 |
| print_tyco_expr (sym, tys) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
399 |
concat [enum "," "(" ")" (map (print_typ BR) tys), (str o deresolve) sym] |
38923 | 400 |
and print_typ fxy (tyco `%% tys) = (case tyco_syntax tyco |
55150 | 401 |
of NONE => print_tyco_expr (Type_Constructor tyco, tys) |
47576 | 402 |
| SOME (_, print) => print print_typ fxy tys) |
33992 | 403 |
| print_typ fxy (ITyVar v) = str ("'" ^ v); |
55150 | 404 |
fun print_dicttyp (class, ty) = print_tyco_expr (Type_Class class, [ty]); |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
405 |
fun print_typscheme_prefix (vs, p) = enum " ->" "" "" |
33992 | 406 |
(map_filter (fn (v, sort) => |
407 |
(print_product (fn class => print_dicttyp (class, ITyVar v)) sort)) vs @| p); |
|
408 |
fun print_typscheme (vs, ty) = print_typscheme_prefix (vs, print_typ NOBR ty); |
|
409 |
fun print_dicttypscheme (vs, class_ty) = print_typscheme_prefix (vs, print_dicttyp class_ty); |
|
41100
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
haftmann
parents:
40627
diff
changeset
|
410 |
val print_classrels = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
411 |
fold_rev (fn classrel => fn p => Pretty.block [p, str ".", (str o deresolve_classrel) classrel]) |
41118
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
412 |
fun print_dict is_pseudo_fun fxy (Dict (classrels, x)) = |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
413 |
print_plain_dict is_pseudo_fun fxy x |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
414 |
|> print_classrels classrels |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
415 |
and print_plain_dict is_pseudo_fun fxy (Dict_Const (inst, dss)) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
416 |
brackify BR ((str o deresolve_inst) inst :: |
55150 | 417 |
(if is_pseudo_fun (Class_Instance inst) then [str "()"] |
77707
a6a81f848135
More explicit type information in dictionary arguments.
haftmann
parents:
77703
diff
changeset
|
418 |
else map_filter (print_dicts is_pseudo_fun BR o snd) dss)) |
63303 | 419 |
| print_plain_dict is_pseudo_fun fxy (Dict_Var { var, index, length, ... }) = |
62539 | 420 |
str (if length = 1 then "_" ^ Name.enforce_case true var |
421 |
else "_" ^ Name.enforce_case true var ^ string_of_int (index + 1)) |
|
38922 | 422 |
and print_dicts is_pseudo_fun = tuplify (print_dict is_pseudo_fun); |
33992 | 423 |
val print_dict_args = map_filter (fn (v, sort) => print_dicts (K false) BR |
63303 | 424 |
(map_index (fn (i, _) => Dict ([], |
425 |
Dict_Var { var = v, index = i, length = length sort, class = nth sort i, unique = true })) sort)); |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
426 |
fun print_term is_pseudo_fun some_thm vars fxy (IConst const) = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
427 |
print_app is_pseudo_fun some_thm vars fxy (const, []) |
35228 | 428 |
| print_term is_pseudo_fun some_thm vars fxy (IVar NONE) = |
31889 | 429 |
str "_" |
35228 | 430 |
| print_term is_pseudo_fun some_thm vars fxy (IVar (SOME v)) = |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
431 |
str (lookup_var vars v) |
35228 | 432 |
| print_term is_pseudo_fun some_thm vars fxy (t as t1 `$ t2) = |
28054 | 433 |
(case Code_Thingol.unfold_const_app t |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
434 |
of SOME app => print_app is_pseudo_fun some_thm vars fxy app |
35228 | 435 |
| NONE => brackify fxy [print_term is_pseudo_fun some_thm vars NOBR t1, |
436 |
print_term is_pseudo_fun some_thm vars BR t2]) |
|
437 |
| print_term is_pseudo_fun some_thm vars fxy (t as _ `|=> _) = |
|
28054 | 438 |
let |
31874 | 439 |
val (binds, t') = Code_Thingol.unfold_pat_abs t; |
35228 | 440 |
val (ps, vars') = fold_map (print_bind is_pseudo_fun some_thm BR o fst) binds vars; |
441 |
in brackets (str "fun" :: ps @ str "->" @@ print_term is_pseudo_fun some_thm vars' NOBR t') end |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
442 |
| print_term is_pseudo_fun some_thm vars fxy (ICase case_expr) = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
443 |
(case Code_Thingol.unfold_const_app (#primitive case_expr) |
55150 | 444 |
of SOME (app as ({ sym = Constant const, ... }, _)) => |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
445 |
if is_none (const_syntax const) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
446 |
then print_case is_pseudo_fun some_thm vars fxy case_expr |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
447 |
else print_app is_pseudo_fun some_thm vars fxy app |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
448 |
| NONE => print_case is_pseudo_fun some_thm vars fxy case_expr) |
77232 | 449 |
and print_app_expr is_pseudo_fun some_thm vars (app as ({ sym, dicts, dom, ... }, ts)) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
450 |
if is_constr sym then |
77232 | 451 |
let val wanted = length dom in |
452 |
if length ts = wanted |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
453 |
then (str o deresolve) sym |
38922 | 454 |
:: the_list (tuplify (print_term is_pseudo_fun some_thm vars) BR ts) |
77703
0262155d2743
more uniform approach towards satisfied applications
haftmann
parents:
77232
diff
changeset
|
455 |
else [print_term is_pseudo_fun some_thm vars BR (Code_Thingol.saturated_application wanted app)] |
33992 | 456 |
end |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
457 |
else if is_pseudo_fun sym |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
458 |
then (str o deresolve) sym @@ str "()" |
77232 | 459 |
else (str o deresolve) sym :: map_filter (print_dicts is_pseudo_fun BR) dicts |
35228 | 460 |
@ map (print_term is_pseudo_fun some_thm vars BR) ts |
461 |
and print_app is_pseudo_fun some_thm vars = gen_print_app (print_app_expr is_pseudo_fun) |
|
38923 | 462 |
(print_term is_pseudo_fun) const_syntax some_thm vars |
33992 | 463 |
and print_bind is_pseudo_fun = gen_print_bind (print_term is_pseudo_fun) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
464 |
and print_case is_pseudo_fun some_thm vars fxy { clauses = [], ... } = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
465 |
(concat o map str) ["failwith", "\"empty case\""] |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
466 |
| print_case is_pseudo_fun some_thm vars fxy (case_expr as { clauses = [_], ... }) = |
28054 | 467 |
let |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
468 |
val (binds, body) = Code_Thingol.unfold_let (ICase case_expr); |
47576 | 469 |
fun print_let ((pat, _), t) vars = |
28054 | 470 |
vars |
35228 | 471 |
|> print_bind is_pseudo_fun some_thm NOBR pat |
28054 | 472 |
|>> (fn p => concat |
35228 | 473 |
[str "let", p, str "=", print_term is_pseudo_fun some_thm vars NOBR t, str "in"]) |
33992 | 474 |
val (ps, vars') = fold_map print_let binds vars; |
31665 | 475 |
in |
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59456
diff
changeset
|
476 |
brackets [Pretty.chunks ps, print_term is_pseudo_fun some_thm vars' NOBR body] |
31665 | 477 |
end |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
478 |
| print_case is_pseudo_fun some_thm vars fxy { term = t, typ = ty, clauses = clause :: clauses, ... } = |
28054 | 479 |
let |
33992 | 480 |
fun print_select delim (pat, body) = |
28054 | 481 |
let |
35228 | 482 |
val (p, vars') = print_bind is_pseudo_fun some_thm NOBR pat vars; |
483 |
in concat [str delim, p, str "->", print_term is_pseudo_fun some_thm vars' NOBR body] end; |
|
28054 | 484 |
in |
31665 | 485 |
brackets ( |
28054 | 486 |
str "match" |
35228 | 487 |
:: print_term is_pseudo_fun some_thm vars NOBR t |
33992 | 488 |
:: print_select "with" clause |
489 |
:: map (print_select "|") clauses |
|
28054 | 490 |
) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
491 |
end; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
492 |
fun print_val_decl print_typscheme (sym, typscheme) = concat |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
493 |
[str "val", str (deresolve sym), str ":", print_typscheme typscheme]; |
33992 | 494 |
fun print_datatype_decl definer (tyco, (vs, cos)) = |
495 |
let |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
496 |
fun print_co ((co, _), []) = str (deresolve_const co) |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
497 |
| print_co ((co, _), tys) = concat [str (deresolve_const co), str "of", |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
498 |
enum " *" "" "" (map (print_typ (INFX (2, X))) tys)]; |
33992 | 499 |
in |
500 |
concat ( |
|
501 |
str definer |
|
55150 | 502 |
:: print_tyco_expr (Type_Constructor tyco, map ITyVar vs) |
33992 | 503 |
:: str "=" |
504 |
:: separate (str "|") (map print_co cos) |
|
505 |
) |
|
506 |
end; |
|
507 |
fun print_def is_pseudo_fun needs_typ definer |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
508 |
(ML_Function (const, (vs_ty as (vs, ty), eqs))) = |
28054 | 509 |
let |
35228 | 510 |
fun print_eqn ((ts, t), (some_thm, _)) = |
28054 | 511 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
512 |
val vars = reserved |
55145
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
513 |
|> intro_base_names_for (is_none o const_syntax) |
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
514 |
deresolve (t :: ts) |
75356 | 515 |
|> intro_vars (build (fold Code_Thingol.add_varnames ts)); |
28054 | 516 |
in concat [ |
38778 | 517 |
(Pretty.block o commas) |
35228 | 518 |
(map (print_term is_pseudo_fun some_thm vars NOBR) ts), |
28054 | 519 |
str "->", |
35228 | 520 |
print_term is_pseudo_fun some_thm vars NOBR t |
28054 | 521 |
] end; |
35228 | 522 |
fun print_eqns is_pseudo [((ts, t), (some_thm, _))] = |
28054 | 523 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
524 |
val vars = reserved |
55145
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
525 |
|> intro_base_names_for (is_none o const_syntax) |
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
526 |
deresolve (t :: ts) |
75356 | 527 |
|> intro_vars (build (fold Code_Thingol.add_varnames ts)); |
28054 | 528 |
in |
529 |
concat ( |
|
29189 | 530 |
(if is_pseudo then [str "()"] |
35228 | 531 |
else map (print_term is_pseudo_fun some_thm vars BR) ts) |
28054 | 532 |
@ str "=" |
35228 | 533 |
@@ print_term is_pseudo_fun some_thm vars NOBR t |
28054 | 534 |
) |
535 |
end |
|
33992 | 536 |
| print_eqns _ ((eq as (([_], _), _)) :: eqs) = |
28054 | 537 |
Pretty.block ( |
538 |
str "=" |
|
539 |
:: Pretty.brk 1 |
|
540 |
:: str "function" |
|
541 |
:: Pretty.brk 1 |
|
33992 | 542 |
:: print_eqn eq |
28054 | 543 |
:: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] |
33992 | 544 |
o single o print_eqn) eqs |
28054 | 545 |
) |
33992 | 546 |
| print_eqns _ (eqs as eq :: eqs') = |
28054 | 547 |
let |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
548 |
val vars = reserved |
55145
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
549 |
|> intro_base_names_for (is_none o const_syntax) |
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52519
diff
changeset
|
550 |
deresolve (map (snd o fst) eqs) |
32924
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
haftmann
parents:
32913
diff
changeset
|
551 |
val dummy_parms = (map str o aux_params vars o map (fst o fst)) eqs; |
28054 | 552 |
in |
553 |
Pretty.block ( |
|
554 |
Pretty.breaks dummy_parms |
|
555 |
@ Pretty.brk 1 |
|
556 |
:: str "=" |
|
557 |
:: Pretty.brk 1 |
|
558 |
:: str "match" |
|
559 |
:: Pretty.brk 1 |
|
38778 | 560 |
:: (Pretty.block o commas) dummy_parms |
28054 | 561 |
:: Pretty.brk 1 |
562 |
:: str "with" |
|
563 |
:: Pretty.brk 1 |
|
33992 | 564 |
:: print_eqn eq |
28054 | 565 |
:: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] |
33992 | 566 |
o single o print_eqn) eqs' |
28054 | 567 |
) |
568 |
end; |
|
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
569 |
val prolog = if needs_typ then |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
570 |
concat [str definer, (str o deresolve_const) const, str ":", print_typ NOBR ty] |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
571 |
else (concat o map str) [definer, deresolve_const const]; |
33992 | 572 |
in pair |
55150 | 573 |
(print_val_decl print_typscheme (Constant const, vs_ty)) |
33992 | 574 |
(concat ( |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
575 |
prolog |
33992 | 576 |
:: print_dict_args vs |
55150 | 577 |
@| print_eqns (is_pseudo_fun (Constant const)) eqs |
33992 | 578 |
)) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
579 |
end |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37242
diff
changeset
|
580 |
| print_def is_pseudo_fun _ definer |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
581 |
(ML_Instance (inst as (tyco, class), { vs, superinsts, inst_params, ... })) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
582 |
let |
77707
a6a81f848135
More explicit type information in dictionary arguments.
haftmann
parents:
77703
diff
changeset
|
583 |
fun print_super_instance (super_class, dss) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
584 |
concat [ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
585 |
(str o deresolve_classrel) (class, super_class), |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
586 |
str "=", |
77707
a6a81f848135
More explicit type information in dictionary arguments.
haftmann
parents:
77703
diff
changeset
|
587 |
print_dict is_pseudo_fun NOBR (Dict ([], Dict_Const ((tyco, super_class), dss))) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
588 |
]; |
52519
598addf65209
explicit hint for domain of class parameters in instance statements
haftmann
parents:
52435
diff
changeset
|
589 |
fun print_classparam_instance ((classparam, (const, _)), (thm, _)) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
590 |
concat [ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
591 |
(str o deresolve_const) classparam, |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
592 |
str "=", |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37242
diff
changeset
|
593 |
print_app (K false) (SOME thm) reserved NOBR (const, []) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
594 |
]; |
33992 | 595 |
in pair |
596 |
(print_val_decl print_dicttypscheme |
|
55150 | 597 |
(Class_Instance inst, (vs, (class, tyco `%% map (ITyVar o fst) vs)))) |
33992 | 598 |
(concat ( |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
599 |
str definer |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
600 |
:: (str o deresolve_inst) inst |
55150 | 601 |
:: (if is_pseudo_fun (Class_Instance inst) then [str "()"] |
43343
e83695ea0e0a
resolving an issue with class instances that are pseudo functions in the OCaml code serializer
bulwahn
parents:
42599
diff
changeset
|
602 |
else print_dict_args vs) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
603 |
@ str "=" |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
604 |
@@ brackets [ |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
605 |
enum_default "()" ";" "{" "}" (map print_super_instance superinsts |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
606 |
@ map print_classparam_instance inst_params), |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
607 |
str ":", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
608 |
print_dicttyp (class, tyco `%% map (ITyVar o fst) vs) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
609 |
] |
33992 | 610 |
)) |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
611 |
end; |
55681 | 612 |
fun print_stmt _ (ML_Exc (const, (vs_ty, n))) = pair |
55150 | 613 |
[print_val_decl print_typscheme (Constant const, vs_ty)] |
33992 | 614 |
((doublesemicolon o map str) ( |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
615 |
"let" |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
616 |
:: deresolve_const const |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
617 |
:: replicate n "_" |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
618 |
@ "=" |
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
619 |
:: "failwith" |
68028 | 620 |
@@ print_ocaml_string const |
33992 | 621 |
)) |
55681 | 622 |
| print_stmt _ (ML_Val binding) = |
33636
a9bb3f063773
accomplish mutual recursion between fun and inst
haftmann
parents:
33519
diff
changeset
|
623 |
let |
33992 | 624 |
val (sig_p, p) = print_def (K false) true "let" binding |
625 |
in pair |
|
626 |
[sig_p] |
|
627 |
(doublesemicolon [p]) |
|
628 |
end |
|
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
629 |
| print_stmt _ (ML_Funs ((export, binding) :: exports_bindings, pseudo_funs)) = |
33992 | 630 |
let |
631 |
val print_def' = print_def (member (op =) pseudo_funs) false; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
632 |
fun print_pseudo_fun sym = concat [ |
29189 | 633 |
str "let", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
634 |
(str o deresolve) sym, |
29189 | 635 |
str "=", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
636 |
(str o deresolve) sym, |
29189 | 637 |
str "();;" |
638 |
]; |
|
33992 | 639 |
val (sig_ps, (ps, p)) = (apsnd split_last o split_list) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
640 |
(print_def' "let rec" binding :: map (print_def' "and" o snd) exports_bindings); |
33992 | 641 |
val pseudo_ps = map print_pseudo_fun pseudo_funs; |
642 |
in pair |
|
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
643 |
(map_filter (fn (export, p) => if Code_Namespace.not_private export then SOME p else NONE) |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
644 |
((export :: map fst exports_bindings) ~~ sig_ps)) |
33992 | 645 |
(Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps)) |
646 |
end |
|
55681 | 647 |
| print_stmt _ (ML_Datas [(tyco, (vs, []))]) = |
33992 | 648 |
let |
55150 | 649 |
val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs); |
33992 | 650 |
in |
651 |
pair |
|
652 |
[concat [str "type", ty_p]] |
|
55681 | 653 |
(doublesemicolon [str "type", ty_p, str "=", str "EMPTY__"]) |
33992 | 654 |
end |
55681 | 655 |
| print_stmt export (ML_Datas (data :: datas)) = |
28054 | 656 |
let |
55681 | 657 |
val decl_ps = print_datatype_decl "type" data |
33992 | 658 |
:: map (print_datatype_decl "and") datas; |
55681 | 659 |
val (ps, p) = split_last decl_ps; |
33992 | 660 |
in pair |
55681 | 661 |
(if Code_Namespace.is_public export |
662 |
then decl_ps |
|
663 |
else map (fn (tyco, (vs, _)) => |
|
664 |
concat [str "type", print_tyco_expr (Type_Constructor tyco, map ITyVar vs)]) |
|
665 |
(data :: datas)) |
|
33992 | 666 |
(Pretty.chunks (ps @| doublesemicolon [p])) |
667 |
end |
|
55681 | 668 |
| print_stmt export (ML_Class (class, (v, (classrels, classparams)))) = |
28054 | 669 |
let |
33992 | 670 |
fun print_field s p = concat [str s, str ":", p]; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
671 |
fun print_super_class_field (classrel as (_, super_class)) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
672 |
print_field (deresolve_classrel classrel) (print_dicttyp (super_class, ITyVar v)); |
33992 | 673 |
fun print_classparam_decl (classparam, ty) = |
674 |
print_val_decl print_typscheme |
|
55150 | 675 |
(Constant classparam, ([(v, [class])], ty)); |
33992 | 676 |
fun print_classparam_field (classparam, ty) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
677 |
print_field (deresolve_const classparam) (print_typ NOBR ty); |
56812 | 678 |
val w = "_" ^ Name.enforce_case true v; |
33992 | 679 |
fun print_classparam_proj (classparam, _) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
680 |
(concat o map str) ["let", deresolve_const classparam, w, "=", |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
681 |
w ^ "." ^ deresolve_const classparam ^ ";;"]; |
33992 | 682 |
val type_decl_p = concat [ |
55682 | 683 |
str "type", |
684 |
print_dicttyp (class, ITyVar v), |
|
28054 | 685 |
str "=", |
33992 | 686 |
enum_default "unit" ";" "{" "}" ( |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
687 |
map print_super_class_field classrels |
33992 | 688 |
@ map print_classparam_field classparams |
689 |
) |
|
28054 | 690 |
]; |
33992 | 691 |
in pair |
71803
14914ae80f70
temporarily revert change which does not work as expected
haftmann
parents:
71798
diff
changeset
|
692 |
(if Code_Namespace.is_public export |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
693 |
then type_decl_p :: map print_classparam_decl classparams |
66325
fd28cb6e6f2c
work around weakness in export calculation when generating OCaml code
haftmann
parents:
63303
diff
changeset
|
694 |
else if null classrels andalso null classparams |
fd28cb6e6f2c
work around weakness in export calculation when generating OCaml code
haftmann
parents:
63303
diff
changeset
|
695 |
then [type_decl_p] (*work around weakness in export calculation*) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
696 |
else [concat [str "type", print_dicttyp (class, ITyVar v)]]) |
33992 | 697 |
(Pretty.chunks ( |
698 |
doublesemicolon [type_decl_p] |
|
699 |
:: map print_classparam_proj classparams |
|
700 |
)) |
|
701 |
end; |
|
702 |
in print_stmt end; |
|
28054 | 703 |
|
55677 | 704 |
fun print_ocaml_module name decls body = |
38933
bd77e092f67c
avoid strange special treatment of empty module names
haftmann
parents:
38928
diff
changeset
|
705 |
Pretty.chunks2 ( |
55677 | 706 |
Pretty.chunks [ |
707 |
str ("module " ^ name ^ " : sig"), |
|
708 |
(indent 2 o Pretty.chunks) decls, |
|
709 |
str "end = struct" |
|
710 |
] |
|
33992 | 711 |
:: body |
712 |
@| str ("end;; (*struct " ^ name ^ "*)") |
|
28054 | 713 |
); |
714 |
||
28064 | 715 |
val literals_ocaml = let |
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34246
diff
changeset
|
716 |
fun numeral_ocaml k = if k < 0 |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69743
diff
changeset
|
717 |
then "(Z.neg " ^ numeral_ocaml (~ k) ^ ")" |
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34246
diff
changeset
|
718 |
else if k <= 1073741823 |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69743
diff
changeset
|
719 |
then "(Z.of_int " ^ string_of_int k ^ ")" |
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69743
diff
changeset
|
720 |
else "(Z.of_string " ^ quote (string_of_int k) ^ ")" |
28064 | 721 |
in Literals { |
68028 | 722 |
literal_string = print_ocaml_string, |
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34246
diff
changeset
|
723 |
literal_numeral = numeral_ocaml, |
34178
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
haftmann
parents:
34071
diff
changeset
|
724 |
literal_list = enum ";" "[" "]", |
28064 | 725 |
infix_cons = (6, "::") |
726 |
} end; |
|
727 |
||
728 |
||
28054 | 729 |
|
730 |
(** SML/OCaml generic part **) |
|
731 |
||
55681 | 732 |
fun ml_program_of_program ctxt module_name reserved identifiers = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
733 |
let |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
734 |
fun namify_const upper base (nsp_const, nsp_type) = |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
735 |
let |
56826
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
haftmann
parents:
56812
diff
changeset
|
736 |
val (base', nsp_const') = Name.variant (Name.enforce_case upper base) nsp_const |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
737 |
in (base', (nsp_const', nsp_type)) end; |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
738 |
fun namify_type base (nsp_const, nsp_type) = |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
739 |
let |
56826
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
haftmann
parents:
56812
diff
changeset
|
740 |
val (base', nsp_type') = Name.variant (Name.enforce_case false base) nsp_type |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
741 |
in (base', (nsp_const, nsp_type')) end; |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
742 |
fun namify_stmt (Code_Thingol.Fun _) = namify_const false |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
743 |
| namify_stmt (Code_Thingol.Datatype _) = namify_type |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
744 |
| namify_stmt (Code_Thingol.Datatypecons _) = namify_const true |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
745 |
| namify_stmt (Code_Thingol.Class _) = namify_type |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
746 |
| namify_stmt (Code_Thingol.Classrel _) = namify_const false |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
747 |
| namify_stmt (Code_Thingol.Classparam _) = namify_const false |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
748 |
| namify_stmt (Code_Thingol.Classinst _) = namify_const false; |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
749 |
fun ml_binding_of_stmt (sym as Constant const, (export, Code_Thingol.Fun ((tysm as (vs, ty), raw_eqs), _))) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
750 |
let |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
751 |
val eqs = filter (snd o snd) raw_eqs; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
752 |
val (eqs', some_sym) = if null (filter_out (null o snd) vs) then case eqs |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
753 |
of [(([], t), some_thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
754 |
then ([(([IVar (SOME "x")], t `$ IVar (SOME "x")), some_thm)], NONE) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
755 |
else (eqs, SOME (sym, member (op =) (Code_Thingol.add_constsyms t []) sym)) |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
756 |
| _ => (eqs, NONE) |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
757 |
else (eqs, NONE) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
758 |
in ((export, ML_Function (const, (tysm, eqs'))), some_sym) end |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
759 |
| ml_binding_of_stmt (sym as Class_Instance inst, (export, Code_Thingol.Classinst (stmt as { vs, ... }))) = |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
760 |
((export, ML_Instance (inst, stmt)), |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
761 |
if forall (null o snd) vs then SOME (sym, false) else NONE) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
762 |
| ml_binding_of_stmt (sym, _) = |
52138
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents:
51143
diff
changeset
|
763 |
error ("Binding block containing illegal statement: " ^ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
764 |
Code_Symbol.quote ctxt sym) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
765 |
fun modify_fun (sym, (export, stmt)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
766 |
let |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
767 |
val ((export', binding), some_value_sym) = ml_binding_of_stmt (sym, (export, stmt)); |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
768 |
val ml_stmt = case binding |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
769 |
of ML_Function (const, ((vs, ty), [])) => |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
770 |
ML_Exc (const, ((vs, ty), |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
771 |
(length o filter_out (null o snd)) vs + (length o fst o Code_Thingol.unfold_fun) ty)) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
772 |
| _ => case some_value_sym |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
773 |
of NONE => ML_Funs ([(export', binding)], []) |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
774 |
| SOME (sym, true) => ML_Funs ([(export, binding)], [sym]) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
775 |
| SOME (sym, false) => ML_Val binding |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
776 |
in SOME (export, ml_stmt) end; |
39031 | 777 |
fun modify_funs stmts = single (SOME |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
778 |
(Code_Namespace.Opaque, ML_Funs (map_split ml_binding_of_stmt stmts |> (apsnd o map_filter o Option.map) fst))) |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
779 |
fun modify_datatypes stmts = |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
780 |
let |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
781 |
val datas = map_filter |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
782 |
(fn (Type_Constructor tyco, (export, Code_Thingol.Datatype stmt)) => SOME (export, (tyco, stmt)) | _ => NONE) stmts |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
783 |
in |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
784 |
if null datas then [] (*for abstract types wrt. code_reflect*) |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
785 |
else datas |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
786 |
|> split_list |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
787 |
|> apfst Code_Namespace.join_exports |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
788 |
|> apsnd ML_Datas |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
789 |
|> SOME |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
790 |
|> single |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63024
diff
changeset
|
791 |
end; |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
792 |
fun modify_class stmts = |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
793 |
the_single (map_filter |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
794 |
(fn (Type_Class class, (export, Code_Thingol.Class stmt)) => SOME (export, (class, stmt)) | _ => NONE) stmts) |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
795 |
|> apsnd ML_Class |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
796 |
|> SOME |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
797 |
|> single; |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
798 |
fun modify_stmts ([stmt as (_, (_, stmt' as Code_Thingol.Fun _))]) = |
39059
3a11a667af75
restored and added surpression of case combinators
haftmann
parents:
39058
diff
changeset
|
799 |
if Code_Thingol.is_case stmt' then [] else [modify_fun stmt] |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
800 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Fun _)) :: _)) = |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
801 |
modify_funs (filter_out (Code_Thingol.is_case o snd o snd) stmts) |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
802 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Datatypecons _)) :: _)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
803 |
modify_datatypes stmts |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
804 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Datatype _)) :: _)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
805 |
modify_datatypes stmts |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
806 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Class _)) :: _)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
807 |
modify_class stmts |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
808 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Classrel _)) :: _)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
809 |
modify_class stmts |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
810 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Classparam _)) :: _)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
811 |
modify_class stmts |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
812 |
| modify_stmts ([stmt as (_, (_, Code_Thingol.Classinst _))]) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
813 |
[modify_fun stmt] |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
814 |
| modify_stmts ((stmts as (_, (_, Code_Thingol.Classinst _)) :: _)) = |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
815 |
modify_funs stmts |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
816 |
| modify_stmts stmts = error ("Illegal mutual dependencies: " ^ |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
817 |
(Library.commas o map (Code_Symbol.quote ctxt o fst)) stmts); |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
818 |
in |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
819 |
Code_Namespace.hierarchical_program ctxt { |
52138
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents:
51143
diff
changeset
|
820 |
module_name = module_name, reserved = reserved, identifiers = identifiers, |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
821 |
empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt, |
55776
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents:
55684
diff
changeset
|
822 |
cyclic_modules = false, class_transitive = true, |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
823 |
class_relation_public = true, empty_data = (), |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
824 |
memorize_data = K I, modify_stmts = modify_stmts } |
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
825 |
end; |
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
826 |
|
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69207
diff
changeset
|
827 |
fun serialize_ml print_ml_module print_ml_stmt ml_extension ctxt |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
828 |
{ module_name, reserved_syms, identifiers, includes, |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69207
diff
changeset
|
829 |
class_syntax, tyco_syntax, const_syntax } program exports = |
28054 | 830 |
let |
39147 | 831 |
|
832 |
(* build program *) |
|
39028
0dd6c5a0beef
use code_namespace module for SML and OCaml code
haftmann
parents:
38966
diff
changeset
|
833 |
val { deresolver, hierarchical_program = ml_program } = |
55681 | 834 |
ml_program_of_program ctxt module_name (Name.make_context reserved_syms) |
55683 | 835 |
identifiers exports program; |
39147 | 836 |
|
837 |
(* print statements *) |
|
55679 | 838 |
fun print_stmt prefix_fragments (_, (export, stmt)) = print_ml_stmt |
47576 | 839 |
tyco_syntax const_syntax (make_vars reserved_syms) |
55681 | 840 |
(Code_Thingol.is_constr program) (deresolver prefix_fragments) export stmt |
841 |
|> apfst (fn decl => if Code_Namespace.not_private export then SOME decl else NONE); |
|
39147 | 842 |
|
843 |
(* print modules *) |
|
47576 | 844 |
fun print_module _ base _ xs = |
39147 | 845 |
let |
846 |
val (raw_decls, body) = split_list xs; |
|
55677 | 847 |
val decls = maps these raw_decls |
39147 | 848 |
in (NONE, print_ml_module base decls body) end; |
849 |
||
850 |
(* serialization *) |
|
851 |
val p = Pretty.chunks2 (map snd includes |
|
852 |
@ map snd (Code_Namespace.print_hierarchical { |
|
853 |
print_module = print_module, print_stmt = print_stmt, |
|
854 |
lift_markup = apsnd } ml_program)); |
|
28054 | 855 |
in |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69207
diff
changeset
|
856 |
(Code_Target.Singleton (ml_extension, p), try (deresolver [])) |
28054 | 857 |
end; |
858 |
||
38966 | 859 |
val serializer_sml : Code_Target.serializer = |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69207
diff
changeset
|
860 |
Code_Target.parse_args (Scan.succeed ()) #> K (serialize_ml print_sml_module print_sml_stmt "ML"); |
28054 | 861 |
|
38966 | 862 |
val serializer_ocaml : Code_Target.serializer = |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
69207
diff
changeset
|
863 |
Code_Target.parse_args (Scan.succeed ()) #> K (serialize_ml print_ocaml_module print_ocaml_stmt "ocaml"); |
28054 | 864 |
|
38966 | 865 |
|
866 |
(** Isar setup **) |
|
867 |
||
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
868 |
fun fun_syntax print_typ fxy [ty1, ty2] = |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
869 |
brackify_infix (1, R) fxy ( |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
870 |
print_typ (INFX (1, X)) ty1, |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
871 |
str "->", |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
872 |
print_typ (INFX (1, R)) ty2 |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
873 |
); |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
874 |
|
59323 | 875 |
val _ = Theory.setup |
876 |
(Code_Target.add_language |
|
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66326
diff
changeset
|
877 |
(target_SML, {serializer = serializer_sml, literals = literals_sml, |
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66326
diff
changeset
|
878 |
check = {env_var = "", |
72511
460d743010bc
clarified signature: overloaded "+" for Path.append;
wenzelm
parents:
71803
diff
changeset
|
879 |
make_destination = fn p => p + Path.explode "ROOT.ML", |
50022
286dfcab9833
restored SML code check which got unintentionally broken: must explicitly check for error during compilation;
haftmann
parents:
48568
diff
changeset
|
880 |
make_command = fn _ => |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66326
diff
changeset
|
881 |
"isabelle process -e 'datatype ref = datatype Unsynchronized.ref' -f 'ROOT.ML' -l Pure"}, |
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66326
diff
changeset
|
882 |
evaluation_args = []}) |
59104 | 883 |
#> Code_Target.add_language |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66326
diff
changeset
|
884 |
(target_OCaml, {serializer = serializer_ocaml, literals = literals_ocaml, |
70352 | 885 |
check = {env_var = "ISABELLE_OCAMLFIND", |
72511
460d743010bc
clarified signature: overloaded "+" for Path.append;
wenzelm
parents:
71803
diff
changeset
|
886 |
make_destination = fn p => p + Path.explode "ROOT.ml" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69743
diff
changeset
|
887 |
(*extension demanded by OCaml compiler*), |
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69743
diff
changeset
|
888 |
make_command = fn _ => |
69950 | 889 |
"\"$ISABELLE_OCAMLFIND\" ocamlopt -w pu -package zarith -linkpkg ROOT.ml </dev/null"}, |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66326
diff
changeset
|
890 |
evaluation_args = []}) |
55150 | 891 |
#> Code_Target.set_printings (Type_Constructor ("fun", |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52138
diff
changeset
|
892 |
[(target_SML, SOME (2, fun_syntax)), (target_OCaml, SOME (2, fun_syntax))])) |
28054 | 893 |
#> fold (Code_Target.add_reserved target_SML) ML_Syntax.reserved_names |
75456
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
894 |
#> fold (Code_Target.add_reserved target_SML) [ |
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
895 |
"ref", (*rebinding is illegal*) |
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
896 |
"o", (*dictionary projections use it already*) |
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
897 |
"nil", (*predefined constructor*) |
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
898 |
"Fail", |
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
899 |
"div", "mod", (*standard infixes*) |
160c9c18a707
»nil« seems to be a reserved constructor word in PolyML
haftmann
parents:
75356
diff
changeset
|
900 |
"IntInf"] |
28054 | 901 |
#> fold (Code_Target.add_reserved target_OCaml) [ |
902 |
"and", "as", "assert", "begin", "class", |
|
903 |
"constraint", "do", "done", "downto", "else", "end", "exception", |
|
904 |
"external", "false", "for", "fun", "function", "functor", "if", |
|
905 |
"in", "include", "inherit", "initializer", "lazy", "let", "match", "method", |
|
906 |
"module", "mutable", "new", "object", "of", "open", "or", "private", "rec", |
|
907 |
"sig", "struct", "then", "to", "true", "try", "type", "val", |
|
908 |
"virtual", "when", "while", "with" |
|
909 |
] |
|
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69743
diff
changeset
|
910 |
#> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod", "Z"]); |
28054 | 911 |
|
912 |
end; (*struct*) |