| author | bulwahn | 
| Tue, 09 Jun 2009 13:40:57 +0200 | |
| changeset 31515 | 62fc203eed88 | 
| parent 31376 | 4356b52b03f7 | 
| child 31665 | a1f4d3b3f6c8 | 
| permissions | -rw-r--r-- | 
| 28054 | 1 | (* Title: Tools/code/code_haskell.ML | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | ||
| 4 | Serializer for Haskell. | |
| 5 | *) | |
| 6 | ||
| 7 | signature CODE_HASKELL = | |
| 8 | sig | |
| 9 | val setup: theory -> theory | |
| 10 | end; | |
| 11 | ||
| 12 | structure Code_Haskell : CODE_HASKELL = | |
| 13 | struct | |
| 14 | ||
| 15 | val target = "Haskell"; | |
| 16 | ||
| 17 | open Basic_Code_Thingol; | |
| 18 | open Code_Printer; | |
| 19 | ||
| 20 | infixr 5 @@; | |
| 21 | infixr 5 @|; | |
| 22 | ||
| 23 | ||
| 24 | (** Haskell serializer **) | |
| 25 | ||
| 26 | fun pr_haskell_bind pr_term = | |
| 27 | let | |
| 28 | fun pr_bind ((NONE, NONE), _) = str "_" | |
| 29 | | pr_bind ((SOME v, NONE), _) = str v | |
| 30 | | pr_bind ((NONE, SOME p), _) = p | |
| 31 | | pr_bind ((SOME v, SOME p), _) = brackets [str v, str "@", p]; | |
| 32 | in gen_pr_bind pr_bind pr_term end; | |
| 33 | ||
| 31054 | 34 | fun pr_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const | 
| 28054 | 35 | init_syms deresolve is_cons contr_classparam_typs deriving_show = | 
| 36 | let | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 37 | val deresolve_base = Long_Name.base_name o deresolve; | 
| 28054 | 38 | fun class_name class = case syntax_class class | 
| 39 | of NONE => deresolve class | |
| 28687 | 40 | | SOME class => class; | 
| 28054 | 41 | fun pr_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs | 
| 42 | of [] => [] | |
| 43 |       | classbinds => Pretty.enum "," "(" ")" (
 | |
| 44 | map (fn (v, class) => | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 45 | str (class_name class ^ " " ^ Code_Printer.lookup_var tyvars v)) classbinds) | 
| 28054 | 46 | @@ str " => "; | 
| 47 | fun pr_typforall tyvars vs = case map fst vs | |
| 48 | of [] => [] | |
| 49 | | vnames => str "forall " :: Pretty.breaks | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 50 | (map (str o Code_Printer.lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1; | 
| 28054 | 51 | fun pr_tycoexpr tyvars fxy (tyco, tys) = | 
| 52 | brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) | |
| 53 | and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco | |
| 54 | of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys) | |
| 55 | | SOME (i, pr) => pr (pr_typ tyvars) fxy tys) | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 56 | | pr_typ tyvars fxy (ITyVar v) = (str o Code_Printer.lookup_var tyvars) v; | 
| 28054 | 57 | fun pr_typdecl tyvars (vs, tycoexpr) = | 
| 58 | Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr); | |
| 59 | fun pr_typscheme tyvars (vs, ty) = | |
| 60 | Pretty.block (pr_typforall tyvars vs @ pr_typcontext tyvars vs @| pr_typ tyvars NOBR ty); | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 61 | fun pr_term tyvars thm vars fxy (IConst c) = | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 62 | pr_app tyvars thm vars fxy (c, []) | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 63 | | pr_term tyvars thm vars fxy (t as (t1 `$ t2)) = | 
| 28054 | 64 | (case Code_Thingol.unfold_const_app t | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 65 | of SOME app => pr_app tyvars thm vars fxy app | 
| 28054 | 66 | | _ => | 
| 67 | brackify fxy [ | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 68 | pr_term tyvars thm vars NOBR t1, | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 69 | pr_term tyvars thm vars BR t2 | 
| 28054 | 70 | ]) | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 71 | | pr_term tyvars thm vars fxy (IVar v) = | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 72 | (str o Code_Printer.lookup_var vars) v | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 73 | | pr_term tyvars thm vars fxy (t as _ `|-> _) = | 
| 28054 | 74 | let | 
| 75 | val (binds, t') = Code_Thingol.unfold_abs t; | |
| 76 | fun pr ((v, pat), ty) = pr_bind tyvars thm BR ((SOME v, pat), ty); | |
| 77 | val (ps, vars') = fold_map pr binds vars; | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 78 | in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm vars' NOBR t') end | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 79 | | pr_term tyvars thm vars fxy (ICase (cases as (_, t0))) = | 
| 28054 | 80 | (case Code_Thingol.unfold_const_app t0 | 
| 81 | of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) | |
| 82 | then pr_case tyvars thm vars fxy cases | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 83 | else pr_app tyvars thm vars fxy c_ts | 
| 28054 | 84 | | NONE => pr_case tyvars thm vars fxy cases) | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 85 | and pr_app' tyvars thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 86 | of [] => (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts | 
| 28054 | 87 | | fingerprint => let | 
| 88 | val ts_fingerprint = ts ~~ curry Library.take (length ts) fingerprint; | |
| 89 | val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) => | |
| 90 | (not o Code_Thingol.locally_monomorphic) t) ts_fingerprint; | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 91 | fun pr_term_anno (t, NONE) _ = pr_term tyvars thm vars BR t | 
| 28054 | 92 | | pr_term_anno (t, SOME _) ty = | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 93 | brackets [pr_term tyvars thm vars NOBR t, str "::", pr_typ tyvars NOBR ty]; | 
| 28054 | 94 | in | 
| 95 | if needs_annotation then | |
| 96 | (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry Library.take (length ts) tys) | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 97 | else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts | 
| 28054 | 98 | end | 
| 31054 | 99 | and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const | 
| 28054 | 100 | and pr_bind tyvars = pr_haskell_bind (pr_term tyvars) | 
| 101 | and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) = | |
| 102 | let | |
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29832diff
changeset | 103 | val (binds, body) = Code_Thingol.unfold_let (ICase cases); | 
| 28054 | 104 | fun pr ((pat, ty), t) vars = | 
| 105 | vars | |
| 106 | |> pr_bind tyvars thm BR ((NONE, SOME pat), ty) | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 107 | |>> (fn p => semicolon [p, str "=", pr_term tyvars thm vars NOBR t]) | 
| 28054 | 108 | val (ps, vars') = fold_map pr binds vars; | 
| 31376 | 109 | in | 
| 28054 | 110 | Pretty.block_enclose ( | 
| 31376 | 111 |               str "(let {",
 | 
| 112 | concat [str "}", str "in", pr_term tyvars thm vars' NOBR body, str ")"] | |
| 28054 | 113 | ) ps | 
| 114 | end | |
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29832diff
changeset | 115 | | pr_case tyvars thm vars fxy (((t, ty), clauses as _ :: _), _) = | 
| 28054 | 116 | let | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29832diff
changeset | 117 | fun pr (pat, body) = | 
| 28054 | 118 | let | 
| 119 | val (p, vars') = pr_bind tyvars thm NOBR ((NONE, SOME pat), ty) vars; | |
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29832diff
changeset | 120 | in semicolon [p, str "->", pr_term tyvars thm vars' NOBR body] end; | 
| 28054 | 121 | in | 
| 122 | Pretty.block_enclose ( | |
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29832diff
changeset | 123 |               concat [str "(case", pr_term tyvars thm vars NOBR t, str "of", str "{"],
 | 
| 28054 | 124 | str "})" | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29832diff
changeset | 125 | ) (map pr clauses) | 
| 28054 | 126 | end | 
| 31376 | 127 | | pr_case tyvars thm vars fxy ((_, []), _) = | 
| 128 | (brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""]; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 129 | fun pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), []))) = | 
| 28054 | 130 | let | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 131 | val tyvars = Code_Printer.intro_vars (map fst vs) init_syms; | 
| 28054 | 132 | val n = (length o fst o Code_Thingol.unfold_fun) ty; | 
| 133 | in | |
| 134 | Pretty.chunks [ | |
| 135 | Pretty.block [ | |
| 136 | (str o suffix " ::" o deresolve_base) name, | |
| 137 | Pretty.brk 1, | |
| 138 | pr_typscheme tyvars (vs, ty), | |
| 139 | str ";" | |
| 140 | ], | |
| 141 | concat ( | |
| 142 | (str o deresolve_base) name | |
| 143 | :: map str (replicate n "_") | |
| 144 | @ str "=" | |
| 145 | :: str "error" | |
| 146 | @@ (str o (fn s => s ^ ";") o ML_Syntax.print_string | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 147 | o Long_Name.base_name o Long_Name.qualifier) name | 
| 28054 | 148 | ) | 
| 149 | ] | |
| 150 | end | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 151 | | pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) = | 
| 28054 | 152 | let | 
| 28350 | 153 | val eqs = filter (snd o snd) raw_eqs; | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 154 | val tyvars = Code_Printer.intro_vars (map fst vs) init_syms; | 
| 28350 | 155 | fun pr_eq ((ts, t), (thm, _)) = | 
| 28054 | 156 | let | 
| 157 | val consts = map_filter | |
| 158 | (fn c => if (is_some o syntax_const) c | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 159 | then NONE else (SOME o Long_Name.base_name o deresolve) c) | 
| 28054 | 160 | ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); | 
| 161 | val vars = init_syms | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 162 | |> Code_Printer.intro_vars consts | 
| 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 163 | |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) | 
| 28054 | 164 | (insert (op =)) ts []); | 
| 165 | in | |
| 166 | semicolon ( | |
| 167 | (str o deresolve_base) name | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 168 | :: map (pr_term tyvars thm vars BR) ts | 
| 28054 | 169 | @ str "=" | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 170 | @@ pr_term tyvars thm vars NOBR t | 
| 28054 | 171 | ) | 
| 172 | end; | |
| 173 | in | |
| 174 | Pretty.chunks ( | |
| 175 | Pretty.block [ | |
| 176 | (str o suffix " ::" o deresolve_base) name, | |
| 177 | Pretty.brk 1, | |
| 178 | pr_typscheme tyvars (vs, ty), | |
| 179 | str ";" | |
| 180 | ] | |
| 181 | :: map pr_eq eqs | |
| 182 | ) | |
| 183 | end | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 184 | | pr_stmt (name, Code_Thingol.Datatype (_, (vs, []))) = | 
| 28054 | 185 | let | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 186 | val tyvars = Code_Printer.intro_vars (map fst vs) init_syms; | 
| 28054 | 187 | in | 
| 188 | semicolon [ | |
| 189 | str "data", | |
| 190 | pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) | |
| 191 | ] | |
| 192 | end | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 193 | | pr_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) = | 
| 28054 | 194 | let | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 195 | val tyvars = Code_Printer.intro_vars (map fst vs) init_syms; | 
| 28054 | 196 | in | 
| 197 | semicolon ( | |
| 198 | str "newtype" | |
| 199 | :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) | |
| 200 | :: str "=" | |
| 201 | :: (str o deresolve_base) co | |
| 202 | :: pr_typ tyvars BR ty | |
| 203 | :: (if deriving_show name then [str "deriving (Read, Show)"] else []) | |
| 204 | ) | |
| 205 | end | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 206 | | pr_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) = | 
| 28054 | 207 | let | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 208 | val tyvars = Code_Printer.intro_vars (map fst vs) init_syms; | 
| 28054 | 209 | fun pr_co (co, tys) = | 
| 210 | concat ( | |
| 211 | (str o deresolve_base) co | |
| 212 | :: map (pr_typ tyvars BR) tys | |
| 213 | ) | |
| 214 | in | |
| 215 | semicolon ( | |
| 216 | str "data" | |
| 217 | :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) | |
| 218 | :: str "=" | |
| 219 | :: pr_co co | |
| 220 | :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos | |
| 221 | @ (if deriving_show name then [str "deriving (Read, Show)"] else []) | |
| 222 | ) | |
| 223 | end | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 224 | | pr_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) = | 
| 28054 | 225 | let | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 226 | val tyvars = Code_Printer.intro_vars [v] init_syms; | 
| 28054 | 227 | fun pr_classparam (classparam, ty) = | 
| 228 | semicolon [ | |
| 28687 | 229 | (str o deresolve_base) classparam, | 
| 28054 | 230 | str "::", | 
| 231 | pr_typ tyvars NOBR ty | |
| 232 | ] | |
| 233 | in | |
| 234 | Pretty.block_enclose ( | |
| 235 | Pretty.block [ | |
| 236 | str "class ", | |
| 237 | Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]), | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 238 | str (deresolve_base name ^ " " ^ Code_Printer.lookup_var tyvars v), | 
| 28054 | 239 |                 str " where {"
 | 
| 240 | ], | |
| 241 | str "};" | |
| 242 | ) (map pr_classparam classparams) | |
| 243 | end | |
| 244 | | pr_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) = | |
| 245 | let | |
| 28687 | 246 | val split_abs_pure = (fn (v, _) `|-> t => SOME (v, t) | _ => NONE); | 
| 247 | val unfold_abs_pure = Code_Thingol.unfoldr split_abs_pure; | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 248 | val tyvars = Code_Printer.intro_vars (map fst vs) init_syms; | 
| 28687 | 249 | fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam | 
| 250 | of NONE => semicolon [ | |
| 251 | (str o deresolve_base) classparam, | |
| 252 | str "=", | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 253 | pr_app tyvars thm init_syms NOBR (c_inst, []) | 
| 28687 | 254 | ] | 
| 255 | | SOME (k, pr) => | |
| 256 | let | |
| 257 | val (c_inst_name, (_, tys)) = c_inst; | |
| 258 | val const = if (is_some o syntax_const) c_inst_name | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 259 | then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name; | 
| 28687 | 260 | val proto_rhs = Code_Thingol.eta_expand k (c_inst, []); | 
| 261 | val (vs, rhs) = unfold_abs_pure proto_rhs; | |
| 262 | val vars = init_syms | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 263 | |> Code_Printer.intro_vars (the_list const) | 
| 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 264 | |> Code_Printer.intro_vars vs; | 
| 31049 | 265 | val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs; | 
| 28687 | 266 | (*dictionaries are not relevant at this late stage*) | 
| 267 | in | |
| 268 | semicolon [ | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 269 | pr_term tyvars thm vars NOBR lhs, | 
| 28687 | 270 | str "=", | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 271 | pr_term tyvars thm vars NOBR rhs | 
| 28687 | 272 | ] | 
| 273 | end; | |
| 28054 | 274 | in | 
| 275 | Pretty.block_enclose ( | |
| 276 | Pretty.block [ | |
| 277 | str "instance ", | |
| 278 | Pretty.block (pr_typcontext tyvars vs), | |
| 279 | str (class_name class ^ " "), | |
| 280 | pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), | |
| 281 |                 str " where {"
 | |
| 282 | ], | |
| 283 | str "};" | |
| 284 | ) (map pr_instdef classparam_insts) | |
| 285 | end; | |
| 286 | in pr_stmt end; | |
| 287 | ||
| 288 | fun haskell_program_of_program labelled_name module_name module_prefix reserved_names raw_module_alias program = | |
| 289 | let | |
| 290 | val module_alias = if is_some module_name then K module_name else raw_module_alias; | |
| 291 | val reserved_names = Name.make_context reserved_names; | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 292 | val mk_name_module = Code_Printer.mk_name_module reserved_names module_prefix module_alias program; | 
| 28054 | 293 | fun add_stmt (name, (stmt, deps)) = | 
| 294 | let | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 295 | val (module_name, base) = Code_Printer.dest_name name; | 
| 28054 | 296 | val module_name' = mk_name_module module_name; | 
| 297 | val mk_name_stmt = yield_singleton Name.variants; | |
| 298 | fun add_fun upper (nsp_fun, nsp_typ) = | |
| 299 | let | |
| 300 | val (base', nsp_fun') = | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 301 | mk_name_stmt (if upper then Code_Printer.first_upper base else base) nsp_fun | 
| 28054 | 302 | in (base', (nsp_fun', nsp_typ)) end; | 
| 303 | fun add_typ (nsp_fun, nsp_typ) = | |
| 304 | let | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 305 | val (base', nsp_typ') = mk_name_stmt (Code_Printer.first_upper base) nsp_typ | 
| 28054 | 306 | in (base', (nsp_fun, nsp_typ')) end; | 
| 307 | val add_name = case stmt | |
| 308 | of Code_Thingol.Fun _ => add_fun false | |
| 309 | | Code_Thingol.Datatype _ => add_typ | |
| 310 | | Code_Thingol.Datatypecons _ => add_fun true | |
| 311 | | Code_Thingol.Class _ => add_typ | |
| 312 | | Code_Thingol.Classrel _ => pair base | |
| 313 | | Code_Thingol.Classparam _ => add_fun false | |
| 314 | | Code_Thingol.Classinst _ => pair base; | |
| 315 | fun add_stmt' base' = case stmt | |
| 316 | of Code_Thingol.Datatypecons _ => | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 317 | cons (name, (Long_Name.append module_name' base', NONE)) | 
| 28054 | 318 | | Code_Thingol.Classrel _ => I | 
| 319 | | Code_Thingol.Classparam _ => | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 320 | cons (name, (Long_Name.append module_name' base', NONE)) | 
| 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 321 | | _ => cons (name, (Long_Name.append module_name' base', SOME stmt)); | 
| 28054 | 322 | in | 
| 323 | Symtab.map_default (module_name', ([], ([], (reserved_names, reserved_names)))) | |
| 324 | (apfst (fold (insert (op = : string * string -> bool)) deps)) | |
| 325 | #> `(fn program => add_name ((snd o snd o the o Symtab.lookup program) module_name')) | |
| 326 | #-> (fn (base', names) => | |
| 327 | (Symtab.map_entry module_name' o apsnd) (fn (stmts, _) => | |
| 328 | (add_stmt' base' stmts, names))) | |
| 329 | end; | |
| 330 | val hs_program = fold add_stmt (AList.make (fn name => | |
| 331 | (Graph.get_node program name, Graph.imm_succs program name)) | |
| 332 | (Graph.strong_conn program |> flat)) Symtab.empty; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 333 | fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 334 | o Symtab.lookup hs_program) ((mk_name_module o fst o Code_Printer.dest_name) name))) name | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 335 |       handle Option => error ("Unknown statement name: " ^ labelled_name name);
 | 
| 28054 | 336 | in (deresolver, hs_program) end; | 
| 337 | ||
| 338 | fun serialize_haskell module_prefix raw_module_name string_classes labelled_name | |
| 29711 
64d41ad4ffc2
avoid name clash of generated modules and includes
 haftmann parents: 
29135diff
changeset | 339 | raw_reserved_names includes raw_module_alias | 
| 31054 | 340 | syntax_class syntax_tyco syntax_const program cs destination = | 
| 28054 | 341 | let | 
| 342 | val stmt_names = Code_Target.stmt_names_of_destination destination; | |
| 343 | val module_name = if null stmt_names then raw_module_name else SOME "Code"; | |
| 29711 
64d41ad4ffc2
avoid name clash of generated modules and includes
 haftmann parents: 
29135diff
changeset | 344 | val reserved_names = fold (insert (op =) o fst) includes raw_reserved_names; | 
| 28054 | 345 | val (deresolver, hs_program) = haskell_program_of_program labelled_name | 
| 346 | module_name module_prefix reserved_names raw_module_alias program; | |
| 347 | val is_cons = Code_Thingol.is_cons program; | |
| 348 | val contr_classparam_typs = Code_Thingol.contr_classparam_typs program; | |
| 349 | fun deriving_show tyco = | |
| 350 | let | |
| 351 | fun deriv _ "fun" = false | |
| 352 | | deriv tycos tyco = member (op =) tycos tyco orelse | |
| 353 | case try (Graph.get_node program) tyco | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 354 | of SOME (Code_Thingol.Datatype (_, (_, cs))) => forall (deriv' (tyco :: tycos)) | 
| 28054 | 355 | (maps snd cs) | 
| 356 | | NONE => true | |
| 357 | and deriv' tycos (tyco `%% tys) = deriv tycos tyco | |
| 358 | andalso forall (deriv' tycos) tys | |
| 359 | | deriv' _ (ITyVar _) = true | |
| 360 | in deriv [] tyco end; | |
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 361 | val reserved_names = Code_Printer.make_vars reserved_names; | 
| 31054 | 362 | fun pr_stmt qualified = pr_haskell_stmt labelled_name | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 363 | syntax_class syntax_tyco syntax_const reserved_names | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 364 | (if qualified then deresolver else Long_Name.base_name o deresolver) | 
| 28054 | 365 | is_cons contr_classparam_typs | 
| 366 | (if string_classes then deriving_show else K false); | |
| 367 | fun pr_module name content = | |
| 368 | (name, Pretty.chunks [ | |
| 369 |         str ("module " ^ name ^ " where {"),
 | |
| 370 | str "", | |
| 371 | content, | |
| 372 | str "", | |
| 373 | str "}" | |
| 374 | ]); | |
| 375 | fun serialize_module1 (module_name', (deps, (stmts, _))) = | |
| 376 | let | |
| 377 | val stmt_names = map fst stmts; | |
| 378 | val deps' = subtract (op =) stmt_names deps | |
| 379 | |> distinct (op =) | |
| 380 | |> map_filter (try deresolver); | |
| 381 | val qualified = is_none module_name andalso | |
| 382 | map deresolver stmt_names @ deps' | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 383 | |> map Long_Name.base_name | 
| 28054 | 384 | |> has_duplicates (op =); | 
| 385 | val imports = deps' | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 386 | |> map Long_Name.qualifier | 
| 28054 | 387 | |> distinct (op =); | 
| 29794 | 388 |         fun pr_import_include (name, _) = str ("import qualified " ^ name ^ ";");
 | 
| 28054 | 389 | val pr_import_module = str o (if qualified | 
| 390 | then prefix "import qualified " | |
| 391 | else prefix "import ") o suffix ";"; | |
| 392 | val content = Pretty.chunks ( | |
| 393 | map pr_import_include includes | |
| 394 | @ map pr_import_module imports | |
| 395 | @ str "" | |
| 396 | :: separate (str "") (map_filter | |
| 397 | (fn (name, (_, SOME stmt)) => SOME (pr_stmt qualified (name, stmt)) | |
| 398 | | (_, (_, NONE)) => NONE) stmts) | |
| 399 | ) | |
| 400 | in pr_module module_name' content end; | |
| 401 | fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks ( | |
| 402 | separate (str "") (map_filter | |
| 403 | (fn (name, (_, SOME stmt)) => if null stmt_names | |
| 404 | orelse member (op =) stmt_names name | |
| 405 | then SOME (pr_stmt false (name, stmt)) | |
| 406 | else NONE | |
| 407 | | (_, (_, NONE)) => NONE) stmts)); | |
| 408 | val serialize_module = | |
| 409 | if null stmt_names then serialize_module1 else pair "" o serialize_module2; | |
| 29832 
b4919260eaec
check for destination directory, do not allocate it gratuitously
 haftmann parents: 
29794diff
changeset | 410 | fun check_destination destination = | 
| 
b4919260eaec
check for destination directory, do not allocate it gratuitously
 haftmann parents: 
29794diff
changeset | 411 | (File.check destination; destination); | 
| 28054 | 412 | fun write_module destination (modlname, content) = | 
| 413 | let | |
| 414 | val filename = case modlname | |
| 415 | of "" => Path.explode "Main.hs" | |
| 416 | | _ => (Path.ext "hs" o Path.explode o implode o separate "/" | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 417 | o Long_Name.explode) modlname; | 
| 28054 | 418 | val pathname = Path.append destination filename; | 
| 419 | val _ = File.mkdir (Path.dir pathname); | |
| 29135 | 420 | in File.write pathname | 
| 421 |         ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n"
 | |
| 422 | ^ Code_Target.code_of_pretty content) | |
| 423 | end | |
| 28054 | 424 | in | 
| 425 | Code_Target.mk_serialization target NONE | |
| 29832 
b4919260eaec
check for destination directory, do not allocate it gratuitously
 haftmann parents: 
29794diff
changeset | 426 | (fn NONE => K () o map (Code_Target.code_writeln o snd) | SOME file => K () o map | 
| 
b4919260eaec
check for destination directory, do not allocate it gratuitously
 haftmann parents: 
29794diff
changeset | 427 | (write_module (check_destination file))) | 
| 28054 | 428 | (rpair [] o cat_lines o map (Code_Target.code_of_pretty o snd)) | 
| 429 | (map (uncurry pr_module) includes | |
| 430 | @ map serialize_module (Symtab.dest hs_program)) | |
| 431 | destination | |
| 432 | end; | |
| 433 | ||
| 28064 | 434 | val literals = let | 
| 435 | fun char_haskell c = | |
| 436 | let | |
| 437 | val s = ML_Syntax.print_char c; | |
| 438 | in if s = "'" then "\\'" else s end; | |
| 439 | in Literals {
 | |
| 440 | literal_char = enclose "'" "'" o char_haskell, | |
| 441 | literal_string = quote o translate_string char_haskell, | |
| 442 | literal_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k | |
| 443 |     else enclose "(" ")" (signed_string_of_int k),
 | |
| 444 | literal_list = Pretty.enum "," "[" "]", | |
| 445 | infix_cons = (5, ":") | |
| 446 | } end; | |
| 447 | ||
| 28054 | 448 | |
| 449 | (** optional monad syntax **) | |
| 450 | ||
| 451 | fun pretty_haskell_monad c_bind = | |
| 452 | let | |
| 28145 | 453 | fun dest_bind t1 t2 = case Code_Thingol.split_abs t2 | 
| 454 | of SOME (((v, pat), ty), t') => | |
| 455 | SOME ((SOME (((SOME v, pat), ty), true), t1), t') | |
| 456 | | NONE => NONE; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 457 | fun dest_monad c_bind_name (IConst (c, _) `$ t1 `$ t2) = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 458 | if c = c_bind_name then dest_bind t1 t2 | 
| 28145 | 459 | else NONE | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 460 | | dest_monad _ t = case Code_Thingol.split_let t | 
| 28145 | 461 | of SOME (((pat, ty), tbind), t') => | 
| 462 | SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t') | |
| 463 | | NONE => NONE; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 464 | fun implode_monad c_bind_name = Code_Thingol.unfoldr (dest_monad c_bind_name); | 
| 28145 | 465 | fun pr_monad pr_bind pr (NONE, t) vars = | 
| 466 | (semicolon [pr vars NOBR t], vars) | |
| 467 | | pr_monad pr_bind pr (SOME (bind, true), t) vars = vars | |
| 468 | |> pr_bind NOBR bind | |
| 469 | |>> (fn p => semicolon [p, str "<-", pr vars NOBR t]) | |
| 470 | | pr_monad pr_bind pr (SOME (bind, false), t) vars = vars | |
| 471 | |> pr_bind NOBR bind | |
| 472 | |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]); | |
| 31054 | 473 | fun pretty _ [c_bind'] pr thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2 | 
| 28145 | 474 | of SOME (bind, t') => let | 
| 31054 | 475 | val (binds, t'') = implode_monad c_bind' t' | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28687diff
changeset | 476 | val (ps, vars') = fold_map (pr_monad (pr_haskell_bind (K pr) thm) pr) (bind :: binds) vars; | 
| 28145 | 477 |         in (brackify fxy o single o Pretty.enclose "do {" "}" o Pretty.breaks) (ps @| pr vars' NOBR t'') end
 | 
| 478 | | NONE => brackify_infix (1, L) fxy | |
| 479 | [pr vars (INFX (1, L)) t1, str ">>=", pr vars (INFX (1, X)) t2] | |
| 31054 | 480 | in (2, ([c_bind], pretty)) end; | 
| 28054 | 481 | |
| 28145 | 482 | fun add_monad target' raw_c_bind thy = | 
| 28054 | 483 | let | 
| 31156 | 484 | val c_bind = Code.read_const thy raw_c_bind; | 
| 28054 | 485 | in if target = target' then | 
| 486 | thy | |
| 28145 | 487 | |> Code_Target.add_syntax_const target c_bind | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 488 | (SOME (pretty_haskell_monad c_bind)) | 
| 28054 | 489 | else error "Only Haskell target allows for monad syntax" end; | 
| 490 | ||
| 491 | ||
| 492 | (** Isar setup **) | |
| 493 | ||
| 494 | fun isar_seri_haskell module = | |
| 495 | Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name) | |
| 496 | -- Scan.optional (Args.$$$ "string_classes" >> K true) false | |
| 497 | >> (fn (module_prefix, string_classes) => | |
| 498 | serialize_haskell module_prefix module string_classes)); | |
| 499 | ||
| 500 | val _ = | |
| 501 | OuterSyntax.command "code_monad" "define code syntax for monads" OuterKeyword.thy_decl ( | |
| 28145 | 502 | OuterParse.term_group -- OuterParse.name >> (fn (raw_bind, target) => | 
| 503 | Toplevel.theory (add_monad target raw_bind)) | |
| 28054 | 504 | ); | 
| 505 | ||
| 506 | val setup = | |
| 28064 | 507 | Code_Target.add_target (target, (isar_seri_haskell, literals)) | 
| 28054 | 508 | #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => | 
| 509 | brackify_infix (1, R) fxy [ | |
| 510 | pr_typ (INFX (1, X)) ty1, | |
| 511 | str "->", | |
| 512 | pr_typ (INFX (1, R)) ty2 | |
| 513 | ])) | |
| 514 | #> fold (Code_Target.add_reserved target) [ | |
| 515 | "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", | |
| 516 | "import", "default", "forall", "let", "in", "class", "qualified", "data", | |
| 517 | "newtype", "instance", "if", "then", "else", "type", "as", "do", "module" | |
| 518 | ] | |
| 519 | #> fold (Code_Target.add_reserved target) [ | |
| 520 | "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int", | |
| 521 | "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded", | |
| 522 | "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor", | |
| 523 | "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException", | |
| 524 | "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException", | |
| 525 | "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure", | |
| 526 | "ExitSuccess", "False", "GT", "HeapOverflow", | |
| 527 | "IOError", "IOException", "IllegalOperation", | |
| 528 | "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError", | |
| 529 | "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow", | |
| 530 | "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError", | |
| 531 | "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow", | |
| 532 | "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow", | |
| 533 | "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all", | |
| 534 | "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan", | |
| 535 | "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen", | |
| 536 | "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break", | |
| 537 | "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const", | |
| 538 | "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod", | |
| 539 | "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem", | |
| 540 | "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo", | |
| 541 | "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip", | |
| 542 | "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational", | |
| 543 | "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble", | |
| 544 | "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj", | |
| 545 | "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head", | |
| 546 | "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha", | |
| 547 | "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite", | |
| 548 | "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'", | |
| 549 | "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log", | |
| 550 | "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum", | |
| 551 | "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem", | |
| 552 | "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo", | |
| 553 | "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred", | |
| 554 | "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn", | |
| 555 | "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat", | |
| 556 | "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile", | |
| 557 | "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn", | |
| 558 | "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational", | |
| 559 | "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse", | |
| 560 | "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence", | |
| 561 | "sequence_", "show", "showChar", "showException", "showField", "showList", | |
| 562 | "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand", | |
| 563 | "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract", | |
| 564 | "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult", | |
| 565 | "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry", | |
| 566 | "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords", | |
| 567 | "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3" | |
| 568 | ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*); | |
| 569 | ||
| 570 | end; (*struct*) |