author | wenzelm |
Fri, 01 Jun 2018 11:10:22 +0200 | |
changeset 68345 | 5bc1e1ac7955 |
parent 68034 | 27ba50c79328 |
child 69623 | ef02c5e051e5 |
permissions | -rw-r--r-- |
37745
6315b6426200
checking generated code for various target languages
haftmann
parents:
37669
diff
changeset
|
1 |
(* Title: Tools/Code/code_scala.ML |
6315b6426200
checking generated code for various target languages
haftmann
parents:
37669
diff
changeset
|
2 |
Author: Florian Haftmann, TU Muenchen |
34294 | 3 |
|
4 |
Serializer for Scala. |
|
5 |
*) |
|
6 |
||
7 |
signature CODE_SCALA = |
|
8 |
sig |
|
37745
6315b6426200
checking generated code for various target languages
haftmann
parents:
37669
diff
changeset
|
9 |
val target: string |
34294 | 10 |
end; |
11 |
||
12 |
structure Code_Scala : CODE_SCALA = |
|
13 |
struct |
|
14 |
||
55150 | 15 |
open Basic_Code_Symbol; |
34294 | 16 |
open Basic_Code_Thingol; |
17 |
open Code_Printer; |
|
18 |
||
19 |
infixr 5 @@; |
|
20 |
infixr 5 @|; |
|
21 |
||
22 |
(** Scala serializer **) |
|
23 |
||
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
24 |
val target = "Scala"; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
25 |
|
68028 | 26 |
val print_scala_string = |
27 |
let |
|
28 |
fun chr i = "\\u" ^ align_right "0" 4 (Int.fmt StringCvt.HEX i) |
|
29 |
fun char c = if c = "'" then "\\'" |
|
30 |
else if c = "\"" then "\\\"" |
|
31 |
else if c = "\\" then "\\\\" |
|
32 |
else |
|
33 |
let |
|
34 |
val i = ord c |
|
35 |
in |
|
36 |
if i < 32 orelse i > 126 |
|
37 |
then chr i |
|
38 |
else if i >= 128 |
|
68034 | 39 |
then error "non-ASCII byte in Scala string literal" |
68028 | 40 |
else c |
41 |
end |
|
42 |
in quote o translate_string char end; |
|
43 |
||
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
44 |
datatype scala_stmt = Fun of typscheme * ((iterm list * iterm) * (thm option * bool)) list |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
45 |
| Datatype of vname list * ((string * vname list) * itype list) list |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
46 |
| Class of (vname * ((class * class) list * (string * itype) list)) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
47 |
* (string * { vs: (vname * sort) list, |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
48 |
inst_params: ((string * (const * int)) * (thm * bool)) list, |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
49 |
superinst_params: ((string * (const * int)) * (thm * bool)) list }) list; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
50 |
|
47609 | 51 |
fun print_scala_stmt tyco_syntax const_syntax reserved |
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
52 |
args_num is_constr (deresolve, deresolve_full) = |
34294 | 53 |
let |
55150 | 54 |
val deresolve_const = deresolve o Constant; |
55 |
val deresolve_tyco = deresolve o Type_Constructor; |
|
56 |
val deresolve_class = deresolve o Type_Class; |
|
56812 | 57 |
fun lookup_tyvar tyvars = lookup_var tyvars o Name.enforce_case true; |
58 |
fun intro_tyvars vs = intro_vars (map (Name.enforce_case true o fst) vs); |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
59 |
fun print_tyco_expr tyvars fxy (sym, tys) = applify "[" "]" |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
60 |
(print_typ tyvars NOBR) fxy ((str o deresolve) sym) tys |
38923 | 61 |
and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco |
55150 | 62 |
of NONE => print_tyco_expr tyvars fxy (Type_Constructor tyco, tys) |
47609 | 63 |
| SOME (_, print) => print (print_typ tyvars) fxy tys) |
37243
6e2ac5358d6e
capitalized type variables; added yield as keyword
haftmann
parents:
37224
diff
changeset
|
64 |
| print_typ tyvars fxy (ITyVar v) = (str o lookup_tyvar tyvars) v; |
55150 | 65 |
fun print_dicttyp tyvars (class, ty) = print_tyco_expr tyvars NOBR (Type_Class class, [ty]); |
37639 | 66 |
fun print_tupled_typ tyvars ([], ty) = |
67 |
print_typ tyvars NOBR ty |
|
68 |
| print_tupled_typ tyvars ([ty1], ty2) = |
|
69 |
concat [print_typ tyvars BR ty1, str "=>", print_typ tyvars NOBR ty2] |
|
70 |
| print_tupled_typ tyvars (tys, ty) = |
|
71 |
concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys), |
|
72 |
str "=>", print_typ tyvars NOBR ty]; |
|
73 |
fun constraint p1 p2 = Pretty.block [p1, str ":", Pretty.brk 1, p2]; |
|
34294 | 74 |
fun print_var vars NONE = str "_" |
63303 | 75 |
| print_var vars (SOME v) = (str o lookup_var vars) v; |
76 |
fun applify_dict tyvars (Dict (_, d)) = applify_plain_dict tyvars d |
|
77 |
and applify_plain_dict tyvars (Dict_Const (inst, dss)) = |
|
78 |
applify_dictss tyvars ((str o deresolve o Class_Instance) inst) dss |
|
79 |
| applify_plain_dict tyvars (Dict_Var { var, class, ... }) = |
|
80 |
Pretty.block [str "implicitly", |
|
81 |
enclose "[" "]" [Pretty.block [(str o deresolve_class) class, |
|
82 |
enclose "[" "]" [(str o lookup_tyvar tyvars) var]]]] |
|
83 |
and applify_dictss tyvars p dss = |
|
84 |
applify "(" ")" (applify_dict tyvars) NOBR p (flat dss) |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
85 |
fun print_term tyvars is_pat some_thm vars fxy (IConst const) = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
86 |
print_app tyvars is_pat some_thm vars fxy (const, []) |
35228 | 87 |
| print_term tyvars is_pat some_thm vars fxy (t as (t1 `$ t2)) = |
34294 | 88 |
(case Code_Thingol.unfold_const_app t |
35228 | 89 |
of SOME app => print_app tyvars is_pat some_thm vars fxy app |
37639 | 90 |
| _ => applify "(" ")" (print_term tyvars is_pat some_thm vars NOBR) fxy |
91 |
(print_term tyvars is_pat some_thm vars BR t1) [t2]) |
|
35228 | 92 |
| print_term tyvars is_pat some_thm vars fxy (IVar v) = |
34294 | 93 |
print_var vars v |
61130
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
94 |
| print_term tyvars is_pat some_thm vars fxy (t as _ `|=> _) = |
34294 | 95 |
let |
61130
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
96 |
val (vs_tys, body) = Code_Thingol.unfold_abs t; |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
97 |
val (ps, vars') = fold_map (print_abs_head tyvars) vs_tys vars; |
34294 | 98 |
in |
61130
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
99 |
brackets (ps @| print_term tyvars false some_thm vars' NOBR body) |
41939 | 100 |
end |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
101 |
| print_term tyvars is_pat some_thm vars fxy (ICase case_expr) = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
102 |
(case Code_Thingol.unfold_const_app (#primitive case_expr) |
55150 | 103 |
of SOME (app as ({ sym = Constant const, ... }, _)) => |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
104 |
if is_none (const_syntax const) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
105 |
then print_case tyvars some_thm vars fxy case_expr |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
106 |
else print_app tyvars is_pat some_thm vars fxy app |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
107 |
| NONE => print_case tyvars some_thm vars fxy case_expr) |
61130
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
108 |
and print_abs_head tyvars (some_v, ty) vars = |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
109 |
let |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
110 |
val vars' = intro_vars (the_list some_v) vars; |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
111 |
in |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
112 |
(concat [ |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
113 |
enclose "(" ")" [constraint (print_var vars' some_v) (print_typ tyvars NOBR ty)], |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
114 |
str "=>" |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
115 |
], vars') |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
59323
diff
changeset
|
116 |
end |
37464 | 117 |
and print_app tyvars is_pat some_thm vars fxy |
63303 | 118 |
(app as ({ sym, typargs, dom, dicts, ... }, ts)) = |
34294 | 119 |
let |
120 |
val k = length ts; |
|
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
121 |
val typargs' = if is_pat then [] else typargs; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
122 |
val syntax = case sym of |
55150 | 123 |
Constant const => const_syntax const |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
124 |
| _ => NONE; |
63303 | 125 |
val applify_dicts = |
126 |
if is_pat orelse is_some syntax orelse is_constr sym |
|
127 |
orelse Code_Thingol.unambiguous_dictss dicts |
|
128 |
then fn p => K p |
|
129 |
else applify_dictss tyvars; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
130 |
val (l, print') = case syntax of |
63303 | 131 |
NONE => (args_num sym, fn fxy => fn ts => applify_dicts |
132 |
(gen_applify (is_constr sym) "(" ")" |
|
37639 | 133 |
(print_term tyvars is_pat some_thm vars NOBR) fxy |
134 |
(applify "[" "]" (print_typ tyvars NOBR) |
|
63303 | 135 |
NOBR ((str o deresolve) sym) typargs') ts) dicts) |
136 |
| SOME (k, Plain_printer s) => (k, fn fxy => fn ts => applify_dicts |
|
137 |
(applify "(" ")" |
|
37881 | 138 |
(print_term tyvars is_pat some_thm vars NOBR) fxy |
139 |
(applify "[" "]" (print_typ tyvars NOBR) |
|
63303 | 140 |
NOBR (str s) typargs') ts) dicts) |
55153
eedd549de3ef
more suitable names, without any notion of "activating"
haftmann
parents:
55150
diff
changeset
|
141 |
| SOME (k, Complex_printer print) => |
38059 | 142 |
(k, fn fxy => fn ts => print (print_term tyvars is_pat some_thm) some_thm vars fxy |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
143 |
(ts ~~ take k dom)) |
38059 | 144 |
in if k = l then print' fxy ts |
34294 | 145 |
else if k < l then |
35228 | 146 |
print_term tyvars is_pat some_thm vars fxy (Code_Thingol.eta_expand l app) |
34294 | 147 |
else let |
148 |
val (ts1, ts23) = chop l ts; |
|
149 |
in |
|
38059 | 150 |
Pretty.block (print' BR ts1 :: map (fn t => Pretty.block |
35228 | 151 |
[str ".apply(", print_term tyvars is_pat some_thm vars NOBR t, str ")"]) ts23) |
34294 | 152 |
end end |
37464 | 153 |
and print_bind tyvars some_thm fxy p = |
154 |
gen_print_bind (print_term tyvars true) some_thm fxy p |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
155 |
and print_case tyvars some_thm vars fxy { clauses = [], ... } = |
48073
1b609a7837ef
prefer sys.error over plain error in Scala to avoid deprecation warning
haftmann
parents:
48072
diff
changeset
|
156 |
(brackify fxy o Pretty.breaks o map str) ["sys.error(\"empty case\")"] |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
157 |
| print_case tyvars some_thm vars fxy (case_expr as { clauses = [_], ... }) = |
34294 | 158 |
let |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
159 |
val (bind :: binds, body) = Code_Thingol.unfold_let (ICase case_expr); |
41781
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
160 |
fun print_match_val ((pat, ty), t) vars = |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
161 |
vars |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
162 |
|> print_bind tyvars some_thm BR pat |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
163 |
|>> (fn p => (false, concat [str "val", constraint p (print_typ tyvars NOBR ty), |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
164 |
str "=", print_term tyvars false some_thm vars NOBR t])); |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
165 |
fun print_match_seq t vars = |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
166 |
((true, print_term tyvars false some_thm vars NOBR t), vars); |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
167 |
fun print_match is_first ((IVar NONE, ty), t) = |
41784 | 168 |
if Code_Thingol.is_IAbs t andalso is_first |
41781
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
169 |
then print_match_val ((IVar NONE, ty), t) |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
170 |
else print_match_seq t |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
171 |
| print_match _ ((pat, ty), t) = |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
172 |
print_match_val ((pat, ty), t); |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
173 |
val (seps_ps, vars') = |
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
174 |
vars |> print_match true bind ||>> fold_map (print_match false) binds |>> uncurry cons; |
38059 | 175 |
val all_seps_ps = seps_ps @ [(true, print_term tyvars false some_thm vars' NOBR body)]; |
176 |
fun insert_seps [(_, p)] = [p] |
|
177 |
| insert_seps ((_, p) :: (seps_ps as (sep, _) :: _)) = |
|
178 |
(if sep then Pretty.block [p, str ";"] else p) :: insert_seps seps_ps |
|
41781
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
haftmann
parents:
41687
diff
changeset
|
179 |
in brackify_block fxy (str "{") (insert_seps all_seps_ps) (str "}") end |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
180 |
| print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } = |
34294 | 181 |
let |
182 |
fun print_select (pat, body) = |
|
183 |
let |
|
37464 | 184 |
val (p_pat, vars') = print_bind tyvars some_thm NOBR pat vars; |
185 |
val p_body = print_term tyvars false some_thm vars' NOBR body |
|
186 |
in concat [str "case", p_pat, str "=>", p_body] end; |
|
46850 | 187 |
in |
188 |
map print_select clauses |
|
189 |
|> Pretty.block_enclose (concat [print_term tyvars false some_thm vars NOBR t, str "match", str "{"], str "}") |
|
190 |
|> single |
|
191 |
|> enclose "(" ")" |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
48003
diff
changeset
|
192 |
end; |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
193 |
fun print_context tyvars vs s = applify "[" "]" |
37639 | 194 |
(fn (v, sort) => (Pretty.block o map str) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
195 |
(lookup_tyvar tyvars v :: maps (fn class => [" : ", deresolve_class class]) sort)) |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
196 |
NOBR (str s) vs; |
66326 | 197 |
fun print_defhead tyvars vars const vs params tys ty = |
55776
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents:
55684
diff
changeset
|
198 |
concat [str "def", constraint (applify "(" ")" (fn (param, ty) => |
37639 | 199 |
constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty)) |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
200 |
NOBR (print_context tyvars vs (deresolve_const const)) (params ~~ tys)) (print_typ tyvars NOBR ty), |
55679 | 201 |
str "="]; |
66326 | 202 |
fun print_def const (vs, ty) [] = |
37639 | 203 |
let |
204 |
val (tys, ty') = Code_Thingol.unfold_fun ty; |
|
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43326
diff
changeset
|
205 |
val params = Name.invent (snd reserved) "a" (length tys); |
37639 | 206 |
val tyvars = intro_tyvars vs reserved; |
207 |
val vars = intro_vars params reserved; |
|
208 |
in |
|
66326 | 209 |
concat [print_defhead tyvars vars const vs params tys ty', |
68028 | 210 |
str ("sys.error(" ^ print_scala_string const ^ ")")] |
37639 | 211 |
end |
66326 | 212 |
| print_def const (vs, ty) eqs = |
37639 | 213 |
let |
214 |
val tycos = fold (fn ((ts, t), _) => |
|
215 |
fold Code_Thingol.add_tyconames (t :: ts)) eqs []; |
|
216 |
val tyvars = reserved |
|
217 |
|> intro_base_names |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
218 |
(is_none o tyco_syntax) deresolve_tyco tycos |
37639 | 219 |
|> intro_tyvars vs; |
220 |
val simple = case eqs |
|
221 |
of [((ts, _), _)] => forall Code_Thingol.is_IVar ts |
|
222 |
| _ => false; |
|
223 |
val vars1 = reserved |
|
55145
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52520
diff
changeset
|
224 |
|> intro_base_names_for (is_none o const_syntax) |
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
haftmann
parents:
52520
diff
changeset
|
225 |
deresolve (map (snd o fst) eqs); |
37639 | 226 |
val params = if simple |
227 |
then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs |
|
228 |
else aux_params vars1 (map (fst o fst) eqs); |
|
229 |
val vars2 = intro_vars params vars1; |
|
230 |
val (tys', ty') = Code_Thingol.unfold_fun_n (length params) ty; |
|
38922 | 231 |
fun tuplify [p] = p |
232 |
| tuplify ps = enum "," "(" ")" ps; |
|
37639 | 233 |
fun print_rhs vars' ((_, t), (some_thm, _)) = |
234 |
print_term tyvars false some_thm vars' NOBR t; |
|
235 |
fun print_clause (eq as ((ts, _), (some_thm, _))) = |
|
34294 | 236 |
let |
37639 | 237 |
val vars' = intro_vars ((fold o Code_Thingol.fold_varnames) |
238 |
(insert (op =)) ts []) vars1; |
|
239 |
in |
|
240 |
concat [str "case", |
|
38922 | 241 |
tuplify (map (print_term tyvars true some_thm vars' NOBR) ts), |
37639 | 242 |
str "=>", print_rhs vars' eq] |
243 |
end; |
|
66326 | 244 |
val head = print_defhead tyvars vars2 const vs params tys' ty'; |
37639 | 245 |
in if simple then |
246 |
concat [head, print_rhs vars2 (hd eqs)] |
|
247 |
else |
|
248 |
Pretty.block_enclose |
|
38922 | 249 |
(concat [head, tuplify (map (str o lookup_var vars2) params), |
37639 | 250 |
str "match", str "{"], str "}") |
251 |
(map print_clause eqs) |
|
252 |
end; |
|
55150 | 253 |
val print_method = str o Library.enclose "`" "`" o deresolve_full o Constant; |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
254 |
fun print_inst class (tyco, { vs, inst_params, superinst_params }) = |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
255 |
let |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
256 |
val tyvars = intro_tyvars vs reserved; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
257 |
val classtyp = (class, tyco `%% map (ITyVar o fst) vs); |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
258 |
fun print_classparam_instance ((classparam, (const as { dom, ... }, dom_length)), (thm, _)) = |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
259 |
let |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
260 |
val aux_dom = Name.invent_names (snd reserved) "a" dom; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
261 |
val auxs = map fst aux_dom; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
262 |
val vars = intro_vars auxs reserved; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
263 |
val (aux_dom1, aux_dom2) = chop dom_length aux_dom; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
264 |
fun abstract_using [] = [] |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
265 |
| abstract_using aux_dom = [enum "," "(" ")" |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
266 |
(map (fn (aux, ty) => constraint ((str o lookup_var vars) aux) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
267 |
(print_typ tyvars NOBR ty)) aux_dom), str "=>"]; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
268 |
val aux_abstr1 = abstract_using aux_dom1; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
269 |
val aux_abstr2 = abstract_using aux_dom2; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
270 |
in |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
271 |
concat ([str "val", print_method classparam, str "="] |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
272 |
@ aux_abstr1 @ aux_abstr2 @| print_app tyvars false (SOME thm) vars NOBR |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
273 |
(const, map (IVar o SOME) auxs)) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
274 |
end; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
275 |
in |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
276 |
Pretty.block_enclose (concat [str "implicit def", |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
277 |
constraint (print_context tyvars vs |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
278 |
((Library.enclose "`" "`" o deresolve_full o Class_Instance) (tyco, class))) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
279 |
(print_dicttyp tyvars classtyp), |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
280 |
str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}") |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
281 |
(map print_classparam_instance (inst_params @ superinst_params)) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
282 |
end; |
66326 | 283 |
fun print_stmt (Constant const, (_, Fun ((vs, ty), raw_eqs))) = |
284 |
print_def const (vs, ty) (filter (snd o snd) raw_eqs) |
|
285 |
| print_stmt (Type_Constructor tyco, (_, Datatype (vs, cos))) = |
|
34294 | 286 |
let |
48003
1d11af40b106
dropped sort constraints on datatype specifications
haftmann
parents:
47609
diff
changeset
|
287 |
val tyvars = intro_tyvars (map (rpair []) vs) reserved; |
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
288 |
fun print_co ((co, vs_args), tys) = |
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
289 |
concat [Pretty.block ((applify "[" "]" (str o lookup_tyvar tyvars) NOBR |
55776
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents:
55684
diff
changeset
|
290 |
((concat o map str) ["final", "case", "class", deresolve_const co]) vs_args) |
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
291 |
@@ enum "," "(" ")" (map (fn (v, arg) => constraint (str v) (print_typ tyvars NOBR arg)) |
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
292 |
(Name.invent_names (snd reserved) "a" tys))), |
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
293 |
str "extends", |
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
294 |
applify "[" "]" (str o lookup_tyvar tyvars) NOBR |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
295 |
((str o deresolve_tyco) tyco) vs |
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
296 |
]; |
34294 | 297 |
in |
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
298 |
Pretty.chunks (applify "[" "]" (str o lookup_tyvar tyvars) |
55776
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents:
55684
diff
changeset
|
299 |
NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve_tyco tyco]) vs |
37639 | 300 |
:: map print_co cos) |
34294 | 301 |
end |
66326 | 302 |
| print_stmt (Type_Class class, (_, Class ((v, (classrels, classparams)), insts))) = |
34294 | 303 |
let |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
304 |
val tyvars = intro_tyvars [(v, [class])] reserved; |
37639 | 305 |
fun add_typarg s = Pretty.block |
306 |
[str s, str "[", (str o lookup_tyvar tyvars) v, str "]"]; |
|
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37337
diff
changeset
|
307 |
fun print_super_classes [] = NONE |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
308 |
| print_super_classes classrels = SOME (concat (str "extends" |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
309 |
:: separate (str "with") (map (add_typarg o deresolve_class o snd) classrels))); |
34294 | 310 |
fun print_classparam_val (classparam, ty) = |
37639 | 311 |
concat [str "val", constraint (print_method classparam) |
312 |
((print_tupled_typ tyvars o Code_Thingol.unfold_fun) ty)]; |
|
34294 | 313 |
fun print_classparam_def (classparam, ty) = |
314 |
let |
|
315 |
val (tys, ty) = Code_Thingol.unfold_fun ty; |
|
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43326
diff
changeset
|
316 |
val [implicit_name] = Name.invent (snd reserved) (lookup_tyvar tyvars v) 1; |
37639 | 317 |
val proto_vars = intro_vars [implicit_name] reserved; |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43326
diff
changeset
|
318 |
val auxs = Name.invent (snd proto_vars) "a" (length tys); |
37639 | 319 |
val vars = intro_vars auxs proto_vars; |
34294 | 320 |
in |
55776
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents:
55684
diff
changeset
|
321 |
concat [str "def", constraint (Pretty.block [applify "(" ")" |
37639 | 322 |
(fn (aux, ty) => constraint ((str o lookup_var vars) aux) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
323 |
(print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve_const classparam)) |
37639 | 324 |
(auxs ~~ tys), str "(implicit ", str implicit_name, str ": ", |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
325 |
add_typarg (deresolve_class class), str ")"]) (print_typ tyvars NOBR ty), str "=", |
37639 | 326 |
applify "(" ")" (str o lookup_var vars) NOBR |
327 |
(Pretty.block [str implicit_name, str ".", print_method classparam]) auxs] |
|
34294 | 328 |
end; |
329 |
in |
|
330 |
Pretty.chunks ( |
|
331 |
(Pretty.block_enclose |
|
55776
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents:
55684
diff
changeset
|
332 |
(concat ([str "trait", (add_typarg o deresolve_class) class] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
333 |
@ the_list (print_super_classes classrels) @ [str "{"]), str "}") |
34294 | 334 |
(map print_classparam_val classparams)) |
335 |
:: map print_classparam_def classparams |
|
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
336 |
@| Pretty.block_enclose |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
337 |
((concat o map str) ["object", deresolve_class class, "{"], str "}") |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
338 |
(map (print_inst class) insts) |
34294 | 339 |
) |
37639 | 340 |
end; |
34294 | 341 |
in print_stmt end; |
342 |
||
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
343 |
fun pickup_instances_for_class program = |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
344 |
let |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
345 |
val tab = |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
346 |
Symtab.empty |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
347 |
|> Code_Symbol.Graph.fold |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
348 |
(fn (_, (Code_Thingol.Classinst { class, tyco, vs, inst_params, superinst_params, ... }, _)) => |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
349 |
Symtab.map_default (class, []) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
350 |
(cons (tyco, { vs = vs, inst_params = inst_params, superinst_params = superinst_params })) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
351 |
| _ => I) program; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
352 |
in Symtab.lookup_list tab end; |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
353 |
|
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
354 |
fun scala_program_of_program ctxt case_insensitive module_name reserved identifiers exports program = |
34294 | 355 |
let |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
356 |
val variant = if case_insensitive |
58520
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
haftmann
parents:
58398
diff
changeset
|
357 |
then Code_Namespace.variant_case_insensitive |
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
haftmann
parents:
58398
diff
changeset
|
358 |
else Name.variant; |
38809
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
359 |
fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) = |
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
360 |
let |
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
361 |
val declare = Name.declare name_fragment; |
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
362 |
in (name_fragment, ((declare nsp_class, declare nsp_object), declare nsp_common)) end; |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
363 |
fun namify_class base ((nsp_class, nsp_object), nsp_common) = |
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
364 |
let |
58520
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
haftmann
parents:
58398
diff
changeset
|
365 |
val (base', nsp_class') = variant base nsp_class |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
366 |
in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end; |
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
367 |
fun namify_object base ((nsp_class, nsp_object), nsp_common) = |
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
368 |
let |
58520
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
haftmann
parents:
58398
diff
changeset
|
369 |
val (base', nsp_object') = variant base nsp_object |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
370 |
in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end; |
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
|
371 |
fun namify_common base ((nsp_class, nsp_object), nsp_common) = |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
372 |
let |
58520
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
haftmann
parents:
58398
diff
changeset
|
373 |
val (base', nsp_common') = variant base nsp_common |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
374 |
in |
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
|
375 |
(base', ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common')) |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
376 |
end; |
38809
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
377 |
fun namify_stmt (Code_Thingol.Fun _) = namify_object |
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
378 |
| namify_stmt (Code_Thingol.Datatype _) = namify_class |
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
|
379 |
| namify_stmt (Code_Thingol.Datatypecons _) = namify_common |
38809
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
380 |
| namify_stmt (Code_Thingol.Class _) = namify_class |
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
381 |
| namify_stmt (Code_Thingol.Classrel _) = namify_object |
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
382 |
| namify_stmt (Code_Thingol.Classparam _) = namify_object |
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
|
383 |
| namify_stmt (Code_Thingol.Classinst _) = namify_common; |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
384 |
val pickup_instances = pickup_instances_for_class program; |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
385 |
fun modify_stmt (_, (_, Code_Thingol.Fun (_, SOME _))) = NONE |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
386 |
| modify_stmt (_, (export, Code_Thingol.Fun (x, NONE))) = SOME (export, Fun x) |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
387 |
| modify_stmt (_, (export, Code_Thingol.Datatype x)) = SOME (export, Datatype x) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
388 |
| modify_stmt (_, (_, Code_Thingol.Datatypecons _)) = NONE |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
389 |
| modify_stmt (Type_Class class, (export, Code_Thingol.Class x)) = |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
390 |
SOME (export, Class (x, pickup_instances class)) |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
391 |
| modify_stmt (_, (_, Code_Thingol.Classrel _)) = NONE |
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
392 |
| modify_stmt (_, (_, Code_Thingol.Classparam _)) = NONE |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
393 |
| modify_stmt (_, (_, Code_Thingol.Classinst _)) = NONE |
38970
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
38966
diff
changeset
|
394 |
in |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
395 |
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
|
396 |
{ module_name = module_name, reserved = reserved, identifiers = identifiers, |
41939 | 397 |
empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module, |
55684
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents:
55683
diff
changeset
|
398 |
namify_stmt = namify_stmt, cyclic_modules = true, |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
399 |
class_transitive = true, class_relation_public = false, empty_data = (), |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
400 |
memorize_data = K I, modify_stmts = map modify_stmt } |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
401 |
exports program |
38970
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
38966
diff
changeset
|
402 |
end; |
34294 | 403 |
|
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
404 |
fun serialize_scala case_insensitive ctxt { module_name, reserved_syms, identifiers, |
55683 | 405 |
includes, class_syntax, tyco_syntax, const_syntax } exports program = |
34294 | 406 |
let |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
407 |
|
38809
7dc73a208722
proper namespace administration for hierarchical modules
haftmann
parents:
38782
diff
changeset
|
408 |
(* build program *) |
39147 | 409 |
val { deresolver, hierarchical_program = scala_program } = |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
410 |
scala_program_of_program ctxt case_insensitive module_name (Name.make_context reserved_syms) |
55683 | 411 |
identifiers exports program; |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
412 |
|
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
413 |
(* print statements *) |
55150 | 414 |
fun lookup_constr tyco constr = case Code_Symbol.Graph.get_node program (Type_Constructor tyco) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
415 |
of Code_Thingol.Datatype (_, constrs) => |
37639 | 416 |
the (AList.lookup (op = o apsnd fst) constrs constr); |
55150 | 417 |
fun classparams_of_class class = case Code_Symbol.Graph.get_node program (Type_Class class) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
418 |
of Code_Thingol.Class (_, (_, classparams)) => classparams; |
55150 | 419 |
fun args_num (sym as Constant const) = case Code_Symbol.Graph.get_node program sym |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
420 |
of Code_Thingol.Fun (((_, ty), []), _) => |
37464 | 421 |
(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
|
422 |
| Code_Thingol.Fun ((_, ((ts, _), _) :: _), _) => length ts |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
423 |
| Code_Thingol.Datatypecons tyco => length (lookup_constr tyco const) |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
424 |
| Code_Thingol.Classparam class => |
37639 | 425 |
(length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
426 |
(classparams_of_class class)) const; |
47609 | 427 |
fun print_stmt prefix_fragments = print_scala_stmt |
39147 | 428 |
tyco_syntax const_syntax (make_vars reserved_syms) args_num |
50626
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents:
48568
diff
changeset
|
429 |
(Code_Thingol.is_constr program) (deresolver prefix_fragments, deresolver []); |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
430 |
|
39147 | 431 |
(* print modules *) |
63350
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
432 |
fun print_module _ base _ ps = Pretty.chunks2 |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
433 |
(str ("object " ^ base ^ " {") |
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
haftmann
parents:
63304
diff
changeset
|
434 |
:: ps @| str ("} /* object " ^ base ^ " */")); |
38769
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
435 |
|
317e64c886d2
preliminary implementation of hierarchical module name space
haftmann
parents:
38059
diff
changeset
|
436 |
(* serialization *) |
39147 | 437 |
val p = Pretty.chunks2 (map snd includes |
438 |
@ Code_Namespace.print_hierarchical { |
|
439 |
print_module = print_module, print_stmt = print_stmt, |
|
440 |
lift_markup = I } scala_program); |
|
39056 | 441 |
fun write width NONE = writeln o format [] width |
442 |
| write width (SOME p) = File.write p o format [] width; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55145
diff
changeset
|
443 |
fun prepare syms width p = ([("", format syms width p)], try (deresolver [])); |
34294 | 444 |
in |
48568 | 445 |
Code_Target.serialization write prepare p |
34294 | 446 |
end; |
447 |
||
38966 | 448 |
val serializer : Code_Target.serializer = |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
449 |
Code_Target.parse_args (Scan.optional (Args.$$$ "case_insensitive" >> K true) false |
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
450 |
>> (fn case_insensitive => serialize_scala case_insensitive)); |
38966 | 451 |
|
34294 | 452 |
val literals = let |
58398 | 453 |
fun numeral_scala k = |
454 |
if ~2147483647 < k andalso k <= 2147483647 |
|
455 |
then signed_string_of_int k |
|
456 |
else quote (signed_string_of_int k) |
|
34294 | 457 |
in Literals { |
68028 | 458 |
literal_string = print_scala_string, |
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34900
diff
changeset
|
459 |
literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")", |
34888 | 460 |
literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps], |
34294 | 461 |
infix_cons = (6, "::") |
462 |
} end; |
|
463 |
||
464 |
||
465 |
(** Isar setup **) |
|
466 |
||
59323 | 467 |
val _ = Theory.setup |
468 |
(Code_Target.add_language |
|
38966 | 469 |
(target, { serializer = serializer, literals = literals, |
41939 | 470 |
check = { env_var = "SCALA_HOME", |
471 |
make_destination = fn p => Path.append p (Path.explode "ROOT.scala"), |
|
41940 | 472 |
make_command = fn _ => |
67207
ad538f6c5d2f
dedicated case option for code generation to Scala
haftmann
parents:
66327
diff
changeset
|
473 |
"isabelle_scala scalac $ISABELLE_SCALAC_OPTIONS ROOT.scala"}, |
67496 | 474 |
evaluation_args = Token.explode0 Keyword.empty_keywords "case_insensitive"}) |
55150 | 475 |
#> 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:
52229
diff
changeset
|
476 |
[(target, SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] => |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52229
diff
changeset
|
477 |
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:
52229
diff
changeset
|
478 |
print_typ BR ty1 (*product type vs. tupled arguments!*), |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52229
diff
changeset
|
479 |
str "=>", |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52229
diff
changeset
|
480 |
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:
52229
diff
changeset
|
481 |
)))])) |
34294 | 482 |
#> fold (Code_Target.add_reserved target) [ |
483 |
"abstract", "case", "catch", "class", "def", "do", "else", "extends", "false", |
|
484 |
"final", "finally", "for", "forSome", "if", "implicit", "import", "lazy", |
|
485 |
"match", "new", "null", "object", "override", "package", "private", "protected", |
|
486 |
"requires", "return", "sealed", "super", "this", "throw", "trait", "try", |
|
37243
6e2ac5358d6e
capitalized type variables; added yield as keyword
haftmann
parents:
37224
diff
changeset
|
487 |
"true", "type", "val", "var", "while", "with", "yield" |
34294 | 488 |
] |
489 |
#> fold (Code_Target.add_reserved target) [ |
|
48073
1b609a7837ef
prefer sys.error over plain error in Scala to avoid deprecation warning
haftmann
parents:
48072
diff
changeset
|
490 |
"apply", "sys", "scala", "BigInt", "Nil", "List" |
59323 | 491 |
]); |
34294 | 492 |
|
493 |
end; (*struct*) |