| author | wenzelm | 
| Sat, 05 Mar 2016 20:47:31 +0100 | |
| changeset 62522 | d32c23d29968 | 
| parent 59936 | b8ffc3dc9e24 | 
| child 65483 | 1cb9fd58d55e | 
| permissions | -rw-r--r-- | 
| 
37745
 
6315b6426200
checking generated code for various target languages
 
haftmann 
parents: 
37669 
diff
changeset
 | 
1  | 
(* Title: Tools/Code/code_haskell.ML  | 
| 28054 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
4  | 
Serializer for Haskell.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature CODE_HASKELL =  | 
|
8  | 
sig  | 
|
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
9  | 
val language_params: string  | 
| 
37745
 
6315b6426200
checking generated code for various target languages
 
haftmann 
parents: 
37669 
diff
changeset
 | 
10  | 
val target: string  | 
| 
58400
 
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
 
haftmann 
parents: 
58397 
diff
changeset
 | 
11  | 
val print_numeral: string -> int -> string  | 
| 28054 | 12  | 
end;  | 
13  | 
||
14  | 
structure Code_Haskell : CODE_HASKELL =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
val target = "Haskell";  | 
|
18  | 
||
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
19  | 
val language_extensions =  | 
| 44940 | 20  | 
["EmptyDataDecls", "RankNTypes", "ScopedTypeVariables"];  | 
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
21  | 
|
| 
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
22  | 
val language_pragma =  | 
| 44940 | 23  | 
  "{-# LANGUAGE " ^ commas language_extensions ^ " #-}";
 | 
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
24  | 
|
| 
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
25  | 
val language_params =  | 
| 44940 | 26  | 
space_implode " " (map (prefix "-X") language_extensions);  | 
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
27  | 
|
| 55150 | 28  | 
open Basic_Code_Symbol;  | 
| 28054 | 29  | 
open Basic_Code_Thingol;  | 
30  | 
open Code_Printer;  | 
|
31  | 
||
32  | 
infixr 5 @@;  | 
|
33  | 
infixr 5 @|;  | 
|
34  | 
||
35  | 
||
36  | 
(** Haskell serializer **)  | 
|
37  | 
||
| 47609 | 38  | 
fun print_haskell_stmt class_syntax tyco_syntax const_syntax  | 
| 
44793
 
fddb09e6f84d
removing previous crude approximation to add type annotations to disambiguate types
 
bulwahn 
parents: 
44792 
diff
changeset
 | 
39  | 
reserved deresolve deriving_show =  | 
| 28054 | 40  | 
let  | 
| 55150 | 41  | 
val deresolve_const = deresolve o Constant;  | 
42  | 
val deresolve_tyco = deresolve o Type_Constructor;  | 
|
43  | 
val deresolve_class = deresolve o Type_Class;  | 
|
| 38923 | 44  | 
fun class_name class = case class_syntax class  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
45  | 
of NONE => deresolve_class class  | 
| 28687 | 46  | 
| SOME class => class;  | 
| 33991 | 47  | 
fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs  | 
| 28054 | 48  | 
of [] => []  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37242 
diff
changeset
 | 
49  | 
      | constraints => enum "," "(" ")" (
 | 
| 28054 | 50  | 
map (fn (v, class) =>  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37242 
diff
changeset
 | 
51  | 
str (class_name class ^ " " ^ lookup_var tyvars v)) constraints)  | 
| 28054 | 52  | 
@@ str " => ";  | 
| 33991 | 53  | 
fun print_typforall tyvars vs = case map fst vs  | 
| 28054 | 54  | 
of [] => []  | 
55  | 
| vnames => str "forall " :: Pretty.breaks  | 
|
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
56  | 
(map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;  | 
| 33991 | 57  | 
fun print_tyco_expr tyvars fxy (tyco, tys) =  | 
58  | 
brackify fxy (str tyco :: map (print_typ tyvars BR) tys)  | 
|
| 47609 | 59  | 
and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
60  | 
of NONE => print_tyco_expr tyvars fxy (deresolve_tyco tyco, tys)  | 
| 47609 | 61  | 
| SOME (_, print) => print (print_typ tyvars) fxy tys)  | 
| 33991 | 62  | 
| print_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v;  | 
| 
48003
 
1d11af40b106
dropped sort constraints on datatype specifications
 
haftmann 
parents: 
47609 
diff
changeset
 | 
63  | 
fun print_typdecl tyvars (tyco, vs) =  | 
| 
 
1d11af40b106
dropped sort constraints on datatype specifications
 
haftmann 
parents: 
47609 
diff
changeset
 | 
64  | 
print_tyco_expr tyvars NOBR (tyco, map ITyVar vs);  | 
| 33991 | 65  | 
fun print_typscheme tyvars (vs, ty) =  | 
66  | 
Pretty.block (print_typforall tyvars vs @ print_typcontext tyvars vs @| print_typ tyvars NOBR ty);  | 
|
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
67  | 
fun print_term tyvars some_thm vars fxy (IConst const) =  | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
68  | 
print_app tyvars some_thm vars fxy (const, [])  | 
| 35228 | 69  | 
| print_term tyvars some_thm vars fxy (t as (t1 `$ t2)) =  | 
| 28054 | 70  | 
(case Code_Thingol.unfold_const_app t  | 
| 35228 | 71  | 
of SOME app => print_app tyvars some_thm vars fxy app  | 
| 28054 | 72  | 
| _ =>  | 
73  | 
brackify fxy [  | 
|
| 35228 | 74  | 
print_term tyvars some_thm vars NOBR t1,  | 
75  | 
print_term tyvars some_thm vars BR t2  | 
|
| 28054 | 76  | 
])  | 
| 35228 | 77  | 
| print_term tyvars some_thm vars fxy (IVar NONE) =  | 
| 31889 | 78  | 
str "_"  | 
| 35228 | 79  | 
| print_term tyvars some_thm vars fxy (IVar (SOME v)) =  | 
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
80  | 
(str o lookup_var vars) v  | 
| 35228 | 81  | 
| print_term tyvars some_thm vars fxy (t as _ `|=> _) =  | 
| 28054 | 82  | 
let  | 
| 31874 | 83  | 
val (binds, t') = Code_Thingol.unfold_pat_abs t;  | 
| 35228 | 84  | 
val (ps, vars') = fold_map (print_bind tyvars some_thm BR o fst) binds vars;  | 
85  | 
in brackets (str "\\" :: ps @ str "->" @@ print_term tyvars some_thm vars' NOBR t') end  | 
|
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
86  | 
| print_term tyvars some_thm vars fxy (ICase case_expr) =  | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
87  | 
(case Code_Thingol.unfold_const_app (#primitive case_expr)  | 
| 55150 | 88  | 
           of SOME (app as ({ sym = Constant const, ... }, _)) =>
 | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
89  | 
if is_none (const_syntax const)  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
90  | 
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
 | 
91  | 
else print_app tyvars some_thm vars fxy app  | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
92  | 
| NONE => print_case tyvars some_thm vars fxy case_expr)  | 
| 58397 | 93  | 
    and print_app_expr tyvars some_thm vars ({ sym, annotation, ... }, ts) =
 | 
| 
44792
 
26b19918e670
adding minimalistic implementation for printing the type annotations
 
bulwahn 
parents: 
44789 
diff
changeset
 | 
94  | 
let  | 
| 
45009
 
99e1965f9c21
syntactic improvements and tuning names in the code generator due to Florian's code review
 
bulwahn 
parents: 
44940 
diff
changeset
 | 
95  | 
val printed_const =  | 
| 58397 | 96  | 
case annotation of  | 
97  | 
SOME ty => brackets [(str o deresolve) sym, str "::", print_typ tyvars NOBR ty]  | 
|
98  | 
| NONE => (str o deresolve) sym  | 
|
| 
44792
 
26b19918e670
adding minimalistic implementation for printing the type annotations
 
bulwahn 
parents: 
44789 
diff
changeset
 | 
99  | 
in  | 
| 
45009
 
99e1965f9c21
syntactic improvements and tuning names in the code generator due to Florian's code review
 
bulwahn 
parents: 
44940 
diff
changeset
 | 
100  | 
printed_const :: map (print_term tyvars some_thm vars BR) ts  | 
| 
44792
 
26b19918e670
adding minimalistic implementation for printing the type annotations
 
bulwahn 
parents: 
44789 
diff
changeset
 | 
101  | 
end  | 
| 38923 | 102  | 
and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) const_syntax  | 
| 35228 | 103  | 
and print_bind tyvars some_thm fxy p = gen_print_bind (print_term tyvars) some_thm fxy p  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
104  | 
    and print_case tyvars some_thm vars fxy { clauses = [], ... } =
 | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
105  | 
(brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""]  | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
106  | 
      | print_case tyvars some_thm vars fxy (case_expr as { clauses = [_], ... }) =
 | 
| 28054 | 107  | 
let  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
108  | 
val (binds, body) = Code_Thingol.unfold_let (ICase case_expr);  | 
| 47609 | 109  | 
fun print_match ((pat, _), t) vars =  | 
| 28054 | 110  | 
vars  | 
| 35228 | 111  | 
|> print_bind tyvars some_thm BR pat  | 
112  | 
|>> (fn p => semicolon [p, str "=", print_term tyvars some_thm vars NOBR t])  | 
|
| 33991 | 113  | 
val (ps, vars') = fold_map print_match binds vars;  | 
| 31665 | 114  | 
          in brackify_block fxy (str "let {")
 | 
115  | 
ps  | 
|
| 35228 | 116  | 
(concat [str "}", str "in", print_term tyvars some_thm vars' NOBR body])  | 
| 28054 | 117  | 
end  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
118  | 
      | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
 | 
| 28054 | 119  | 
let  | 
| 33991 | 120  | 
fun print_select (pat, body) =  | 
| 28054 | 121  | 
let  | 
| 35228 | 122  | 
val (p, vars') = print_bind tyvars some_thm NOBR pat vars;  | 
123  | 
in semicolon [p, str "->", print_term tyvars some_thm vars' NOBR body] end;  | 
|
| 36576 | 124  | 
in Pretty.block_enclose  | 
125  | 
            (concat [str "(case", print_term tyvars some_thm vars NOBR t, str "of", str "{"], str "})")
 | 
|
| 33991 | 126  | 
(map print_select clauses)  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
127  | 
end;  | 
| 55150 | 128  | 
fun print_stmt (Constant const, Code_Thingol.Fun (((vs, ty), raw_eqs), _)) =  | 
| 28054 | 129  | 
let  | 
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
130  | 
val tyvars = intro_vars (map fst vs) reserved;  | 
| 34269 | 131  | 
fun print_err n =  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
132  | 
(semicolon o map str) (  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
133  | 
deresolve_const const  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
134  | 
:: replicate n "_"  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
135  | 
@ "="  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
136  | 
:: "error"  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
137  | 
@@ ML_Syntax.print_string const  | 
| 34269 | 138  | 
);  | 
| 35228 | 139  | 
fun print_eqn ((ts, t), (some_thm, _)) =  | 
| 28054 | 140  | 
let  | 
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
141  | 
val vars = reserved  | 
| 
55145
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52801 
diff
changeset
 | 
142  | 
|> intro_base_names_for (is_none o const_syntax)  | 
| 
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52801 
diff
changeset
 | 
143  | 
deresolve (t :: ts)  | 
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
144  | 
|> intro_vars ((fold o Code_Thingol.fold_varnames)  | 
| 32925 | 145  | 
(insert (op =)) ts []);  | 
| 28054 | 146  | 
in  | 
147  | 
semicolon (  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
148  | 
(str o deresolve_const) const  | 
| 35228 | 149  | 
:: map (print_term tyvars some_thm vars BR) ts  | 
| 28054 | 150  | 
@ str "="  | 
| 35228 | 151  | 
@@ print_term tyvars some_thm vars NOBR t  | 
| 28054 | 152  | 
)  | 
153  | 
end;  | 
|
154  | 
in  | 
|
155  | 
Pretty.chunks (  | 
|
| 33991 | 156  | 
semicolon [  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
157  | 
(str o suffix " ::" o deresolve_const) const,  | 
| 33991 | 158  | 
print_typscheme tyvars (vs, ty)  | 
| 28054 | 159  | 
]  | 
| 34269 | 160  | 
:: (case filter (snd o snd) raw_eqs  | 
161  | 
of [] => [print_err ((length o fst o Code_Thingol.unfold_fun) ty)]  | 
|
162  | 
| eqs => map print_eqn eqs)  | 
|
| 28054 | 163  | 
)  | 
164  | 
end  | 
|
| 55150 | 165  | 
| print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [])) =  | 
| 28054 | 166  | 
let  | 
| 
48003
 
1d11af40b106
dropped sort constraints on datatype specifications
 
haftmann 
parents: 
47609 
diff
changeset
 | 
167  | 
val tyvars = intro_vars vs reserved;  | 
| 28054 | 168  | 
in  | 
169  | 
semicolon [  | 
|
170  | 
str "data",  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
171  | 
print_typdecl tyvars (deresolve_tyco tyco, vs)  | 
| 28054 | 172  | 
]  | 
173  | 
end  | 
|
| 55150 | 174  | 
| print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, [((co, _), [ty])])) =  | 
| 28054 | 175  | 
let  | 
| 
48003
 
1d11af40b106
dropped sort constraints on datatype specifications
 
haftmann 
parents: 
47609 
diff
changeset
 | 
176  | 
val tyvars = intro_vars vs reserved;  | 
| 28054 | 177  | 
in  | 
178  | 
semicolon (  | 
|
179  | 
str "newtype"  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
180  | 
:: print_typdecl tyvars (deresolve_tyco tyco, vs)  | 
| 28054 | 181  | 
:: str "="  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
182  | 
:: (str o deresolve_const) co  | 
| 33991 | 183  | 
:: print_typ tyvars BR ty  | 
| 55374 | 184  | 
:: (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else [])  | 
| 28054 | 185  | 
)  | 
186  | 
end  | 
|
| 55150 | 187  | 
| print_stmt (Type_Constructor tyco, Code_Thingol.Datatype (vs, co :: cos)) =  | 
| 28054 | 188  | 
let  | 
| 
48003
 
1d11af40b106
dropped sort constraints on datatype specifications
 
haftmann 
parents: 
47609 
diff
changeset
 | 
189  | 
val tyvars = intro_vars vs reserved;  | 
| 37449 | 190  | 
fun print_co ((co, _), tys) =  | 
| 28054 | 191  | 
concat (  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
192  | 
(str o deresolve_const) co  | 
| 33991 | 193  | 
:: map (print_typ tyvars BR) tys  | 
| 28054 | 194  | 
)  | 
195  | 
in  | 
|
196  | 
semicolon (  | 
|
197  | 
str "data"  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
198  | 
:: print_typdecl tyvars (deresolve_tyco tyco, vs)  | 
| 28054 | 199  | 
:: str "="  | 
| 33991 | 200  | 
:: print_co co  | 
201  | 
:: map ((fn p => Pretty.block [str "| ", p]) o print_co) cos  | 
|
| 55374 | 202  | 
@ (if deriving_show tyco then [str "deriving (Prelude.Read, Prelude.Show)"] else [])  | 
| 28054 | 203  | 
)  | 
204  | 
end  | 
|
| 55150 | 205  | 
| print_stmt (Type_Class class, Code_Thingol.Class (v, (classrels, classparams))) =  | 
| 28054 | 206  | 
let  | 
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
207  | 
val tyvars = intro_vars [v] reserved;  | 
| 33991 | 208  | 
fun print_classparam (classparam, ty) =  | 
| 28054 | 209  | 
semicolon [  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
210  | 
(str o deresolve_const) classparam,  | 
| 28054 | 211  | 
str "::",  | 
| 33991 | 212  | 
print_typ tyvars NOBR ty  | 
| 28054 | 213  | 
]  | 
214  | 
in  | 
|
215  | 
Pretty.block_enclose (  | 
|
216  | 
Pretty.block [  | 
|
217  | 
str "class ",  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
218  | 
Pretty.block (print_typcontext tyvars [(v, map snd classrels)]),  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
219  | 
str (deresolve_class class ^ " " ^ lookup_var tyvars v),  | 
| 28054 | 220  | 
                str " where {"
 | 
221  | 
],  | 
|
222  | 
str "};"  | 
|
| 33991 | 223  | 
) (map print_classparam classparams)  | 
| 28054 | 224  | 
end  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
225  | 
      | print_stmt (_, Code_Thingol.Classinst { class, tyco, vs, inst_params, ... }) =
 | 
| 28054 | 226  | 
let  | 
| 
32924
 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 
haftmann 
parents: 
32913 
diff
changeset
 | 
227  | 
val tyvars = intro_vars (map fst vs) reserved;  | 
| 
52519
 
598addf65209
explicit hint for domain of class parameters in instance statements
 
haftmann 
parents: 
52435 
diff
changeset
 | 
228  | 
fun print_classparam_instance ((classparam, (const, _)), (thm, _)) =  | 
| 
55373
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
229  | 
case const_syntax classparam of  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
230  | 
NONE => semicolon [  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
231  | 
(str o Long_Name.base_name o deresolve_const) classparam,  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
232  | 
str "=",  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
233  | 
print_app tyvars (SOME thm) reserved NOBR (const, [])  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
234  | 
]  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
235  | 
| SOME (_, Code_Printer.Plain_printer s) => semicolon [  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
236  | 
(str o Long_Name.base_name) s,  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
237  | 
str "=",  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
238  | 
print_app tyvars (SOME thm) reserved NOBR (const, [])  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
239  | 
]  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
240  | 
| SOME (k, Code_Printer.Complex_printer _) =>  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
241  | 
let  | 
| 58397 | 242  | 
                    val { sym = Constant c, dom, ... } = const;
 | 
| 
55373
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
243  | 
val (vs, rhs) = (apfst o map) fst  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
244  | 
(Code_Thingol.unfold_abs (Code_Thingol.eta_expand k (const, [])));  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
245  | 
val s = if (is_some o const_syntax) c  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
246  | 
then NONE else (SOME o Long_Name.base_name o deresolve_const) c;  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
247  | 
val vars = reserved  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
248  | 
|> intro_vars (map_filter I (s :: vs));  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
249  | 
                    val lhs = IConst { sym = Constant classparam, typargs = [],
 | 
| 58397 | 250  | 
dicts = [], dom = dom, annotation = NONE } `$$ map IVar vs;  | 
| 
55373
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
251  | 
(*dictionaries are not relevant at this late stage,  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
252  | 
and these consts never need type annotations for disambiguation *)  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
253  | 
in  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
254  | 
semicolon [  | 
| 
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
255  | 
print_term tyvars (SOME thm) vars NOBR lhs,  | 
| 28687 | 256  | 
str "=",  | 
| 
55373
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
257  | 
print_term tyvars (SOME thm) vars NOBR rhs  | 
| 28687 | 258  | 
]  | 
| 
55373
 
2b4204cb7904
method names in instance declarations are always unqualified
 
haftmann 
parents: 
55153 
diff
changeset
 | 
259  | 
end;  | 
| 28054 | 260  | 
in  | 
261  | 
Pretty.block_enclose (  | 
|
262  | 
Pretty.block [  | 
|
263  | 
str "instance ",  | 
|
| 33991 | 264  | 
Pretty.block (print_typcontext tyvars vs),  | 
| 28054 | 265  | 
str (class_name class ^ " "),  | 
| 33991 | 266  | 
print_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),  | 
| 28054 | 267  | 
                str " where {"
 | 
268  | 
],  | 
|
269  | 
str "};"  | 
|
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
48003 
diff
changeset
 | 
270  | 
) (map print_classparam_instance inst_params)  | 
| 28054 | 271  | 
end;  | 
| 33991 | 272  | 
in print_stmt end;  | 
| 28054 | 273  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
274  | 
fun haskell_program_of_program ctxt module_prefix module_name reserved identifiers =  | 
| 39202 | 275  | 
let  | 
276  | 
fun namify_fun upper base (nsp_fun, nsp_typ) =  | 
|
277  | 
let  | 
|
| 
56826
 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 
haftmann 
parents: 
56812 
diff
changeset
 | 
278  | 
val (base', nsp_fun') = Name.variant (Name.enforce_case upper base) nsp_fun;  | 
| 39202 | 279  | 
in (base', (nsp_fun', nsp_typ)) end;  | 
280  | 
fun namify_typ base (nsp_fun, nsp_typ) =  | 
|
281  | 
let  | 
|
| 56812 | 282  | 
val (base', nsp_typ') = Name.variant (Name.enforce_case true base) nsp_typ;  | 
| 39202 | 283  | 
in (base', (nsp_fun, nsp_typ')) end;  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
284  | 
fun namify_stmt (Code_Thingol.Fun (_, SOME _)) = pair  | 
| 39202 | 285  | 
| namify_stmt (Code_Thingol.Fun _) = namify_fun false  | 
286  | 
| namify_stmt (Code_Thingol.Datatype _) = namify_typ  | 
|
287  | 
| namify_stmt (Code_Thingol.Datatypecons _) = namify_fun true  | 
|
288  | 
| namify_stmt (Code_Thingol.Class _) = namify_typ  | 
|
289  | 
| namify_stmt (Code_Thingol.Classrel _) = pair  | 
|
290  | 
| namify_stmt (Code_Thingol.Classparam _) = namify_fun false  | 
|
291  | 
| namify_stmt (Code_Thingol.Classinst _) = pair;  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
292  | 
fun select_stmt (Code_Thingol.Fun (_, SOME _)) = false  | 
| 39202 | 293  | 
| select_stmt (Code_Thingol.Fun _) = true  | 
294  | 
| select_stmt (Code_Thingol.Datatype _) = true  | 
|
295  | 
| select_stmt (Code_Thingol.Datatypecons _) = false  | 
|
296  | 
| select_stmt (Code_Thingol.Class _) = true  | 
|
297  | 
| select_stmt (Code_Thingol.Classrel _) = false  | 
|
298  | 
| select_stmt (Code_Thingol.Classparam _) = false  | 
|
299  | 
| select_stmt (Code_Thingol.Classinst _) = true;  | 
|
300  | 
in  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
301  | 
Code_Namespace.flat_program ctxt  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
51143 
diff
changeset
 | 
302  | 
      { module_prefix = module_prefix, module_name = module_name, reserved = reserved,
 | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
51143 
diff
changeset
 | 
303  | 
identifiers = identifiers, empty_nsp = (reserved, reserved), namify_stmt = namify_stmt,  | 
| 39202 | 304  | 
modify_stmt = fn stmt => if select_stmt stmt then SOME stmt else NONE }  | 
305  | 
end;  | 
|
306  | 
||
| 
48431
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
307  | 
val prelude_import_operators = [  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
308  | 
"==", "/=", "<", "<=", ">=", ">", "+", "-", "*", "/", "**", ">>=", ">>", "=<<", "&&", "||", "^", "^^", ".", "$", "$!", "++", "!!"  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
309  | 
];  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
310  | 
|
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
311  | 
val prelude_import_unqualified = [  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
312  | 
"Eq",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
313  | 
"error",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
314  | 
"id",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
315  | 
"return",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
316  | 
"not",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
317  | 
"fst", "snd",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
318  | 
"map", "filter", "concat", "concatMap", "reverse", "zip", "null", "takeWhile", "dropWhile", "all", "any",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
319  | 
"Integer", "negate", "abs", "divMod",  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
320  | 
"String"  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
321  | 
];  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
322  | 
|
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
323  | 
val prelude_import_unqualified_constr = [  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
324  | 
  ("Bool", ["True", "False"]),
 | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
325  | 
  ("Maybe", ["Nothing", "Just"])
 | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
326  | 
];  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
327  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
328  | 
fun serialize_haskell module_prefix string_classes ctxt { module_name,
 | 
| 55683 | 329  | 
reserved_syms, identifiers, includes, class_syntax, tyco_syntax, const_syntax } exports program =  | 
| 28054 | 330  | 
let  | 
| 39204 | 331  | 
|
332  | 
(* build program *)  | 
|
| 38926 | 333  | 
val reserved = fold (insert (op =) o fst) includes reserved_syms;  | 
| 39204 | 334  | 
    val { deresolver, flat_program = haskell_program } = haskell_program_of_program
 | 
| 55683 | 335  | 
ctxt module_prefix module_name (Name.make_context reserved) identifiers exports program;  | 
| 39204 | 336  | 
|
337  | 
(* print statements *)  | 
|
| 28054 | 338  | 
fun deriving_show tyco =  | 
339  | 
let  | 
|
340  | 
fun deriv _ "fun" = false  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
341  | 
| deriv tycos tyco = member (op =) tycos tyco  | 
| 55150 | 342  | 
orelse case try (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
 | 
343  | 
of SOME (Code_Thingol.Datatype (_, cs)) => forall (deriv' (tyco :: tycos))  | 
| 28054 | 344  | 
(maps snd cs)  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
345  | 
| NONE => true  | 
| 28054 | 346  | 
and deriv' tycos (tyco `%% tys) = deriv tycos tyco  | 
347  | 
andalso forall (deriv' tycos) tys  | 
|
348  | 
| deriv' _ (ITyVar _) = true  | 
|
349  | 
in deriv [] tyco end;  | 
|
| 47609 | 350  | 
fun print_stmt deresolve = print_haskell_stmt  | 
| 39204 | 351  | 
class_syntax tyco_syntax const_syntax (make_vars reserved)  | 
| 
45009
 
99e1965f9c21
syntactic improvements and tuning names in the code generator due to Florian's code review
 
bulwahn 
parents: 
44940 
diff
changeset
 | 
352  | 
deresolve (if string_classes then deriving_show else K false);  | 
| 39204 | 353  | 
|
354  | 
(* print modules *)  | 
|
| 55679 | 355  | 
fun print_module_frame module_name exports ps =  | 
| 39204 | 356  | 
(module_name, Pretty.chunks2 (  | 
| 55679 | 357  | 
concat [str "module",  | 
358  | 
case exports of  | 
|
359  | 
            SOME ps => Pretty.block [str module_name, enclose "(" ")" (commas ps)]
 | 
|
360  | 
| NONE => str module_name,  | 
|
361  | 
          str "where {"
 | 
|
362  | 
]  | 
|
| 39204 | 363  | 
:: ps  | 
364  | 
@| str "}"  | 
|
365  | 
));  | 
|
| 55679 | 366  | 
fun print_qualified_import module_name =  | 
367  | 
semicolon [str "import qualified", str module_name];  | 
|
| 
48431
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
368  | 
val import_common_ps =  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
369  | 
      enclose "import Prelude (" ");" (commas (map str
 | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
370  | 
        (map (Library.enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
 | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
371  | 
          @ map (fn (tyco, constrs) => (enclose (tyco ^ "(") ")" o commas o map str) constrs) prelude_import_unqualified_constr))
 | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
372  | 
:: print_qualified_import "Prelude"  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
373  | 
:: map (print_qualified_import o fst) includes;  | 
| 39204 | 374  | 
fun print_module module_name (gr, imports) =  | 
| 28054 | 375  | 
let  | 
| 
48431
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
376  | 
val deresolve = deresolver module_name;  | 
| 55679 | 377  | 
val deresolve_import = SOME o str o deresolve;  | 
378  | 
val deresolve_import_attached = SOME o str o suffix "(..)" o deresolve;  | 
|
| 55681 | 379  | 
fun print_import (sym, (_, Code_Thingol.Fun _)) = deresolve_import sym  | 
380  | 
| print_import (sym, (Code_Namespace.Public, Code_Thingol.Datatype _)) = deresolve_import_attached sym  | 
|
381  | 
| print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Datatype _)) = deresolve_import sym  | 
|
382  | 
| print_import (sym, (Code_Namespace.Public, Code_Thingol.Class _)) = deresolve_import_attached sym  | 
|
383  | 
| print_import (sym, (Code_Namespace.Opaque, Code_Thingol.Class _)) = deresolve_import sym  | 
|
384  | 
| print_import (sym, (_, Code_Thingol.Classinst _)) = NONE;  | 
|
| 
48431
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
385  | 
val import_ps = import_common_ps @ map (print_qualified_import o fst) imports;  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
386  | 
fun print_stmt' sym = case Code_Symbol.Graph.get_node gr sym  | 
| 39204 | 387  | 
of (_, NONE) => NONE  | 
| 55679 | 388  | 
| (_, SOME (export, stmt)) =>  | 
| 55681 | 389  | 
SOME (if Code_Namespace.not_private export then print_import (sym, (export, stmt)) else NONE, markup_stmt sym (print_stmt deresolve (sym, stmt)));  | 
| 55679 | 390  | 
val (export_ps, body_ps) = (flat o rev o Code_Symbol.Graph.strong_conn) gr  | 
391  | 
|> map_filter print_stmt'  | 
|
392  | 
|> split_list  | 
|
393  | 
|> apfst (map_filter I);  | 
|
| 39204 | 394  | 
in  | 
| 55679 | 395  | 
print_module_frame module_name (SOME export_ps)  | 
| 39204 | 396  | 
((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)  | 
397  | 
end;  | 
|
398  | 
||
399  | 
(*serialization*)  | 
|
400  | 
fun write_module width (SOME destination) (module_name, content) =  | 
|
| 38915 | 401  | 
let  | 
| 
42003
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41952 
diff
changeset
 | 
402  | 
val _ = File.check_dir destination;  | 
| 39204 | 403  | 
val filepath = (Path.append destination o Path.ext "hs" o Path.explode o implode  | 
404  | 
o separate "/" o Long_Name.explode) module_name;  | 
|
| 
41246
 
e1da70df68c1
allocate intermediate directories in module hierarchy
 
haftmann 
parents: 
40804 
diff
changeset
 | 
405  | 
val _ = Isabelle_System.mkdirs (Path.dir filepath);  | 
| 39209 | 406  | 
in  | 
407  | 
(File.write filepath o format [] width o Pretty.chunks2)  | 
|
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
408  | 
[str language_pragma, content]  | 
| 39209 | 409  | 
end  | 
| 39056 | 410  | 
| write_module width NONE (_, content) = writeln (format [] width content);  | 
| 28054 | 411  | 
in  | 
| 38916 | 412  | 
Code_Target.serialization  | 
| 38915 | 413  | 
(fn width => fn destination => K () o map (write_module width destination))  | 
| 40705 | 414  | 
(fn present => fn width => rpair (try (deresolver ""))  | 
| 48568 | 415  | 
o (map o apsnd) (format present width))  | 
| 55679 | 416  | 
(map (fn (module_name, content) => print_module_frame module_name NONE [content]) includes  | 
| 39204 | 417  | 
@ map (fn module_name => print_module module_name (Graph.get_node haskell_program module_name))  | 
418  | 
((flat o rev o Graph.strong_conn) haskell_program))  | 
|
| 28054 | 419  | 
end;  | 
420  | 
||
| 38966 | 421  | 
val serializer : Code_Target.serializer =  | 
| 39206 | 422  | 
Code_Target.parse_args (Scan.optional (Args.$$$ "root" -- Args.colon |-- Args.name) ""  | 
| 38966 | 423  | 
-- Scan.optional (Args.$$$ "string_classes" >> K true) false  | 
424  | 
>> (fn (module_prefix, string_classes) =>  | 
|
425  | 
serialize_haskell module_prefix string_classes));  | 
|
426  | 
||
| 
58400
 
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
 
haftmann 
parents: 
58397 
diff
changeset
 | 
427  | 
fun print_numeral typ = Library.enclose "(" (" :: " ^ typ ^ ")") o signed_string_of_int;
 | 
| 
 
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
 
haftmann 
parents: 
58397 
diff
changeset
 | 
428  | 
|
| 28064 | 429  | 
val literals = let  | 
430  | 
fun char_haskell c =  | 
|
431  | 
let  | 
|
432  | 
val s = ML_Syntax.print_char c;  | 
|
433  | 
in if s = "'" then "\\'" else s end;  | 
|
434  | 
in Literals {
 | 
|
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34085 
diff
changeset
 | 
435  | 
literal_char = Library.enclose "'" "'" o char_haskell,  | 
| 28064 | 436  | 
literal_string = quote o translate_string char_haskell,  | 
| 
58400
 
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
 
haftmann 
parents: 
58397 
diff
changeset
 | 
437  | 
literal_numeral = print_numeral "Integer",  | 
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34085 
diff
changeset
 | 
438  | 
literal_list = enum "," "[" "]",  | 
| 28064 | 439  | 
infix_cons = (5, ":")  | 
440  | 
} end;  | 
|
441  | 
||
| 28054 | 442  | 
|
443  | 
(** optional monad syntax **)  | 
|
444  | 
||
445  | 
fun pretty_haskell_monad c_bind =  | 
|
446  | 
let  | 
|
| 31874 | 447  | 
fun dest_bind t1 t2 = case Code_Thingol.split_pat_abs t2  | 
| 31889 | 448  | 
of SOME ((pat, ty), t') =>  | 
449  | 
SOME ((SOME ((pat, ty), true), t1), t')  | 
|
| 28145 | 450  | 
| NONE => NONE;  | 
| 55150 | 451  | 
    fun dest_monad (IConst { sym = Constant c, ... } `$ t1 `$ t2) =
 | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
452  | 
if c = c_bind then dest_bind t1 t2  | 
| 28145 | 453  | 
else NONE  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
454  | 
| dest_monad t = case Code_Thingol.split_let t  | 
| 28145 | 455  | 
of SOME (((pat, ty), tbind), t') =>  | 
| 31889 | 456  | 
SOME ((SOME ((pat, ty), false), tbind), t')  | 
| 28145 | 457  | 
| NONE => NONE;  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
458  | 
val implode_monad = Code_Thingol.unfoldr dest_monad;  | 
| 33991 | 459  | 
fun print_monad print_bind print_term (NONE, t) vars =  | 
460  | 
(semicolon [print_term vars NOBR t], vars)  | 
|
461  | 
| print_monad print_bind print_term (SOME ((bind, _), true), t) vars = vars  | 
|
462  | 
|> print_bind NOBR bind  | 
|
463  | 
|>> (fn p => semicolon [p, str "<-", print_term vars NOBR t])  | 
|
464  | 
| print_monad print_bind print_term (SOME ((bind, _), false), t) vars = vars  | 
|
465  | 
|> print_bind NOBR bind  | 
|
| 37832 | 466  | 
          |>> (fn p => semicolon [str "let", str "{", p, str "=", print_term vars NOBR t, str "}"]);
 | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
467  | 
fun pretty _ print_term thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2  | 
| 28145 | 468  | 
of SOME (bind, t') => let  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
469  | 
val (binds, t'') = implode_monad t'  | 
| 33991 | 470  | 
val (ps, vars') = fold_map (print_monad (gen_print_bind (K print_term) thm) print_term)  | 
471  | 
(bind :: binds) vars;  | 
|
472  | 
in  | 
|
| 58454 | 473  | 
          brackify_block fxy (str "do { ")
 | 
| 33991 | 474  | 
(ps @| print_term vars' NOBR t'')  | 
| 58454 | 475  | 
(str " }")  | 
| 33991 | 476  | 
end  | 
| 28145 | 477  | 
| NONE => brackify_infix (1, L) fxy  | 
| 
37242
 
97097e589715
brackify_infix etc.: no break before infix operator -- eases survival in Scala
 
haftmann 
parents: 
36960 
diff
changeset
 | 
478  | 
(print_term vars (INFX (1, L)) t1, str ">>=", print_term vars (INFX (1, X)) t2)  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
479  | 
in (2, pretty) end;  | 
| 28054 | 480  | 
|
| 28145 | 481  | 
fun add_monad target' raw_c_bind thy =  | 
| 28054 | 482  | 
let  | 
| 31156 | 483  | 
val c_bind = Code.read_const thy raw_c_bind;  | 
| 52801 | 484  | 
in  | 
485  | 
if target = target' then  | 
|
486  | 
thy  | 
|
| 55150 | 487  | 
|> Code_Target.set_printings (Constant (c_bind, [(target,  | 
| 52801 | 488  | 
SOME (Code_Printer.complex_const_syntax (pretty_haskell_monad c_bind)))]))  | 
489  | 
else error "Only Haskell target allows for monad syntax"  | 
|
490  | 
end;  | 
|
| 28054 | 491  | 
|
492  | 
||
493  | 
(** Isar setup **)  | 
|
494  | 
||
| 59323 | 495  | 
val _ = Theory.setup  | 
496  | 
(Code_Target.add_language  | 
|
| 38966 | 497  | 
    (target, { serializer = serializer, literals = literals,
 | 
| 
41952
 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 
wenzelm 
parents: 
41940 
diff
changeset
 | 
498  | 
      check = { env_var = "ISABELLE_GHC", make_destination = I,
 | 
| 41940 | 499  | 
make_command = fn module_name =>  | 
| 
44926
 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 
noschinl 
parents: 
44852 
diff
changeset
 | 
500  | 
"\"$ISABELLE_GHC\" " ^ language_params ^ " -odir build -hidir build -stubdir build -e \"\" " ^  | 
| 41940 | 501  | 
module_name ^ ".hs" } })  | 
| 55150 | 502  | 
  #> Code_Target.set_printings (Type_Constructor ("fun",
 | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52138 
diff
changeset
 | 
503  | 
[(target, SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>  | 
| 
37242
 
97097e589715
brackify_infix etc.: no break before infix operator -- eases survival in Scala
 
haftmann 
parents: 
36960 
diff
changeset
 | 
504  | 
brackify_infix (1, R) fxy (  | 
| 33991 | 505  | 
print_typ (INFX (1, X)) ty1,  | 
| 28054 | 506  | 
str "->",  | 
| 33991 | 507  | 
print_typ (INFX (1, R)) ty2  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52138 
diff
changeset
 | 
508  | 
)))]))  | 
| 28054 | 509  | 
#> fold (Code_Target.add_reserved target) [  | 
510  | 
"hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",  | 
|
511  | 
"import", "default", "forall", "let", "in", "class", "qualified", "data",  | 
|
512  | 
"newtype", "instance", "if", "then", "else", "type", "as", "do", "module"  | 
|
513  | 
]  | 
|
| 
48431
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
514  | 
#> fold (Code_Target.add_reserved target) prelude_import_unqualified  | 
| 
 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 
haftmann 
parents: 
48072 
diff
changeset
 | 
515  | 
#> fold (Code_Target.add_reserved target o fst) prelude_import_unqualified_constr  | 
| 59323 | 516  | 
#> fold (fold (Code_Target.add_reserved target) o snd) prelude_import_unqualified_constr);  | 
517  | 
||
518  | 
val _ =  | 
|
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
59323 
diff
changeset
 | 
519  | 
  Outer_Syntax.command @{command_keyword code_monad} "define code syntax for monads"
 | 
| 59323 | 520  | 
(Parse.term -- Parse.name >> (fn (raw_bind, target) =>  | 
521  | 
Toplevel.theory (add_monad target raw_bind)));  | 
|
| 28054 | 522  | 
|
523  | 
end; (*struct*)  |