haftmann@28060: (* Title: Tools/code/code_printer.ML haftmann@28060: Author: Florian Haftmann, TU Muenchen haftmann@28060: haftmann@28060: Generic operations for pretty printing of target language code. haftmann@28060: *) haftmann@28060: haftmann@28060: signature CODE_PRINTER = haftmann@28060: sig haftmann@32908: type itype = Code_Thingol.itype haftmann@32908: type iterm = Code_Thingol.iterm haftmann@32908: type const = Code_Thingol.const haftmann@32908: type dict = Code_Thingol.dict haftmann@32908: haftmann@28663: val nerror: thm -> string -> 'a haftmann@28060: haftmann@28060: val @@ : 'a * 'a -> 'a list haftmann@28060: val @| : 'a list * 'a -> 'a list haftmann@28060: val str: string -> Pretty.T haftmann@28060: val concat: Pretty.T list -> Pretty.T haftmann@28060: val brackets: Pretty.T list -> Pretty.T haftmann@28060: val semicolon: Pretty.T list -> Pretty.T haftmann@28060: val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T haftmann@28060: haftmann@30648: val first_upper: string -> string haftmann@30648: val first_lower: string -> string haftmann@30648: type var_ctxt haftmann@30648: val make_vars: string list -> var_ctxt haftmann@30648: val intro_vars: string list -> var_ctxt -> var_ctxt haftmann@30648: val lookup_var: var_ctxt -> string -> string haftmann@32913: val intro_base_names: (string -> bool) -> (string -> string) haftmann@32913: -> string list -> var_ctxt -> var_ctxt haftmann@32908: val aux_params: var_ctxt -> iterm list list -> string list haftmann@30648: haftmann@31056: type literals haftmann@31056: val Literals: { literal_char: string -> string, literal_string: string -> string, haftmann@31056: literal_numeral: bool -> int -> string, haftmann@31056: literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string } haftmann@31056: -> literals haftmann@31056: val literal_char: literals -> string -> string haftmann@31056: val literal_string: literals -> string -> string haftmann@31056: val literal_numeral: literals -> bool -> int -> string haftmann@31056: val literal_list: literals -> Pretty.T list -> Pretty.T haftmann@31056: val infix_cons: literals -> int * string haftmann@31056: haftmann@28060: type lrx haftmann@28060: val L: lrx haftmann@28060: val R: lrx haftmann@28060: val X: lrx haftmann@28060: type fixity haftmann@28060: val BR: fixity haftmann@28060: val NOBR: fixity haftmann@28060: val INFX: int * lrx -> fixity haftmann@28060: val APP: fixity haftmann@28060: val brackify: fixity -> Pretty.T list -> Pretty.T haftmann@28060: val brackify_infix: int * lrx -> fixity -> Pretty.T list -> Pretty.T haftmann@31665: val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T haftmann@28060: haftmann@28060: type tyco_syntax haftmann@28060: type const_syntax haftmann@31056: type proto_const_syntax haftmann@28060: val parse_infix: ('a -> 'b) -> lrx * int -> string haftmann@28663: -> int * ((fixity -> 'b -> Pretty.T) haftmann@28663: -> fixity -> 'a list -> Pretty.T) haftmann@28060: val parse_syntax: ('a -> 'b) -> OuterParse.token list haftmann@28663: -> (int * ((fixity -> 'b -> Pretty.T) haftmann@28663: -> fixity -> 'a list -> Pretty.T)) option * OuterParse.token list haftmann@28663: val simple_const_syntax: (int * ((fixity -> iterm -> Pretty.T) haftmann@31056: -> fixity -> (iterm * itype) list -> Pretty.T)) option -> proto_const_syntax option haftmann@31056: val activate_const_syntax: theory -> literals haftmann@31056: -> proto_const_syntax -> Code_Thingol.naming -> const_syntax * Code_Thingol.naming haftmann@30648: val gen_pr_app: (thm -> var_ctxt -> const * iterm list -> Pretty.T list) haftmann@30648: -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@31056: -> (string -> const_syntax option) haftmann@30648: -> thm -> var_ctxt -> fixity -> const * iterm list -> Pretty.T haftmann@31889: val gen_pr_bind: (thm -> var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@28663: -> thm -> fixity haftmann@31889: -> iterm -> var_ctxt -> Pretty.T * var_ctxt haftmann@28060: haftmann@30648: val mk_name_module: Name.context -> string option -> (string -> string option) haftmann@30648: -> 'a Graph.T -> string -> string haftmann@30648: val dest_name: string -> string * string haftmann@28060: end; haftmann@28060: haftmann@28060: structure Code_Printer : CODE_PRINTER = haftmann@28060: struct haftmann@28060: haftmann@28060: open Code_Thingol; haftmann@28060: wenzelm@32091: fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm_without_context thm); haftmann@28060: haftmann@28060: (** assembling text pieces **) haftmann@28060: haftmann@28060: infixr 5 @@; haftmann@28060: infixr 5 @|; haftmann@28060: fun x @@ y = [x, y]; haftmann@28060: fun xs @| y = xs @ [y]; haftmann@28060: val str = PrintMode.setmp [] Pretty.str; haftmann@28060: val concat = Pretty.block o Pretty.breaks; haftmann@28060: val brackets = Pretty.enclose "(" ")" o Pretty.breaks; haftmann@28060: fun semicolon ps = Pretty.block [concat ps, str ";"]; haftmann@28060: fun enum_default default sep opn cls [] = str default haftmann@28060: | enum_default default sep opn cls xs = Pretty.enum sep opn cls xs; haftmann@28060: haftmann@28060: haftmann@30648: (** names and variable name contexts **) haftmann@30648: haftmann@30648: type var_ctxt = string Symtab.table * Name.context; haftmann@30648: haftmann@30648: fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty, haftmann@30648: Name.make_context names); haftmann@30648: haftmann@30648: fun intro_vars names (namemap, namectxt) = haftmann@30648: let haftmann@30648: val (names', namectxt') = Name.variants names namectxt; haftmann@30648: val namemap' = fold2 (curry Symtab.update) names names' namemap; haftmann@30648: in (namemap', namectxt') end; haftmann@30648: haftmann@30648: fun lookup_var (namemap, _) name = case Symtab.lookup namemap name haftmann@30648: of SOME name' => name' haftmann@30648: | NONE => error ("Invalid name in context: " ^ quote name); haftmann@30648: haftmann@30648: val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode; haftmann@30648: val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode; haftmann@30648: haftmann@32908: fun aux_params vars lhss = haftmann@32908: let haftmann@32908: fun fish_param _ (w as SOME _) = w haftmann@32908: | fish_param (IVar (SOME v)) NONE = SOME v haftmann@32908: | fish_param _ NONE = NONE; haftmann@32908: fun fillup_param _ (_, SOME v) = v haftmann@32908: | fillup_param x (i, NONE) = x ^ string_of_int i; haftmann@32908: val fished1 = fold (map2 fish_param) lhss (replicate (length (hd lhss)) NONE); haftmann@32908: val x = Name.variant (map_filter I fished1) "x"; haftmann@32908: val fished2 = map_index (fillup_param x) fished1; haftmann@32908: val (fished3, _) = Name.variants fished2 Name.context; haftmann@32908: val vars' = intro_vars fished3 vars; haftmann@32908: in map (lookup_var vars') fished3 end; haftmann@32908: haftmann@32913: fun intro_base_names no_syntax deresolve names = names haftmann@32913: |> map_filter (fn name => if no_syntax name then haftmann@32913: let val name' = deresolve name in haftmann@32913: if Long_Name.is_qualified name' then NONE else SOME name' haftmann@32913: end else NONE) haftmann@32913: |> intro_vars; haftmann@32913: haftmann@30648: haftmann@31056: (** pretty literals **) haftmann@31056: haftmann@31056: datatype literals = Literals of { haftmann@31056: literal_char: string -> string, haftmann@31056: literal_string: string -> string, haftmann@31056: literal_numeral: bool -> int -> string, haftmann@31056: literal_list: Pretty.T list -> Pretty.T, haftmann@31056: infix_cons: int * string haftmann@31056: }; haftmann@31056: haftmann@31056: fun dest_Literals (Literals lits) = lits; haftmann@31056: haftmann@31056: val literal_char = #literal_char o dest_Literals; haftmann@31056: val literal_string = #literal_string o dest_Literals; haftmann@31056: val literal_numeral = #literal_numeral o dest_Literals; haftmann@31056: val literal_list = #literal_list o dest_Literals; haftmann@31056: val infix_cons = #infix_cons o dest_Literals; haftmann@31056: haftmann@31056: haftmann@28060: (** syntax printer **) haftmann@28060: haftmann@28060: (* binding priorities *) haftmann@28060: haftmann@28060: datatype lrx = L | R | X; haftmann@28060: haftmann@28060: datatype fixity = haftmann@28060: BR haftmann@28060: | NOBR haftmann@28060: | INFX of (int * lrx); haftmann@28060: haftmann@28060: val APP = INFX (~1, L); haftmann@28060: haftmann@28060: fun fixity_lrx L L = false haftmann@28060: | fixity_lrx R R = false haftmann@28060: | fixity_lrx _ _ = true; haftmann@28060: haftmann@28060: fun fixity NOBR _ = false haftmann@28060: | fixity _ NOBR = false haftmann@28060: | fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = haftmann@28060: pr < pr_ctxt haftmann@28060: orelse pr = pr_ctxt haftmann@28060: andalso fixity_lrx lr lr_ctxt haftmann@28060: orelse pr_ctxt = ~1 haftmann@28060: | fixity BR (INFX _) = false haftmann@28060: | fixity _ _ = true; haftmann@28060: haftmann@28060: fun gen_brackify _ [p] = p haftmann@28060: | gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps haftmann@28060: | gen_brackify false (ps as _::_) = Pretty.block ps; haftmann@28060: haftmann@28060: fun brackify fxy_ctxt = haftmann@28060: gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks; haftmann@28060: haftmann@28060: fun brackify_infix infx fxy_ctxt = haftmann@28060: gen_brackify (fixity (INFX infx) fxy_ctxt) o Pretty.breaks; haftmann@28060: haftmann@31665: fun brackify_block fxy_ctxt p1 ps p2 = haftmann@31665: let val p = Pretty.block_enclose (p1, p2) ps haftmann@31665: in if fixity BR fxy_ctxt haftmann@31665: then Pretty.enclose "(" ")" [p] haftmann@31665: else p haftmann@31665: end; haftmann@31665: haftmann@28060: haftmann@28060: (* generic syntax *) haftmann@28060: haftmann@28060: type tyco_syntax = int * ((fixity -> itype -> Pretty.T) haftmann@28060: -> fixity -> itype list -> Pretty.T); haftmann@30648: type const_syntax = int * ((var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@31056: -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T); haftmann@31056: type proto_const_syntax = int * (string list * (literals -> string list haftmann@31056: -> (var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@31056: -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T)); haftmann@28060: haftmann@31056: fun simple_const_syntax (SOME (n, f)) = SOME (n, haftmann@31056: ([], (fn _ => fn _ => fn pr => fn thm => fn vars => f (pr vars)))) haftmann@31056: | simple_const_syntax NONE = NONE; haftmann@28060: haftmann@31056: fun activate_const_syntax thy literals (n, (cs, f)) naming = haftmann@31056: fold_map (Code_Thingol.ensure_declared_const thy) cs naming haftmann@31056: |-> (fn cs' => pair (n, f literals cs')); haftmann@31056: haftmann@31056: fun gen_pr_app pr_app pr_term syntax_const thm vars fxy (app as ((c, (_, tys)), ts)) = haftmann@28060: case syntax_const c haftmann@28708: of NONE => brackify fxy (pr_app thm vars app) haftmann@28060: | SOME (k, pr) => haftmann@28060: let haftmann@33955: fun pr' fxy ts = pr (pr_term thm) thm vars fxy (ts ~~ curry (uncurry take) k tys); haftmann@28060: in if k = length ts haftmann@28060: then pr' fxy ts haftmann@28060: else if k < length ts haftmann@28060: then case chop k ts of (ts1, ts2) => haftmann@28708: brackify fxy (pr' APP ts1 :: map (pr_term thm vars BR) ts2) haftmann@28708: else pr_term thm vars fxy (Code_Thingol.eta_expand k app) haftmann@28060: end; haftmann@28060: haftmann@31889: fun gen_pr_bind pr_term thm (fxy : fixity) pat vars = haftmann@28060: let haftmann@31889: val vs = Code_Thingol.fold_varnames (insert (op =)) pat []; haftmann@31874: val vars' = intro_vars vs vars; haftmann@31889: in (pr_term thm vars' fxy pat, vars') end; haftmann@28060: haftmann@28060: haftmann@28060: (* mixfix syntax *) haftmann@28060: haftmann@28060: datatype 'a mixfix = haftmann@28060: Arg of fixity haftmann@28060: | Pretty of Pretty.T; haftmann@28060: haftmann@28060: fun mk_mixfix prep_arg (fixity_this, mfx) = haftmann@28060: let haftmann@28060: fun is_arg (Arg _) = true haftmann@28060: | is_arg _ = false; haftmann@28060: val i = (length o filter is_arg) mfx; haftmann@28060: fun fillin _ [] [] = haftmann@28060: [] haftmann@28060: | fillin pr (Arg fxy :: mfx) (a :: args) = haftmann@28060: (pr fxy o prep_arg) a :: fillin pr mfx args haftmann@28060: | fillin pr (Pretty p :: mfx) args = haftmann@28060: p :: fillin pr mfx args; haftmann@28060: in haftmann@28060: (i, fn pr => fn fixity_ctxt => fn args => haftmann@28060: gen_brackify (fixity fixity_this fixity_ctxt) (fillin pr mfx args)) haftmann@28060: end; haftmann@28060: haftmann@28060: fun parse_infix prep_arg (x, i) s = haftmann@28060: let haftmann@28060: val l = case x of L => INFX (i, L) | _ => INFX (i, X); haftmann@28060: val r = case x of R => INFX (i, R) | _ => INFX (i, X); haftmann@28060: in haftmann@28060: mk_mixfix prep_arg (INFX (i, x), haftmann@28060: [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r]) haftmann@28060: end; haftmann@28060: haftmann@28060: fun parse_mixfix prep_arg s = haftmann@28060: let haftmann@28060: val sym_any = Scan.one Symbol.is_regular; haftmann@28060: val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat ( haftmann@28060: ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR)) haftmann@28060: || ($$ "_" >> K (Arg BR)) haftmann@28060: || ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length)) haftmann@28060: || (Scan.repeat1 haftmann@28060: ( $$ "'" |-- sym_any haftmann@28060: || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")") haftmann@28060: sym_any) >> (Pretty o str o implode))); haftmann@28060: in case Scan.finite Symbol.stopper parse (Symbol.explode s) haftmann@28060: of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p) haftmann@28060: | ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p) haftmann@28060: | _ => Scan.!! haftmann@28060: (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail () haftmann@28060: end; haftmann@28060: haftmann@28060: val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr"); haftmann@28060: haftmann@28060: fun parse_syntax prep_arg xs = haftmann@28060: Scan.option (( haftmann@28060: ((OuterParse.$$$ infixK >> K X) haftmann@28060: || (OuterParse.$$$ infixlK >> K L) haftmann@28060: || (OuterParse.$$$ infixrK >> K R)) haftmann@28060: -- OuterParse.nat >> parse_infix prep_arg haftmann@28060: || Scan.succeed (parse_mixfix prep_arg)) haftmann@28060: -- OuterParse.string haftmann@28060: >> (fn (parse, s) => parse s)) xs; haftmann@28060: haftmann@28060: val _ = List.app OuterKeyword.keyword [infixK, infixlK, infixrK]; haftmann@28060: haftmann@28064: haftmann@30648: (** module name spaces **) haftmann@30648: haftmann@30648: val dest_name = haftmann@30648: apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode; haftmann@30648: haftmann@32924: fun mk_name_module reserved module_prefix module_alias program = haftmann@30648: let haftmann@30648: fun mk_alias name = case module_alias name haftmann@30648: of SOME name' => name' haftmann@30648: | NONE => name haftmann@30648: |> Long_Name.explode haftmann@32924: |> map (fn name => (the_single o fst) (Name.variants [name] reserved)) haftmann@30648: |> Long_Name.implode; haftmann@30648: fun mk_prefix name = case module_prefix haftmann@30648: of SOME module_prefix => Long_Name.append module_prefix name haftmann@30648: | NONE => name; haftmann@30648: val tab = haftmann@30648: Symtab.empty haftmann@30648: |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name)) haftmann@30648: o fst o dest_name o fst) haftmann@30648: program haftmann@30648: in the o Symtab.lookup tab end; haftmann@30648: haftmann@28060: end; (*struct*)