haftmann@28054: (* Title: Tools/code/code_haskell.ML haftmann@28054: Author: Florian Haftmann, TU Muenchen haftmann@28054: haftmann@28054: Serializer for Haskell. haftmann@28054: *) haftmann@28054: haftmann@28054: signature CODE_HASKELL = haftmann@28054: sig haftmann@28054: val setup: theory -> theory haftmann@28054: end; haftmann@28054: haftmann@28054: structure Code_Haskell : CODE_HASKELL = haftmann@28054: struct haftmann@28054: haftmann@28054: val target = "Haskell"; haftmann@28054: haftmann@28054: open Basic_Code_Thingol; haftmann@28054: open Code_Printer; haftmann@28054: haftmann@28054: infixr 5 @@; haftmann@28054: infixr 5 @|; haftmann@28054: haftmann@28054: haftmann@28054: (** Haskell serializer **) haftmann@28054: haftmann@31054: fun pr_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const haftmann@32924: reserved deresolve contr_classparam_typs deriving_show = haftmann@28054: let wenzelm@30364: val deresolve_base = Long_Name.base_name o deresolve; haftmann@28054: fun class_name class = case syntax_class class haftmann@28054: of NONE => deresolve class haftmann@28687: | SOME class => class; haftmann@28054: fun pr_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs haftmann@28054: of [] => [] haftmann@28054: | classbinds => Pretty.enum "," "(" ")" ( haftmann@28054: map (fn (v, class) => haftmann@32924: str (class_name class ^ " " ^ lookup_var tyvars v)) classbinds) haftmann@28054: @@ str " => "; haftmann@28054: fun pr_typforall tyvars vs = case map fst vs haftmann@28054: of [] => [] haftmann@28054: | vnames => str "forall " :: Pretty.breaks haftmann@32924: (map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1; haftmann@28054: fun pr_tycoexpr tyvars fxy (tyco, tys) = haftmann@28054: brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) haftmann@28054: and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco haftmann@28054: of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys) haftmann@28054: | SOME (i, pr) => pr (pr_typ tyvars) fxy tys) haftmann@32924: | pr_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v; haftmann@28054: fun pr_typdecl tyvars (vs, tycoexpr) = haftmann@28054: Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr); haftmann@28054: fun pr_typscheme tyvars (vs, ty) = haftmann@28054: Pretty.block (pr_typforall tyvars vs @ pr_typcontext tyvars vs @| pr_typ tyvars NOBR ty); haftmann@28708: fun pr_term tyvars thm vars fxy (IConst c) = haftmann@28708: pr_app tyvars thm vars fxy (c, []) haftmann@28708: | pr_term tyvars thm vars fxy (t as (t1 `$ t2)) = haftmann@28054: (case Code_Thingol.unfold_const_app t haftmann@28708: of SOME app => pr_app tyvars thm vars fxy app haftmann@28054: | _ => haftmann@28054: brackify fxy [ haftmann@28708: pr_term tyvars thm vars NOBR t1, haftmann@28708: pr_term tyvars thm vars BR t2 haftmann@28054: ]) haftmann@31889: | pr_term tyvars thm vars fxy (IVar NONE) = haftmann@31889: str "_" haftmann@31889: | pr_term tyvars thm vars fxy (IVar (SOME v)) = haftmann@32924: (str o lookup_var vars) v haftmann@31724: | pr_term tyvars thm vars fxy (t as _ `|=> _) = haftmann@28054: let haftmann@31874: val (binds, t') = Code_Thingol.unfold_pat_abs t; haftmann@31889: val (ps, vars') = fold_map (pr_bind tyvars thm BR o fst) binds vars; haftmann@28708: in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm vars' NOBR t') end haftmann@28708: | pr_term tyvars thm vars fxy (ICase (cases as (_, t0))) = haftmann@28054: (case Code_Thingol.unfold_const_app t0 haftmann@28054: of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c) haftmann@28054: then pr_case tyvars thm vars fxy cases haftmann@28708: else pr_app tyvars thm vars fxy c_ts haftmann@28054: | NONE => pr_case tyvars thm vars fxy cases) haftmann@28708: and pr_app' tyvars thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c haftmann@28708: of [] => (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts haftmann@28054: | fingerprint => let haftmann@33955: val ts_fingerprint = ts ~~ curry (uncurry take) (length ts) fingerprint; haftmann@28054: val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) => haftmann@28054: (not o Code_Thingol.locally_monomorphic) t) ts_fingerprint; haftmann@28708: fun pr_term_anno (t, NONE) _ = pr_term tyvars thm vars BR t haftmann@28054: | pr_term_anno (t, SOME _) ty = haftmann@28708: brackets [pr_term tyvars thm vars NOBR t, str "::", pr_typ tyvars NOBR ty]; haftmann@28054: in haftmann@28054: if needs_annotation then haftmann@33955: (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry (uncurry take) (length ts) tys) haftmann@28708: else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts haftmann@28054: end haftmann@31054: and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const haftmann@31889: and pr_bind tyvars thm fxy p = gen_pr_bind (pr_term tyvars) thm fxy p haftmann@28054: and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) = haftmann@28054: let haftmann@29952: val (binds, body) = Code_Thingol.unfold_let (ICase cases); haftmann@28054: fun pr ((pat, ty), t) vars = haftmann@28054: vars haftmann@31889: |> pr_bind tyvars thm BR pat haftmann@28708: |>> (fn p => semicolon [p, str "=", pr_term tyvars thm vars NOBR t]) haftmann@28054: val (ps, vars') = fold_map pr binds vars; haftmann@31665: in brackify_block fxy (str "let {") haftmann@31665: ps haftmann@31665: (concat [str "}", str "in", pr_term tyvars thm vars' NOBR body]) haftmann@28054: end haftmann@29952: | pr_case tyvars thm vars fxy (((t, ty), clauses as _ :: _), _) = haftmann@28054: let haftmann@29952: fun pr (pat, body) = haftmann@28054: let haftmann@31889: val (p, vars') = pr_bind tyvars thm NOBR pat vars; haftmann@29952: in semicolon [p, str "->", pr_term tyvars thm vars' NOBR body] end; haftmann@31665: in brackify_block fxy haftmann@31665: (concat [str "case", pr_term tyvars thm vars NOBR t, str "of", str "{"]) haftmann@31665: (map pr clauses) haftmann@31665: (str "}") haftmann@28054: end haftmann@31376: | pr_case tyvars thm vars fxy ((_, []), _) = haftmann@31376: (brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""]; haftmann@28663: fun pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), []))) = haftmann@28054: let haftmann@32924: val tyvars = intro_vars (map fst vs) reserved; haftmann@28054: val n = (length o fst o Code_Thingol.unfold_fun) ty; haftmann@28054: in haftmann@28054: Pretty.chunks [ haftmann@28054: Pretty.block [ haftmann@28054: (str o suffix " ::" o deresolve_base) name, haftmann@28054: Pretty.brk 1, haftmann@28054: pr_typscheme tyvars (vs, ty), haftmann@28054: str ";" haftmann@28054: ], haftmann@28054: concat ( haftmann@28054: (str o deresolve_base) name haftmann@28054: :: map str (replicate n "_") haftmann@28054: @ str "=" haftmann@28054: :: str "error" haftmann@28054: @@ (str o (fn s => s ^ ";") o ML_Syntax.print_string wenzelm@30364: o Long_Name.base_name o Long_Name.qualifier) name haftmann@28054: ) haftmann@28054: ] haftmann@28054: end haftmann@28663: | pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) = haftmann@28054: let haftmann@28350: val eqs = filter (snd o snd) raw_eqs; haftmann@32924: val tyvars = intro_vars (map fst vs) reserved; haftmann@28350: fun pr_eq ((ts, t), (thm, _)) = haftmann@28054: let haftmann@32913: val consts = fold Code_Thingol.add_constnames (t :: ts) []; haftmann@32924: val vars = reserved haftmann@32924: |> intro_base_names haftmann@32913: (is_none o syntax_const) deresolve consts haftmann@32924: |> intro_vars ((fold o Code_Thingol.fold_varnames) haftmann@32925: (insert (op =)) ts []); haftmann@28054: in haftmann@28054: semicolon ( haftmann@28054: (str o deresolve_base) name haftmann@28708: :: map (pr_term tyvars thm vars BR) ts haftmann@28054: @ str "=" haftmann@28708: @@ pr_term tyvars thm vars NOBR t haftmann@28054: ) haftmann@28054: end; haftmann@28054: in haftmann@28054: Pretty.chunks ( haftmann@28054: Pretty.block [ haftmann@28054: (str o suffix " ::" o deresolve_base) name, haftmann@28054: Pretty.brk 1, haftmann@28054: pr_typscheme tyvars (vs, ty), haftmann@28054: str ";" haftmann@28054: ] haftmann@28054: :: map pr_eq eqs haftmann@28054: ) haftmann@28054: end haftmann@28663: | pr_stmt (name, Code_Thingol.Datatype (_, (vs, []))) = haftmann@28054: let haftmann@32924: val tyvars = intro_vars (map fst vs) reserved; haftmann@28054: in haftmann@28054: semicolon [ haftmann@28054: str "data", haftmann@28054: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) haftmann@28054: ] haftmann@28054: end haftmann@28663: | pr_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) = haftmann@28054: let haftmann@32924: val tyvars = intro_vars (map fst vs) reserved; haftmann@28054: in haftmann@28054: semicolon ( haftmann@28054: str "newtype" haftmann@28054: :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) haftmann@28054: :: str "=" haftmann@28054: :: (str o deresolve_base) co haftmann@28054: :: pr_typ tyvars BR ty haftmann@28054: :: (if deriving_show name then [str "deriving (Read, Show)"] else []) haftmann@28054: ) haftmann@28054: end haftmann@28663: | pr_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) = haftmann@28054: let haftmann@32924: val tyvars = intro_vars (map fst vs) reserved; haftmann@28054: fun pr_co (co, tys) = haftmann@28054: concat ( haftmann@28054: (str o deresolve_base) co haftmann@28054: :: map (pr_typ tyvars BR) tys haftmann@28054: ) haftmann@28054: in haftmann@28054: semicolon ( haftmann@28054: str "data" haftmann@28054: :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs)) haftmann@28054: :: str "=" haftmann@28054: :: pr_co co haftmann@28054: :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos haftmann@28054: @ (if deriving_show name then [str "deriving (Read, Show)"] else []) haftmann@28054: ) haftmann@28054: end haftmann@28663: | pr_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) = haftmann@28054: let haftmann@32924: val tyvars = intro_vars [v] reserved; haftmann@28054: fun pr_classparam (classparam, ty) = haftmann@28054: semicolon [ haftmann@28687: (str o deresolve_base) classparam, haftmann@28054: str "::", haftmann@28054: pr_typ tyvars NOBR ty haftmann@28054: ] haftmann@28054: in haftmann@28054: Pretty.block_enclose ( haftmann@28054: Pretty.block [ haftmann@28054: str "class ", haftmann@28054: Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]), haftmann@32924: str (deresolve_base name ^ " " ^ lookup_var tyvars v), haftmann@28054: str " where {" haftmann@28054: ], haftmann@28054: str "};" haftmann@28054: ) (map pr_classparam classparams) haftmann@28054: end haftmann@28054: | pr_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) = haftmann@28054: let haftmann@32924: val tyvars = intro_vars (map fst vs) reserved; haftmann@28687: fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam haftmann@28687: of NONE => semicolon [ haftmann@28687: (str o deresolve_base) classparam, haftmann@28687: str "=", haftmann@32924: pr_app tyvars thm reserved NOBR (c_inst, []) haftmann@28687: ] haftmann@28687: | SOME (k, pr) => haftmann@28687: let haftmann@28687: val (c_inst_name, (_, tys)) = c_inst; haftmann@28687: val const = if (is_some o syntax_const) c_inst_name wenzelm@30364: then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name; haftmann@28687: val proto_rhs = Code_Thingol.eta_expand k (c_inst, []); haftmann@31874: val (vs, rhs) = (apfst o map) fst (Code_Thingol.unfold_abs proto_rhs); haftmann@32924: val vars = reserved haftmann@32924: |> intro_vars (the_list const) haftmann@32924: |> intro_vars (map_filter I vs); haftmann@31889: val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs; haftmann@28687: (*dictionaries are not relevant at this late stage*) haftmann@28687: in haftmann@28687: semicolon [ haftmann@28708: pr_term tyvars thm vars NOBR lhs, haftmann@28687: str "=", haftmann@28708: pr_term tyvars thm vars NOBR rhs haftmann@28687: ] haftmann@28687: end; haftmann@28054: in haftmann@28054: Pretty.block_enclose ( haftmann@28054: Pretty.block [ haftmann@28054: str "instance ", haftmann@28054: Pretty.block (pr_typcontext tyvars vs), haftmann@28054: str (class_name class ^ " "), haftmann@28054: pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), haftmann@28054: str " where {" haftmann@28054: ], haftmann@28054: str "};" haftmann@28054: ) (map pr_instdef classparam_insts) haftmann@28054: end; haftmann@28054: in pr_stmt end; haftmann@28054: haftmann@32924: fun haskell_program_of_program labelled_name module_name module_prefix reserved raw_module_alias program = haftmann@28054: let haftmann@28054: val module_alias = if is_some module_name then K module_name else raw_module_alias; haftmann@32924: val reserved = Name.make_context reserved; haftmann@32924: val mk_name_module = mk_name_module reserved module_prefix module_alias program; haftmann@28054: fun add_stmt (name, (stmt, deps)) = haftmann@28054: let haftmann@32924: val (module_name, base) = dest_name name; haftmann@28054: val module_name' = mk_name_module module_name; haftmann@28054: val mk_name_stmt = yield_singleton Name.variants; haftmann@28054: fun add_fun upper (nsp_fun, nsp_typ) = haftmann@28054: let haftmann@28054: val (base', nsp_fun') = haftmann@32924: mk_name_stmt (if upper then first_upper base else base) nsp_fun haftmann@28054: in (base', (nsp_fun', nsp_typ)) end; haftmann@28054: fun add_typ (nsp_fun, nsp_typ) = haftmann@28054: let haftmann@32924: val (base', nsp_typ') = mk_name_stmt (first_upper base) nsp_typ haftmann@28054: in (base', (nsp_fun, nsp_typ')) end; haftmann@28054: val add_name = case stmt haftmann@28054: of Code_Thingol.Fun _ => add_fun false haftmann@28054: | Code_Thingol.Datatype _ => add_typ haftmann@28054: | Code_Thingol.Datatypecons _ => add_fun true haftmann@28054: | Code_Thingol.Class _ => add_typ haftmann@28054: | Code_Thingol.Classrel _ => pair base haftmann@28054: | Code_Thingol.Classparam _ => add_fun false haftmann@28054: | Code_Thingol.Classinst _ => pair base; haftmann@28054: fun add_stmt' base' = case stmt haftmann@28054: of Code_Thingol.Datatypecons _ => wenzelm@30364: cons (name, (Long_Name.append module_name' base', NONE)) haftmann@28054: | Code_Thingol.Classrel _ => I haftmann@28054: | Code_Thingol.Classparam _ => wenzelm@30364: cons (name, (Long_Name.append module_name' base', NONE)) wenzelm@30364: | _ => cons (name, (Long_Name.append module_name' base', SOME stmt)); haftmann@28054: in haftmann@32924: Symtab.map_default (module_name', ([], ([], (reserved, reserved)))) haftmann@28054: (apfst (fold (insert (op = : string * string -> bool)) deps)) haftmann@28054: #> `(fn program => add_name ((snd o snd o the o Symtab.lookup program) module_name')) haftmann@28054: #-> (fn (base', names) => haftmann@28054: (Symtab.map_entry module_name' o apsnd) (fn (stmts, _) => haftmann@28054: (add_stmt' base' stmts, names))) haftmann@28054: end; haftmann@28054: val hs_program = fold add_stmt (AList.make (fn name => haftmann@28054: (Graph.get_node program name, Graph.imm_succs program name)) haftmann@28054: (Graph.strong_conn program |> flat)) Symtab.empty; haftmann@28663: fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the haftmann@32924: o Symtab.lookup hs_program) ((mk_name_module o fst o dest_name) name))) name haftmann@28663: handle Option => error ("Unknown statement name: " ^ labelled_name name); haftmann@28054: in (deresolver, hs_program) end; haftmann@28054: haftmann@28054: fun serialize_haskell module_prefix raw_module_name string_classes labelled_name haftmann@32924: raw_reserved includes raw_module_alias haftmann@31054: syntax_class syntax_tyco syntax_const program cs destination = haftmann@28054: let haftmann@28054: val stmt_names = Code_Target.stmt_names_of_destination destination; haftmann@28054: val module_name = if null stmt_names then raw_module_name else SOME "Code"; haftmann@32924: val reserved = fold (insert (op =) o fst) includes raw_reserved; haftmann@28054: val (deresolver, hs_program) = haskell_program_of_program labelled_name haftmann@32924: module_name module_prefix reserved raw_module_alias program; haftmann@28054: val contr_classparam_typs = Code_Thingol.contr_classparam_typs program; haftmann@28054: fun deriving_show tyco = haftmann@28054: let haftmann@28054: fun deriv _ "fun" = false haftmann@28054: | deriv tycos tyco = member (op =) tycos tyco orelse haftmann@28054: case try (Graph.get_node program) tyco haftmann@28663: of SOME (Code_Thingol.Datatype (_, (_, cs))) => forall (deriv' (tyco :: tycos)) haftmann@28054: (maps snd cs) haftmann@28054: | NONE => true haftmann@28054: and deriv' tycos (tyco `%% tys) = deriv tycos tyco haftmann@28054: andalso forall (deriv' tycos) tys haftmann@28054: | deriv' _ (ITyVar _) = true haftmann@28054: in deriv [] tyco end; haftmann@32924: val reserved = make_vars reserved; haftmann@31054: fun pr_stmt qualified = pr_haskell_stmt labelled_name haftmann@32924: syntax_class syntax_tyco syntax_const reserved wenzelm@30364: (if qualified then deresolver else Long_Name.base_name o deresolver) haftmann@32903: contr_classparam_typs haftmann@28054: (if string_classes then deriving_show else K false); haftmann@28054: fun pr_module name content = haftmann@28054: (name, Pretty.chunks [ haftmann@28054: str ("module " ^ name ^ " where {"), haftmann@28054: str "", haftmann@28054: content, haftmann@28054: str "", haftmann@28054: str "}" haftmann@28054: ]); haftmann@28054: fun serialize_module1 (module_name', (deps, (stmts, _))) = haftmann@28054: let haftmann@28054: val stmt_names = map fst stmts; haftmann@33421: val qualified = is_none module_name; haftmann@33421: val imports = subtract (op =) stmt_names deps haftmann@28054: |> distinct (op =) haftmann@33421: |> map_filter (try deresolver) wenzelm@30364: |> map Long_Name.qualifier haftmann@28054: |> distinct (op =); haftmann@29794: fun pr_import_include (name, _) = str ("import qualified " ^ name ^ ";"); haftmann@28054: val pr_import_module = str o (if qualified haftmann@28054: then prefix "import qualified " haftmann@28054: else prefix "import ") o suffix ";"; haftmann@28054: val content = Pretty.chunks ( haftmann@28054: map pr_import_include includes haftmann@28054: @ map pr_import_module imports haftmann@28054: @ str "" haftmann@28054: :: separate (str "") (map_filter haftmann@28054: (fn (name, (_, SOME stmt)) => SOME (pr_stmt qualified (name, stmt)) haftmann@28054: | (_, (_, NONE)) => NONE) stmts) haftmann@28054: ) haftmann@28054: in pr_module module_name' content end; haftmann@28054: fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks ( haftmann@28054: separate (str "") (map_filter haftmann@28054: (fn (name, (_, SOME stmt)) => if null stmt_names haftmann@28054: orelse member (op =) stmt_names name haftmann@28054: then SOME (pr_stmt false (name, stmt)) haftmann@28054: else NONE haftmann@28054: | (_, (_, NONE)) => NONE) stmts)); haftmann@28054: val serialize_module = haftmann@28054: if null stmt_names then serialize_module1 else pair "" o serialize_module2; haftmann@29832: fun check_destination destination = haftmann@29832: (File.check destination; destination); haftmann@28054: fun write_module destination (modlname, content) = haftmann@28054: let haftmann@28054: val filename = case modlname haftmann@28054: of "" => Path.explode "Main.hs" haftmann@28054: | _ => (Path.ext "hs" o Path.explode o implode o separate "/" wenzelm@30364: o Long_Name.explode) modlname; haftmann@28054: val pathname = Path.append destination filename; haftmann@28054: val _ = File.mkdir (Path.dir pathname); haftmann@29135: in File.write pathname haftmann@29135: ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n" haftmann@29135: ^ Code_Target.code_of_pretty content) haftmann@29135: end haftmann@28054: in haftmann@28054: Code_Target.mk_serialization target NONE haftmann@29832: (fn NONE => K () o map (Code_Target.code_writeln o snd) | SOME file => K () o map haftmann@29832: (write_module (check_destination file))) haftmann@28054: (rpair [] o cat_lines o map (Code_Target.code_of_pretty o snd)) haftmann@28054: (map (uncurry pr_module) includes haftmann@28054: @ map serialize_module (Symtab.dest hs_program)) haftmann@28054: destination haftmann@28054: end; haftmann@28054: haftmann@28064: val literals = let haftmann@28064: fun char_haskell c = haftmann@28064: let haftmann@28064: val s = ML_Syntax.print_char c; haftmann@28064: in if s = "'" then "\\'" else s end; haftmann@28064: in Literals { haftmann@28064: literal_char = enclose "'" "'" o char_haskell, haftmann@28064: literal_string = quote o translate_string char_haskell, haftmann@28064: literal_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k haftmann@28064: else enclose "(" ")" (signed_string_of_int k), haftmann@28064: literal_list = Pretty.enum "," "[" "]", haftmann@28064: infix_cons = (5, ":") haftmann@28064: } end; haftmann@28064: haftmann@28054: haftmann@28054: (** optional monad syntax **) haftmann@28054: haftmann@28054: fun pretty_haskell_monad c_bind = haftmann@28054: let haftmann@31874: fun dest_bind t1 t2 = case Code_Thingol.split_pat_abs t2 haftmann@31889: of SOME ((pat, ty), t') => haftmann@31889: SOME ((SOME ((pat, ty), true), t1), t') haftmann@28145: | NONE => NONE; haftmann@28663: fun dest_monad c_bind_name (IConst (c, _) `$ t1 `$ t2) = haftmann@28663: if c = c_bind_name then dest_bind t1 t2 haftmann@28145: else NONE haftmann@28663: | dest_monad _ t = case Code_Thingol.split_let t haftmann@28145: of SOME (((pat, ty), tbind), t') => haftmann@31889: SOME ((SOME ((pat, ty), false), tbind), t') haftmann@28145: | NONE => NONE; haftmann@28663: fun implode_monad c_bind_name = Code_Thingol.unfoldr (dest_monad c_bind_name); haftmann@28145: fun pr_monad pr_bind pr (NONE, t) vars = haftmann@28145: (semicolon [pr vars NOBR t], vars) haftmann@31889: | pr_monad pr_bind pr (SOME ((bind, _), true), t) vars = vars haftmann@28145: |> pr_bind NOBR bind haftmann@28145: |>> (fn p => semicolon [p, str "<-", pr vars NOBR t]) haftmann@31889: | pr_monad pr_bind pr (SOME ((bind, _), false), t) vars = vars haftmann@28145: |> pr_bind NOBR bind haftmann@28145: |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]); haftmann@31054: fun pretty _ [c_bind'] pr thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2 haftmann@28145: of SOME (bind, t') => let haftmann@31054: val (binds, t'') = implode_monad c_bind' t' haftmann@31889: val (ps, vars') = fold_map (pr_monad (gen_pr_bind (K pr) thm) pr) (bind :: binds) vars; haftmann@28145: in (brackify fxy o single o Pretty.enclose "do {" "}" o Pretty.breaks) (ps @| pr vars' NOBR t'') end haftmann@28145: | NONE => brackify_infix (1, L) fxy haftmann@28145: [pr vars (INFX (1, L)) t1, str ">>=", pr vars (INFX (1, X)) t2] haftmann@31054: in (2, ([c_bind], pretty)) end; haftmann@28054: haftmann@28145: fun add_monad target' raw_c_bind thy = haftmann@28054: let haftmann@31156: val c_bind = Code.read_const thy raw_c_bind; haftmann@28054: in if target = target' then haftmann@28054: thy haftmann@28145: |> Code_Target.add_syntax_const target c_bind haftmann@28663: (SOME (pretty_haskell_monad c_bind)) haftmann@28054: else error "Only Haskell target allows for monad syntax" end; haftmann@28054: haftmann@28054: haftmann@28054: (** Isar setup **) haftmann@28054: haftmann@28054: fun isar_seri_haskell module = haftmann@28054: Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name) haftmann@28054: -- Scan.optional (Args.$$$ "string_classes" >> K true) false haftmann@28054: >> (fn (module_prefix, string_classes) => haftmann@28054: serialize_haskell module_prefix module string_classes)); haftmann@28054: haftmann@28054: val _ = haftmann@28054: OuterSyntax.command "code_monad" "define code syntax for monads" OuterKeyword.thy_decl ( haftmann@28145: OuterParse.term_group -- OuterParse.name >> (fn (raw_bind, target) => haftmann@28145: Toplevel.theory (add_monad target raw_bind)) haftmann@28054: ); haftmann@28054: haftmann@28054: val setup = haftmann@28064: Code_Target.add_target (target, (isar_seri_haskell, literals)) haftmann@28054: #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => haftmann@28054: brackify_infix (1, R) fxy [ haftmann@28054: pr_typ (INFX (1, X)) ty1, haftmann@28054: str "->", haftmann@28054: pr_typ (INFX (1, R)) ty2 haftmann@28054: ])) haftmann@28054: #> fold (Code_Target.add_reserved target) [ haftmann@28054: "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", haftmann@28054: "import", "default", "forall", "let", "in", "class", "qualified", "data", haftmann@28054: "newtype", "instance", "if", "then", "else", "type", "as", "do", "module" haftmann@28054: ] haftmann@28054: #> fold (Code_Target.add_reserved target) [ haftmann@28054: "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int", haftmann@28054: "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded", haftmann@28054: "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor", haftmann@28054: "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException", haftmann@28054: "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException", haftmann@28054: "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure", haftmann@28054: "ExitSuccess", "False", "GT", "HeapOverflow", haftmann@28054: "IOError", "IOException", "IllegalOperation", haftmann@28054: "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError", haftmann@28054: "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow", haftmann@28054: "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError", haftmann@28054: "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow", haftmann@28054: "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow", haftmann@28054: "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all", haftmann@28054: "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan", haftmann@28054: "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen", haftmann@28054: "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break", haftmann@28054: "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const", haftmann@28054: "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod", haftmann@28054: "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem", haftmann@28054: "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo", haftmann@28054: "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip", haftmann@28054: "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational", haftmann@28054: "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble", haftmann@28054: "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj", haftmann@28054: "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head", haftmann@28054: "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha", haftmann@28054: "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite", haftmann@28054: "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'", haftmann@28054: "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log", haftmann@28054: "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum", haftmann@28054: "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem", haftmann@28054: "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo", haftmann@28054: "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred", haftmann@28054: "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn", haftmann@28054: "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat", haftmann@28054: "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile", haftmann@28054: "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn", haftmann@28054: "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational", haftmann@28054: "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse", haftmann@28054: "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence", haftmann@28054: "sequence_", "show", "showChar", "showException", "showField", "showList", haftmann@28054: "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand", haftmann@28054: "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract", haftmann@28054: "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult", haftmann@28054: "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry", haftmann@28054: "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords", haftmann@28054: "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3" haftmann@28054: ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*); haftmann@28054: haftmann@28054: end; (*struct*)