| author | haftmann | 
| Wed, 22 May 2013 22:56:17 +0200 | |
| changeset 52119 | 90ba620333d0 | 
| parent 51143 | 0a2371e7ced3 | 
| child 52138 | e21426f244aa | 
| permissions | -rw-r--r-- | 
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37669diff
changeset | 1 | (* Title: Tools/Code/code_scala.ML | 
| 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37669diff
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: 
37669diff
changeset | 9 | val target: string | 
| 34294 | 10 | val setup: theory -> theory | 
| 11 | end; | |
| 12 | ||
| 13 | structure Code_Scala : CODE_SCALA = | |
| 14 | struct | |
| 15 | ||
| 16 | val target = "Scala"; | |
| 17 | ||
| 18 | open Basic_Code_Thingol; | |
| 19 | open Code_Printer; | |
| 20 | ||
| 21 | infixr 5 @@; | |
| 22 | infixr 5 @|; | |
| 23 | ||
| 24 | ||
| 25 | (** Scala serializer **) | |
| 26 | ||
| 47609 | 27 | fun print_scala_stmt tyco_syntax const_syntax reserved | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 28 | args_num is_constr (deresolve, deresolve_full) = | 
| 34294 | 29 | let | 
| 37639 | 30 | fun lookup_tyvar tyvars = lookup_var tyvars o first_upper; | 
| 31 | fun intro_tyvars vs = intro_vars (map (first_upper o fst) vs); | |
| 32 | fun print_tyco_expr tyvars fxy (tyco, tys) = applify "[" "]" | |
| 33 | (print_typ tyvars NOBR) fxy ((str o deresolve) tyco) tys | |
| 38923 | 34 | and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco | 
| 37639 | 35 | of NONE => print_tyco_expr tyvars fxy (tyco, tys) | 
| 47609 | 36 | | SOME (_, print) => print (print_typ tyvars) fxy tys) | 
| 37243 
6e2ac5358d6e
capitalized type variables; added yield as keyword
 haftmann parents: 
37224diff
changeset | 37 | | print_typ tyvars fxy (ITyVar v) = (str o lookup_tyvar tyvars) v; | 
| 37639 | 38 | fun print_dicttyp tyvars (class, ty) = print_tyco_expr tyvars NOBR (class, [ty]); | 
| 39 | fun print_tupled_typ tyvars ([], ty) = | |
| 40 | print_typ tyvars NOBR ty | |
| 41 | | print_tupled_typ tyvars ([ty1], ty2) = | |
| 42 | concat [print_typ tyvars BR ty1, str "=>", print_typ tyvars NOBR ty2] | |
| 43 | | print_tupled_typ tyvars (tys, ty) = | |
| 44 |           concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys),
 | |
| 45 | str "=>", print_typ tyvars NOBR ty]; | |
| 46 | fun constraint p1 p2 = Pretty.block [p1, str ":", Pretty.brk 1, p2]; | |
| 34294 | 47 | fun print_var vars NONE = str "_" | 
| 48 | | print_var vars (SOME v) = (str o lookup_var vars) v | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 49 | fun print_term tyvars is_pat some_thm vars fxy (IConst const) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 50 | print_app tyvars is_pat some_thm vars fxy (const, []) | 
| 35228 | 51 | | print_term tyvars is_pat some_thm vars fxy (t as (t1 `$ t2)) = | 
| 34294 | 52 | (case Code_Thingol.unfold_const_app t | 
| 35228 | 53 | of SOME app => print_app tyvars is_pat some_thm vars fxy app | 
| 37639 | 54 |             | _ => applify "(" ")" (print_term tyvars is_pat some_thm vars NOBR) fxy
 | 
| 55 | (print_term tyvars is_pat some_thm vars BR t1) [t2]) | |
| 35228 | 56 | | print_term tyvars is_pat some_thm vars fxy (IVar v) = | 
| 34294 | 57 | print_var vars v | 
| 35228 | 58 | | print_term tyvars is_pat some_thm vars fxy ((v, ty) `|=> t) = | 
| 34294 | 59 | let | 
| 60 | val vars' = intro_vars (the_list v) vars; | |
| 61 | in | |
| 62 | concat [ | |
| 37639 | 63 |               enclose "(" ")" [constraint (print_var vars' v) (print_typ tyvars NOBR ty)],
 | 
| 34294 | 64 | str "=>", | 
| 35228 | 65 | print_term tyvars false some_thm vars' NOBR t | 
| 34294 | 66 | ] | 
| 41939 | 67 | end | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 68 | | print_term tyvars is_pat some_thm vars fxy (ICase case_expr) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 69 | (case Code_Thingol.unfold_const_app (#primitive case_expr) | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 70 |            of SOME (app as ({ name = c, ... }, _)) => if is_none (const_syntax c)
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 71 | then print_case tyvars some_thm vars fxy case_expr | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 72 | else print_app tyvars is_pat some_thm vars fxy app | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 73 | | NONE => print_case tyvars some_thm vars fxy case_expr) | 
| 37464 | 74 | and print_app tyvars is_pat some_thm vars fxy | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 75 |         (app as ({ name = c, typargs, dom, ... }, ts)) =
 | 
| 34294 | 76 | let | 
| 77 | val k = length ts; | |
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 78 | val typargs' = if is_pat then [] else typargs; | 
| 38923 | 79 | val (l, print') = case const_syntax c | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 80 |          of NONE => (args_num c, fn fxy => fn ts => gen_applify (is_constr c ) "(" ")"
 | 
| 37639 | 81 | (print_term tyvars is_pat some_thm vars NOBR) fxy | 
| 82 | (applify "[" "]" (print_typ tyvars NOBR) | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 83 | NOBR ((str o deresolve) c) typargs') ts) | 
| 38059 | 84 |           | SOME (Plain_const_syntax (k, s)) => (k, fn fxy => fn ts => applify "(" ")"
 | 
| 37881 | 85 | (print_term tyvars is_pat some_thm vars NOBR) fxy | 
| 86 | (applify "[" "]" (print_typ tyvars NOBR) | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 87 | NOBR (str s) typargs') ts) | 
| 37881 | 88 | | SOME (Complex_const_syntax (k, print)) => | 
| 38059 | 89 | (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: 
48003diff
changeset | 90 | (ts ~~ take k dom)) | 
| 38059 | 91 | in if k = l then print' fxy ts | 
| 34294 | 92 | else if k < l then | 
| 35228 | 93 | print_term tyvars is_pat some_thm vars fxy (Code_Thingol.eta_expand l app) | 
| 34294 | 94 | else let | 
| 95 | val (ts1, ts23) = chop l ts; | |
| 96 | in | |
| 38059 | 97 | Pretty.block (print' BR ts1 :: map (fn t => Pretty.block | 
| 35228 | 98 |           [str ".apply(", print_term tyvars is_pat some_thm vars NOBR t, str ")"]) ts23)
 | 
| 34294 | 99 | end end | 
| 37464 | 100 | and print_bind tyvars some_thm fxy p = | 
| 101 | gen_print_bind (print_term tyvars true) some_thm fxy p | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 102 |     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: 
48072diff
changeset | 103 | (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: 
48003diff
changeset | 104 |       | print_case tyvars some_thm vars fxy (case_expr as { clauses = [_], ... }) =
 | 
| 34294 | 105 | let | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 106 | 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: 
41687diff
changeset | 107 | fun print_match_val ((pat, ty), t) vars = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 108 | vars | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 109 | |> print_bind tyvars some_thm BR pat | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 110 | |>> (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: 
41687diff
changeset | 111 | str "=", print_term tyvars false some_thm vars NOBR t])); | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 112 | fun print_match_seq t vars = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 113 | ((true, print_term tyvars false some_thm vars NOBR t), vars); | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 114 | fun print_match is_first ((IVar NONE, ty), t) = | 
| 41784 | 115 | if Code_Thingol.is_IAbs t andalso is_first | 
| 41781 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 116 | then print_match_val ((IVar NONE, ty), t) | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 117 | else print_match_seq t | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 118 | | print_match _ ((pat, ty), t) = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 119 | print_match_val ((pat, ty), t); | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 120 | val (seps_ps, vars') = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 121 | vars |> print_match true bind ||>> fold_map (print_match false) binds |>> uncurry cons; | 
| 38059 | 122 | val all_seps_ps = seps_ps @ [(true, print_term tyvars false some_thm vars' NOBR body)]; | 
| 123 | fun insert_seps [(_, p)] = [p] | |
| 124 | | insert_seps ((_, p) :: (seps_ps as (sep, _) :: _)) = | |
| 125 | (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: 
41687diff
changeset | 126 |           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: 
48003diff
changeset | 127 |       | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
 | 
| 34294 | 128 | let | 
| 129 | fun print_select (pat, body) = | |
| 130 | let | |
| 37464 | 131 | val (p_pat, vars') = print_bind tyvars some_thm NOBR pat vars; | 
| 132 | val p_body = print_term tyvars false some_thm vars' NOBR body | |
| 133 | in concat [str "case", p_pat, str "=>", p_body] end; | |
| 46850 | 134 | in | 
| 135 | map print_select clauses | |
| 136 |             |> Pretty.block_enclose (concat [print_term tyvars false some_thm vars NOBR t, str "match", str "{"], str "}")
 | |
| 137 | |> single | |
| 138 |             |> enclose "(" ")"
 | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 139 | end; | 
| 37639 | 140 | fun print_context tyvars vs name = applify "[" "]" | 
| 141 | (fn (v, sort) => (Pretty.block o map str) | |
| 41781 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 142 | (lookup_tyvar tyvars v :: maps (fn sort => [" : ", deresolve sort]) sort)) | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 143 | NOBR ((str o deresolve) name) vs; | 
| 37639 | 144 | fun print_defhead tyvars vars name vs params tys ty = | 
| 145 |       Pretty.block [str "def ", constraint (applify "(" ")" (fn (param, ty) =>
 | |
| 146 | constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty)) | |
| 147 | NOBR (print_context tyvars vs name) (params ~~ tys)) (print_typ tyvars NOBR ty), | |
| 148 | str " ="]; | |
| 149 | fun print_def name (vs, ty) [] = | |
| 150 | let | |
| 151 | val (tys, ty') = Code_Thingol.unfold_fun ty; | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 152 | val params = Name.invent (snd reserved) "a" (length tys); | 
| 37639 | 153 | val tyvars = intro_tyvars vs reserved; | 
| 154 | val vars = intro_vars params reserved; | |
| 155 | in | |
| 156 | concat [print_defhead tyvars vars name vs params tys ty', | |
| 48073 
1b609a7837ef
prefer sys.error over plain error in Scala to avoid deprecation warning
 haftmann parents: 
48072diff
changeset | 157 |               str ("sys.error(\"" ^ name ^ "\")")]
 | 
| 37639 | 158 | end | 
| 159 | | print_def name (vs, ty) eqs = | |
| 160 | let | |
| 161 | val tycos = fold (fn ((ts, t), _) => | |
| 162 | fold Code_Thingol.add_tyconames (t :: ts)) eqs []; | |
| 163 | val tyvars = reserved | |
| 164 | |> intro_base_names | |
| 38923 | 165 | (is_none o tyco_syntax) deresolve tycos | 
| 37639 | 166 | |> intro_tyvars vs; | 
| 167 | val simple = case eqs | |
| 168 | of [((ts, _), _)] => forall Code_Thingol.is_IVar ts | |
| 169 | | _ => false; | |
| 170 | val consts = fold Code_Thingol.add_constnames | |
| 171 | (map (snd o fst) eqs) []; | |
| 172 | val vars1 = reserved | |
| 173 | |> intro_base_names | |
| 38923 | 174 | (is_none o const_syntax) deresolve consts | 
| 37639 | 175 | val params = if simple | 
| 176 | then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs | |
| 177 | else aux_params vars1 (map (fst o fst) eqs); | |
| 178 | val vars2 = intro_vars params vars1; | |
| 179 | val (tys', ty') = Code_Thingol.unfold_fun_n (length params) ty; | |
| 38922 | 180 | fun tuplify [p] = p | 
| 181 |               | tuplify ps = enum "," "(" ")" ps;
 | |
| 37639 | 182 | fun print_rhs vars' ((_, t), (some_thm, _)) = | 
| 183 | print_term tyvars false some_thm vars' NOBR t; | |
| 184 | fun print_clause (eq as ((ts, _), (some_thm, _))) = | |
| 34294 | 185 | let | 
| 37639 | 186 | val vars' = intro_vars ((fold o Code_Thingol.fold_varnames) | 
| 187 | (insert (op =)) ts []) vars1; | |
| 188 | in | |
| 189 | concat [str "case", | |
| 38922 | 190 | tuplify (map (print_term tyvars true some_thm vars' NOBR) ts), | 
| 37639 | 191 | str "=>", print_rhs vars' eq] | 
| 192 | end; | |
| 193 | val head = print_defhead tyvars vars2 name vs params tys' ty'; | |
| 194 | in if simple then | |
| 195 | concat [head, print_rhs vars2 (hd eqs)] | |
| 196 | else | |
| 197 | Pretty.block_enclose | |
| 38922 | 198 | (concat [head, tuplify (map (str o lookup_var vars2) params), | 
| 37639 | 199 |                 str "match", str "{"], str "}")
 | 
| 200 | (map print_clause eqs) | |
| 201 | end; | |
| 39023 | 202 | val print_method = str o Library.enclose "`" "`" o deresolve_full; | 
| 37639 | 203 | fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) = | 
| 204 | print_def name (vs, ty) (filter (snd o snd) raw_eqs) | |
| 34294 | 205 | | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) = | 
| 206 | let | |
| 48003 
1d11af40b106
dropped sort constraints on datatype specifications
 haftmann parents: 
47609diff
changeset | 207 | val tyvars = intro_tyvars (map (rpair []) vs) reserved; | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 208 | fun print_co ((co, vs_args), tys) = | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 209 | concat [Pretty.block ((applify "[" "]" (str o lookup_tyvar tyvars) NOBR | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 210 | ((concat o map str) ["final", "case", "class", deresolve co]) vs_args) | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 211 |                 @@ 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: 
48568diff
changeset | 212 | (Name.invent_names (snd reserved) "a" tys))), | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 213 | str "extends", | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 214 | applify "[" "]" (str o lookup_tyvar tyvars) NOBR | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 215 | ((str o deresolve) name) vs | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 216 | ]; | 
| 34294 | 217 | in | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 218 | Pretty.chunks (applify "[" "]" (str o lookup_tyvar tyvars) | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 219 | NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve name]) vs | 
| 37639 | 220 | :: map print_co cos) | 
| 34294 | 221 | end | 
| 37447 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 haftmann parents: 
37446diff
changeset | 222 | | print_stmt (name, Code_Thingol.Class (_, (v, (super_classes, classparams)))) = | 
| 34294 | 223 | let | 
| 37639 | 224 | val tyvars = intro_tyvars [(v, [name])] reserved; | 
| 225 | fun add_typarg s = Pretty.block | |
| 226 | [str s, str "[", (str o lookup_tyvar tyvars) v, str "]"]; | |
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37337diff
changeset | 227 | fun print_super_classes [] = NONE | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37337diff
changeset | 228 | | print_super_classes classes = SOME (concat (str "extends" | 
| 37639 | 229 | :: separate (str "with") (map (add_typarg o deresolve o fst) classes))); | 
| 34294 | 230 | fun print_classparam_val (classparam, ty) = | 
| 37639 | 231 | concat [str "val", constraint (print_method classparam) | 
| 232 | ((print_tupled_typ tyvars o Code_Thingol.unfold_fun) ty)]; | |
| 34294 | 233 | fun print_classparam_def (classparam, ty) = | 
| 234 | let | |
| 235 | val (tys, ty) = Code_Thingol.unfold_fun ty; | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 236 | val [implicit_name] = Name.invent (snd reserved) (lookup_tyvar tyvars v) 1; | 
| 37639 | 237 | val proto_vars = intro_vars [implicit_name] reserved; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 238 | val auxs = Name.invent (snd proto_vars) "a" (length tys); | 
| 37639 | 239 | val vars = intro_vars auxs proto_vars; | 
| 34294 | 240 | in | 
| 37639 | 241 |                 concat [str "def", constraint (Pretty.block [applify "(" ")"
 | 
| 242 | (fn (aux, ty) => constraint ((str o lookup_var vars) aux) | |
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 243 | (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve classparam)) | 
| 37639 | 244 | (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ", | 
| 245 | add_typarg (deresolve name), str ")"]) (print_typ tyvars NOBR ty), str "=", | |
| 246 |                   applify "(" ")" (str o lookup_var vars) NOBR
 | |
| 247 | (Pretty.block [str implicit_name, str ".", print_method classparam]) auxs] | |
| 34294 | 248 | end; | 
| 249 | in | |
| 250 | Pretty.chunks ( | |
| 251 | (Pretty.block_enclose | |
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 252 | (concat ([str "trait", (add_typarg o deresolve) name] | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37337diff
changeset | 253 |                   @ the_list (print_super_classes super_classes) @ [str "{"]), str "}")
 | 
| 34294 | 254 | (map print_classparam_val classparams)) | 
| 255 | :: map print_classparam_def classparams | |
| 256 | ) | |
| 257 | end | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 258 | | print_stmt (name, Code_Thingol.Classinst | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 259 |           { class, tyco, vs, inst_params, superinst_params, ... }) =
 | 
| 34294 | 260 | let | 
| 37639 | 261 | val tyvars = intro_tyvars vs reserved; | 
| 262 | val classtyp = (class, tyco `%% map (ITyVar o fst) vs); | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 263 |             fun print_classparam_instance ((classparam, const as { dom, ... }), (thm, _)) =
 | 
| 37450 | 264 | let | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 265 | val aux_dom = Name.invent_names (snd reserved) "a" dom; | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 266 | val auxs = map fst aux_dom; | 
| 37450 | 267 | val vars = intro_vars auxs reserved; | 
| 37639 | 268 |                 val aux_abstr = if null auxs then [] else [enum "," "(" ")"
 | 
| 269 | (map (fn (aux, ty) => constraint ((str o lookup_var vars) aux) | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 270 | (print_typ tyvars NOBR ty)) aux_dom), str "=>"]; | 
| 41939 | 271 | in | 
| 37639 | 272 | concat ([str "val", print_method classparam, str "="] | 
| 273 | @ aux_abstr @| print_app tyvars false (SOME thm) vars NOBR | |
| 274 | (const, map (IVar o SOME) auxs)) | |
| 37450 | 275 | end; | 
| 37639 | 276 | in | 
| 277 | Pretty.block_enclose (concat [str "implicit def", | |
| 278 | constraint (print_context tyvars vs name) (print_dicttyp tyvars classtyp), | |
| 279 |               str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}")
 | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 280 | (map print_classparam_instance (inst_params @ superinst_params)) | 
| 37639 | 281 | end; | 
| 34294 | 282 | in print_stmt end; | 
| 283 | ||
| 38779 | 284 | fun scala_program_of_program labelled_name reserved module_alias program = | 
| 34294 | 285 | let | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 286 | fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) = | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 287 | let | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 288 | val declare = Name.declare name_fragment; | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 289 | in (name_fragment, ((declare nsp_class, declare nsp_object), declare nsp_common)) end; | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 290 | fun namify_class base ((nsp_class, nsp_object), nsp_common) = | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 291 | let | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
41940diff
changeset | 292 | val (base', nsp_class') = Name.variant base nsp_class | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 293 | in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end; | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 294 | fun namify_object base ((nsp_class, nsp_object), nsp_common) = | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 295 | let | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
41940diff
changeset | 296 | val (base', nsp_object') = Name.variant base nsp_object | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 297 | in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end; | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 298 | fun namify_common upper base ((nsp_class, nsp_object), nsp_common) = | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 299 | let | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 300 | val (base', nsp_common') = | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
41940diff
changeset | 301 | Name.variant (if upper then first_upper base else base) nsp_common | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 302 | in | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 303 | (base', | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 304 | ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common')) | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 305 | end; | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 306 | fun namify_stmt (Code_Thingol.Fun _) = namify_object | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 307 | | namify_stmt (Code_Thingol.Datatype _) = namify_class | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 308 | | namify_stmt (Code_Thingol.Datatypecons _) = namify_common true | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 309 | | namify_stmt (Code_Thingol.Class _) = namify_class | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 310 | | namify_stmt (Code_Thingol.Classrel _) = namify_object | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 311 | | namify_stmt (Code_Thingol.Classparam _) = namify_object | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 312 | | namify_stmt (Code_Thingol.Classinst _) = namify_common false; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 313 | fun memorize_implicits name = | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 314 | let | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 315 | fun is_classinst stmt = case stmt | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 316 | of Code_Thingol.Classinst _ => true | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 317 | | _ => false; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 318 | val implicits = filter (is_classinst o Graph.get_node program) | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43329diff
changeset | 319 | (Graph.immediate_succs program name); | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 320 | in union (op =) implicits end; | 
| 39059 
3a11a667af75
restored and added surpression of case combinators
 haftmann parents: 
39058diff
changeset | 321 | fun modify_stmt (_, Code_Thingol.Fun (_, (_, SOME _))) = NONE | 
| 
3a11a667af75
restored and added surpression of case combinators
 haftmann parents: 
39058diff
changeset | 322 | | modify_stmt (_, Code_Thingol.Datatypecons _) = NONE | 
| 39024 
30d5dd2f30b6
simultaneous modification of statements: statement names
 haftmann parents: 
39023diff
changeset | 323 | | modify_stmt (_, Code_Thingol.Classrel _) = NONE | 
| 
30d5dd2f30b6
simultaneous modification of statements: statement names
 haftmann parents: 
39023diff
changeset | 324 | | modify_stmt (_, Code_Thingol.Classparam _) = NONE | 
| 
30d5dd2f30b6
simultaneous modification of statements: statement names
 haftmann parents: 
39023diff
changeset | 325 | | modify_stmt (_, stmt) = SOME stmt; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 326 | in | 
| 41939 | 327 | Code_Namespace.hierarchical_program labelled_name | 
| 328 |       { module_alias = module_alias, reserved = reserved,
 | |
| 329 | empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module, | |
| 330 | namify_stmt = namify_stmt, cyclic_modules = true, empty_data = [], | |
| 331 | memorize_data = memorize_implicits, modify_stmts = map modify_stmt } program | |
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 332 | end; | 
| 34294 | 333 | |
| 38928 | 334 | fun serialize_scala { labelled_name, reserved_syms, includes,
 | 
| 41343 | 335 | module_alias, class_syntax, tyco_syntax, const_syntax } program = | 
| 34294 | 336 | let | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 337 | |
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 338 | (* build program *) | 
| 39147 | 339 |     val { deresolver, hierarchical_program = scala_program } =
 | 
| 39030 
2bb34f36db80
include names need not be considered as reserved any longer
 haftmann parents: 
39024diff
changeset | 340 | scala_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program; | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 341 | |
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 342 | (* print statements *) | 
| 37639 | 343 | fun lookup_constr tyco constr = case Graph.get_node program tyco | 
| 344 | of Code_Thingol.Datatype (_, (_, constrs)) => | |
| 345 | the (AList.lookup (op = o apsnd fst) constrs constr); | |
| 346 | fun classparams_of_class class = case Graph.get_node program class | |
| 347 | of Code_Thingol.Class (_, (_, (_, classparams))) => classparams; | |
| 34294 | 348 | fun args_num c = case Graph.get_node program c | 
| 37464 | 349 | of Code_Thingol.Fun (_, (((_, ty), []), _)) => | 
| 350 | (length o fst o Code_Thingol.unfold_fun) ty | |
| 37437 | 351 | | Code_Thingol.Fun (_, ((_, ((ts, _), _) :: _), _)) => length ts | 
| 37639 | 352 | | Code_Thingol.Datatypecons (_, tyco) => length (lookup_constr tyco c) | 
| 34294 | 353 | | Code_Thingol.Classparam (_, class) => | 
| 37639 | 354 | (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =) | 
| 355 | (classparams_of_class class)) c; | |
| 47609 | 356 | fun print_stmt prefix_fragments = print_scala_stmt | 
| 39147 | 357 | tyco_syntax const_syntax (make_vars reserved_syms) args_num | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 358 | (Code_Thingol.is_constr program) (deresolver prefix_fragments, deresolver []); | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 359 | |
| 39147 | 360 | (* print modules *) | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 361 | fun print_implicit prefix_fragments implicit = | 
| 38782 | 362 | let | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 363 | val s = deresolver prefix_fragments implicit; | 
| 38782 | 364 | in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end; | 
| 39147 | 365 | fun print_module prefix_fragments base implicits ps = Pretty.chunks2 | 
| 366 |       ([str ("object " ^ base ^ " {")]
 | |
| 367 | @ (case map_filter (print_implicit prefix_fragments) implicits | |
| 368 | of [] => [] | implicit_ps => (single o Pretty.block) | |
| 369 | (str "import /*implicits*/" :: Pretty.brk 1 :: commas implicit_ps)) | |
| 370 |         @ ps @ [str ("} /* object " ^ base ^ " */")]);
 | |
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 371 | |
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 372 | (* serialization *) | 
| 39147 | 373 | val p = Pretty.chunks2 (map snd includes | 
| 374 |       @ Code_Namespace.print_hierarchical {
 | |
| 375 | print_module = print_module, print_stmt = print_stmt, | |
| 376 | lift_markup = I } scala_program); | |
| 39056 | 377 | fun write width NONE = writeln o format [] width | 
| 378 | | write width (SOME p) = File.write p o format [] width; | |
| 48568 | 379 |     fun prepare names width p = ([("", format names width p)], try (deresolver []));
 | 
| 34294 | 380 | in | 
| 48568 | 381 | Code_Target.serialization write prepare p | 
| 34294 | 382 | end; | 
| 383 | ||
| 38966 | 384 | val serializer : Code_Target.serializer = | 
| 385 | Code_Target.parse_args (Scan.succeed ()) #> K serialize_scala; | |
| 386 | ||
| 34294 | 387 | val literals = let | 
| 37224 | 388 | fun char_scala c = if c = "'" then "\\'" | 
| 389 | else if c = "\"" then "\\\"" | |
| 390 | else if c = "\\" then "\\\\" | |
| 391 | else let val k = ord c | |
| 392 | in if k < 32 orelse k > 126 then "\\" ^ radixstring (8, "0", k) else c end | |
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34900diff
changeset | 393 | fun numeral_scala k = if k < 0 | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37932diff
changeset | 394 | then if k > ~ 2147483647 then "- " ^ string_of_int (~ k) | 
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34900diff
changeset | 395 |       else quote ("- " ^ string_of_int (~ k))
 | 
| 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34900diff
changeset | 396 | else if k <= 2147483647 then string_of_int k | 
| 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34900diff
changeset | 397 | else quote (string_of_int k) | 
| 34294 | 398 | in Literals {
 | 
| 399 | literal_char = Library.enclose "'" "'" o char_scala, | |
| 400 | literal_string = quote o translate_string char_scala, | |
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34900diff
changeset | 401 |   literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
 | 
| 34888 | 402 |   literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps],
 | 
| 34294 | 403 | infix_cons = (6, "::") | 
| 404 | } end; | |
| 405 | ||
| 406 | ||
| 407 | (** Isar setup **) | |
| 408 | ||
| 409 | val setup = | |
| 37821 | 410 | Code_Target.add_target | 
| 38966 | 411 |     (target, { serializer = serializer, literals = literals,
 | 
| 41939 | 412 |       check = { env_var = "SCALA_HOME",
 | 
| 413 | make_destination = fn p => Path.append p (Path.explode "ROOT.scala"), | |
| 41940 | 414 | make_command = fn _ => | 
| 415 | "env JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' \"$SCALA_HOME/bin/scalac\" ROOT.scala" } }) | |
| 38923 | 416 | #> Code_Target.add_tyco_syntax target "fun" | 
| 37464 | 417 | (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] => | 
| 418 | brackify_infix (1, R) fxy ( | |
| 419 | print_typ BR ty1 (*product type vs. tupled arguments!*), | |
| 420 | str "=>", | |
| 421 | print_typ (INFX (1, R)) ty2 | |
| 422 | ))) | |
| 34294 | 423 | #> fold (Code_Target.add_reserved target) [ | 
| 424 | "abstract", "case", "catch", "class", "def", "do", "else", "extends", "false", | |
| 425 | "final", "finally", "for", "forSome", "if", "implicit", "import", "lazy", | |
| 426 | "match", "new", "null", "object", "override", "package", "private", "protected", | |
| 427 | "requires", "return", "sealed", "super", "this", "throw", "trait", "try", | |
| 37243 
6e2ac5358d6e
capitalized type variables; added yield as keyword
 haftmann parents: 
37224diff
changeset | 428 | "true", "type", "val", "var", "while", "with", "yield" | 
| 34294 | 429 | ] | 
| 430 | #> fold (Code_Target.add_reserved target) [ | |
| 48073 
1b609a7837ef
prefer sys.error over plain error in Scala to avoid deprecation warning
 haftmann parents: 
48072diff
changeset | 431 | "apply", "sys", "scala", "BigInt", "Nil", "List" | 
| 34294 | 432 | ]; | 
| 433 | ||
| 434 | end; (*struct*) |