haftmann@37744: (* 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@35228: val eqn_error: thm option -> 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@34178: val enclose: string -> string -> Pretty.T list -> Pretty.T haftmann@38778: val commas: Pretty.T list -> Pretty.T list haftmann@34178: val enum: string -> string -> string -> Pretty.T list -> Pretty.T haftmann@34178: val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T haftmann@28060: val semicolon: Pretty.T list -> Pretty.T haftmann@33989: val doublesemicolon: Pretty.T list -> Pretty.T haftmann@34178: val indent: int -> Pretty.T -> Pretty.T haftmann@39056: val markup_stmt: string -> Pretty.T -> Pretty.T haftmann@39056: val format: string list -> int -> Pretty.T -> string 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@37958: literal_numeral: 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@34944: val literal_numeral: literals -> 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@37242: val brackify_infix: int * lrx -> fixity -> Pretty.T * Pretty.T * Pretty.T -> Pretty.T haftmann@31665: val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T haftmann@50626: val gen_applify: bool -> string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T haftmann@37638: val applify: string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T haftmann@38922: val tuplify: (fixity -> 'a -> Pretty.T) -> fixity -> 'a list -> Pretty.T option haftmann@38922: haftmann@52070: datatype ('a, 'b, 'c, 'd, 'e) symbol_attr = haftmann@52070: Constant of 'a | Type_Constructor of 'b | Type_Class of 'c | Class_Instance of 'd | Module of 'e haftmann@52070: val map_symbol_attr: ('a -> 'f) -> ('b -> 'g) -> ('c -> 'h) -> ('d -> 'i) -> ('e -> 'j) haftmann@52070: -> ('a, 'b, 'c, 'd, 'e) symbol_attr -> ('f, 'g, 'h, 'i, 'j) symbol_attr haftmann@52070: val maps_symbol_attr: ('a -> 'f list) -> ('b -> 'g list) haftmann@52070: -> ('c -> 'h list) -> ('d -> 'i list) -> ('e -> 'j list) haftmann@52070: -> ('a, 'b, 'c, 'd, 'e) symbol_attr -> ('f, 'g, 'h, 'i, 'j) symbol_attr list haftmann@52070: val maps_symbol_attr': ('a -> ('k * 'f) list) -> ('b -> ('k * 'g) list) haftmann@52070: -> ('c -> ('k * 'h) list) -> ('d -> ('k * 'i) list) -> ('e -> ('k * 'j) list) haftmann@52070: -> ('a, 'b, 'c, 'd, 'e) symbol_attr -> ('k * ('f, 'g, 'h, 'i, 'j) symbol_attr) list haftmann@52070: type ('a, 'b, 'c, 'd, 'e) symbol_data haftmann@52070: val empty_symbol_data: ('a, 'b, 'c, 'd, 'e) symbol_data haftmann@52070: val merge_symbol_data: ('a, 'b, 'c, 'd, 'e) symbol_data * ('a, 'b, 'c, 'd, 'e) symbol_data haftmann@52070: -> ('a, 'b, 'c, 'd, 'e) symbol_data haftmann@52070: val lookup_constant_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> string -> 'a option haftmann@52070: val lookup_type_constructor_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> string -> 'b option haftmann@52070: val lookup_type_class_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> string -> 'c option haftmann@52070: val lookup_class_instance_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> string * string -> 'd option haftmann@52070: val lookup_module_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> string -> 'e option haftmann@52070: val dest_constant_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> (string * 'a) list haftmann@52070: val dest_type_constructor_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> (string * 'b) list haftmann@52070: val dest_type_class_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> (string * 'c) list haftmann@52070: val dest_class_instance_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> ((string * string) * 'd) list haftmann@52070: val dest_module_data: ('a, 'b, 'c, 'd, 'e) symbol_data -> (string * 'e) list haftmann@52070: val set_symbol_data: (string * 'a option, string * 'b option, string * 'c option, haftmann@52070: (string * string) * 'd option, string * 'e option) symbol_attr haftmann@52070: -> ('a, 'b, 'c, 'd, 'e) symbol_data -> ('a, 'b, 'c, 'd, 'e) symbol_data haftmann@52070: haftmann@34152: type simple_const_syntax haftmann@37881: type complex_const_syntax haftmann@28060: type const_syntax haftmann@37881: type activated_complex_const_syntax haftmann@37881: datatype activated_const_syntax = Plain_const_syntax of int * string haftmann@37881: | Complex_const_syntax of activated_complex_const_syntax haftmann@52070: type tyco_syntax haftmann@37881: val requires_args: const_syntax -> int wenzelm@52079: val parse_const_syntax: const_syntax parser wenzelm@52079: val parse_tyco_syntax: tyco_syntax parser haftmann@37881: val plain_const_syntax: string -> const_syntax haftmann@37876: val simple_const_syntax: simple_const_syntax -> const_syntax haftmann@37881: val complex_const_syntax: complex_const_syntax -> const_syntax haftmann@31056: val activate_const_syntax: theory -> literals haftmann@37881: -> string -> const_syntax -> Code_Thingol.naming -> activated_const_syntax * Code_Thingol.naming haftmann@35228: val gen_print_app: (thm option -> var_ctxt -> const * iterm list -> Pretty.T list) haftmann@35228: -> (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@37876: -> (string -> activated_const_syntax option) haftmann@35228: -> thm option -> var_ctxt -> fixity -> const * iterm list -> Pretty.T haftmann@35228: val gen_print_bind: (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@35228: -> thm option -> fixity haftmann@31889: -> iterm -> var_ctxt -> Pretty.T * var_ctxt haftmann@28060: end; haftmann@28060: haftmann@28060: structure Code_Printer : CODE_PRINTER = haftmann@28060: struct haftmann@28060: haftmann@28060: open Code_Thingol; haftmann@28060: haftmann@39034: (** generic nonsense *) haftmann@39034: wenzelm@43947: fun eqn_error (SOME thm) s = wenzelm@43947: error (s ^ ",\nin equation " ^ Display.string_of_thm_without_context thm) haftmann@35228: | eqn_error NONE s = error s; haftmann@28060: haftmann@39034: val code_presentationN = "code_presentation"; haftmann@39057: val stmt_nameN = "stmt_name"; haftmann@39034: val _ = Markup.add_mode code_presentationN YXML.output_markup; haftmann@39034: haftmann@39034: haftmann@34071: (** assembling and printing 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@39660: val str = Print_Mode.setmp [] Pretty.str; haftmann@28060: val concat = Pretty.block o Pretty.breaks; haftmann@39660: val commas = Print_Mode.setmp [] Pretty.commas; haftmann@39660: fun enclose l r = Print_Mode.setmp [] (Pretty.enclose l r); haftmann@34178: val brackets = enclose "(" ")" o Pretty.breaks; haftmann@39660: fun enum sep l r = Print_Mode.setmp [] (Pretty.enum sep l r); haftmann@34178: fun enum_default default sep l r [] = str default haftmann@34178: | enum_default default sep l r xs = enum sep l r xs; haftmann@28060: fun semicolon ps = Pretty.block [concat ps, str ";"]; haftmann@33989: fun doublesemicolon ps = Pretty.block [concat ps, str ";;"]; haftmann@39660: fun indent i = Print_Mode.setmp [] (Pretty.indent i); haftmann@28060: haftmann@39660: fun markup_stmt name = Print_Mode.setmp [code_presentationN] haftmann@39069: (Pretty.mark (code_presentationN, [(stmt_nameN, name)])); haftmann@39062: haftmann@39558: fun filter_presentation [] tree = haftmann@39558: Buffer.empty wenzelm@43947: |> fold XML.add_content tree haftmann@39558: | filter_presentation presentation_names tree = haftmann@39057: let wenzelm@43947: fun is_selected (name, attrs) = haftmann@39558: name = code_presentationN haftmann@39558: andalso member (op =) presentation_names (the (Properties.get attrs stmt_nameN)); haftmann@39558: fun add_content_with_space tree (is_first, buf) = haftmann@39558: buf haftmann@39558: |> not is_first ? Buffer.add "\n\n" haftmann@39558: |> XML.add_content tree haftmann@39558: |> pair false; haftmann@39558: fun filter (XML.Elem (name_attrs, xs)) = haftmann@39558: fold (if is_selected name_attrs then add_content_with_space else filter) xs haftmann@47576: | filter (XML.Text _) = I; haftmann@39558: in snd (fold filter tree (true, Buffer.empty)) end; haftmann@39062: haftmann@39062: fun format presentation_names width = haftmann@39660: Print_Mode.setmp [code_presentationN] (Pretty.string_of_margin width) haftmann@39062: #> YXML.parse_body haftmann@39558: #> filter_presentation presentation_names haftmann@39678: #> Buffer.add "\n" haftmann@39558: #> Buffer.content; haftmann@34071: 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 wenzelm@43326: val (names', namectxt') = fold_map Name.variant names namectxt; haftmann@30648: val namemap' = fold2 (curry Symtab.update) names names' namemap; haftmann@30648: in (namemap', namectxt') end; haftmann@30648: wenzelm@43947: fun lookup_var (namemap, _) name = wenzelm@43947: case Symtab.lookup namemap name of wenzelm@43947: SOME name' => name' haftmann@30648: | NONE => error ("Invalid name in context: " ^ quote name); haftmann@30648: wenzelm@40627: val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o raw_explode; wenzelm@40627: val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o raw_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); wenzelm@43324: val x = singleton (Name.variant_list (map_filter I fished1)) "x"; haftmann@32908: val fished2 = map_index (fillup_param x) fished1; wenzelm@43326: val (fished3, _) = fold_map Name.variant fished2 Name.context; haftmann@32908: val vars' = intro_vars fished3 vars; haftmann@32908: in map (lookup_var vars') fished3 end; haftmann@32908: wenzelm@43947: 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@34944: literal_numeral: 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@33989: | fixity (INFX (pr, lr)) (INFX (print_ctxt, lr_ctxt)) = haftmann@33989: pr < print_ctxt haftmann@33989: orelse pr = print_ctxt haftmann@28060: andalso fixity_lrx lr lr_ctxt haftmann@33989: orelse print_ctxt = ~1 haftmann@28060: | fixity BR (INFX _) = false haftmann@28060: | fixity _ _ = true; haftmann@28060: haftmann@28060: fun gen_brackify _ [p] = p haftmann@34178: | gen_brackify true (ps as _::_) = 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@37242: fun brackify_infix infx fxy_ctxt (l, m, r) = haftmann@50627: gen_brackify (fixity (INFX infx) fxy_ctxt) [l, str " ", m, Pretty.brk 1, r]; 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@34178: then enclose "(" ")" [p] haftmann@31665: else p haftmann@31665: end; haftmann@31665: haftmann@50626: fun gen_applify strict opn cls f fxy_ctxt p [] = haftmann@50626: if strict haftmann@50628: then gen_brackify (fixity BR fxy_ctxt) [p, str (opn ^ cls)] haftmann@50626: else p haftmann@50626: | gen_applify strict opn cls f fxy_ctxt p ps = haftmann@50627: gen_brackify (fixity BR fxy_ctxt) (p @@ enum "," opn cls (map f ps)); haftmann@34247: haftmann@50626: fun applify opn = gen_applify false opn; haftmann@50626: haftmann@38922: fun tuplify _ _ [] = NONE haftmann@38922: | tuplify print fxy [x] = SOME (print fxy x) haftmann@38922: | tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs)); haftmann@38922: haftmann@28060: haftmann@52070: (* data for symbols *) haftmann@52070: haftmann@52070: datatype ('a, 'b, 'c, 'd, 'e) symbol_attr = haftmann@52070: Constant of 'a | Type_Constructor of 'b | Type_Class of 'c | Class_Instance of 'd | Module of 'e; haftmann@52070: haftmann@52070: fun map_symbol_attr const tyco class inst module (Constant x) = Constant (const x) haftmann@52070: | map_symbol_attr const tyco class inst module (Type_Constructor x) = Type_Constructor (tyco x) haftmann@52070: | map_symbol_attr const tyco class inst module (Type_Class x) = Type_Class (class x) haftmann@52070: | map_symbol_attr const tyco class inst module (Class_Instance x) = Class_Instance (inst x) haftmann@52070: | map_symbol_attr const tyco class inst module (Module x) = Module (module x); haftmann@52070: haftmann@52070: fun maps_symbol_attr const tyco class inst module (Constant x) = map Constant (const x) haftmann@52070: | maps_symbol_attr const tyco class inst module (Type_Constructor x) = map Type_Constructor (tyco x) haftmann@52070: | maps_symbol_attr const tyco class inst module (Type_Class x) = map Type_Class (class x) haftmann@52070: | maps_symbol_attr const tyco class inst module (Class_Instance x) = map Class_Instance (inst x) haftmann@52070: | maps_symbol_attr const tyco class inst module (Module x) = map Module (module x); haftmann@52070: haftmann@52070: fun maps_symbol_attr' const tyco class inst module (Constant x) = (map o apsnd) Constant (const x) haftmann@52070: | maps_symbol_attr' const tyco class inst module (Type_Constructor x) = (map o apsnd) Type_Constructor (tyco x) haftmann@52070: | maps_symbol_attr' const tyco class inst module (Type_Class x) = (map o apsnd) Type_Class (class x) haftmann@52070: | maps_symbol_attr' const tyco class inst module (Class_Instance x) = (map o apsnd) Class_Instance (inst x) haftmann@52070: | maps_symbol_attr' const tyco class inst module (Module x) = (map o apsnd) Module (module x); haftmann@52070: haftmann@52070: datatype ('a, 'b, 'c, 'd, 'e) symbol_data = haftmann@52070: Symbol_Data of { constant: 'a Symtab.table, type_constructor: 'b Symtab.table, haftmann@52070: type_class: 'c Symtab.table, class_instance: 'd Symreltab.table, module: 'e Symtab.table }; haftmann@52070: haftmann@52070: fun make_symbol_data constant type_constructor type_class class_instance module = haftmann@52070: Symbol_Data { constant = constant, type_constructor = type_constructor, haftmann@52070: type_class = type_class, class_instance = class_instance, module = module }; haftmann@52070: fun dest_symbol_data (Symbol_Data x) = x; haftmann@52070: fun map_symbol_data map_constant map_type_constructor map_type_class map_class_instance map_module haftmann@52070: (Symbol_Data { constant, type_constructor, type_class, class_instance, module }) = haftmann@52070: Symbol_Data { constant = map_constant constant, type_constructor = map_type_constructor type_constructor, haftmann@52070: type_class = map_type_class type_class, class_instance = map_class_instance class_instance, haftmann@52070: module = map_module module }; haftmann@28060: haftmann@52070: val empty_symbol_data = Symbol_Data { constant = Symtab.empty, type_constructor = Symtab.empty, haftmann@52070: type_class = Symtab.empty, class_instance = Symreltab.empty, module = Symtab.empty }; haftmann@52070: fun merge_symbol_data (Symbol_Data { constant = constant1, type_constructor = type_constructor1, haftmann@52070: type_class = type_class1, class_instance = class_instance1, module = module1 }, haftmann@52070: Symbol_Data { constant = constant2, type_constructor = type_constructor2, haftmann@52070: type_class = type_class2, class_instance = class_instance2, module = module2 }) = haftmann@52070: make_symbol_data (Symtab.join (K snd) (constant1, constant2)) haftmann@52070: (Symtab.join (K snd) (type_constructor1, type_constructor2)) haftmann@52070: (Symtab.join (K snd) (type_class1, type_class2)) haftmann@52070: (Symreltab.join (K snd) (class_instance1, class_instance2)) haftmann@52070: (Symtab.join (K snd) (module1, module2)); (*prefer later entries: K snd*) haftmann@52070: haftmann@52070: fun lookup_constant_data x = (Symtab.lookup o #constant o dest_symbol_data) x; haftmann@52070: fun lookup_type_constructor_data x = (Symtab.lookup o #type_constructor o dest_symbol_data) x; haftmann@52070: fun lookup_type_class_data x = (Symtab.lookup o #type_class o dest_symbol_data) x; haftmann@52070: fun lookup_class_instance_data x = (Symreltab.lookup o #class_instance o dest_symbol_data) x; haftmann@52070: fun lookup_module_data x = (Symtab.lookup o #module o dest_symbol_data) x; haftmann@52070: haftmann@52070: fun dest_constant_data x = (Symtab.dest o #constant o dest_symbol_data) x; haftmann@52070: fun dest_type_constructor_data x = (Symtab.dest o #type_constructor o dest_symbol_data) x; haftmann@52070: fun dest_type_class_data x = (Symtab.dest o #type_class o dest_symbol_data) x; haftmann@52070: fun dest_class_instance_data x = (Symreltab.dest o #class_instance o dest_symbol_data) x; haftmann@52070: fun dest_module_data x = (Symtab.dest o #module o dest_symbol_data) x; haftmann@52070: haftmann@52070: fun set_sym (sym, NONE) = Symtab.delete_safe sym haftmann@52070: | set_sym (sym, SOME y) = Symtab.update (sym, y); haftmann@52070: fun set_symrel (symrel, NONE) = Symreltab.delete_safe symrel haftmann@52070: | set_symrel (symrel, SOME y) = Symreltab.update (symrel, y); haftmann@52070: haftmann@52070: fun set_symbol_data (Constant x) = map_symbol_data (set_sym x) I I I I haftmann@52070: | set_symbol_data (Type_Constructor x) = map_symbol_data I (set_sym x) I I I haftmann@52070: | set_symbol_data (Type_Class x) = map_symbol_data I I (set_sym x) I I haftmann@52070: | set_symbol_data (Class_Instance x) = map_symbol_data I I I (set_symrel x) I haftmann@52070: | set_symbol_data (Module x) = map_symbol_data I I I I (set_sym x); haftmann@52070: haftmann@52070: haftmann@52070: (* generic syntax *) haftmann@34152: haftmann@34152: type simple_const_syntax = int * ((fixity -> iterm -> Pretty.T) haftmann@34152: -> fixity -> (iterm * itype) list -> Pretty.T); haftmann@37881: haftmann@37881: type complex_const_syntax = int * (string list * (literals -> string list haftmann@31056: -> (var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@35228: -> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T)); haftmann@37881: haftmann@37881: datatype const_syntax = plain_const_syntax of string haftmann@37881: | complex_const_syntax of complex_const_syntax; haftmann@37881: haftmann@37881: fun requires_args (plain_const_syntax _) = 0 haftmann@37881: | requires_args (complex_const_syntax (k, _)) = k; haftmann@28060: haftmann@34176: fun simple_const_syntax syn = wenzelm@43947: complex_const_syntax wenzelm@43947: (apsnd (fn f => ([], (fn _ => fn _ => fn print => fn _ => fn vars => f (print vars)))) syn); haftmann@28060: haftmann@37881: type activated_complex_const_syntax = int * ((var_ctxt -> fixity -> iterm -> Pretty.T) haftmann@37881: -> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T) haftmann@37881: haftmann@37881: datatype activated_const_syntax = Plain_const_syntax of int * string haftmann@37881: | Complex_const_syntax of activated_complex_const_syntax; haftmann@31056: haftmann@37881: fun activate_const_syntax thy literals c (plain_const_syntax s) naming = haftmann@37881: (Plain_const_syntax (Code.args_number thy c, s), naming) haftmann@37881: | activate_const_syntax thy literals c (complex_const_syntax (n, (cs, f))) naming = haftmann@37881: fold_map (Code_Thingol.ensure_declared_const thy) cs naming haftmann@37881: |-> (fn cs' => pair (Complex_const_syntax (n, f literals cs'))); haftmann@37881: wenzelm@43947: fun gen_print_app print_app_expr print_term const_syntax some_thm vars fxy haftmann@48072: (app as ({ name = c, dom, ... }, ts)) = wenzelm@43947: case const_syntax c of wenzelm@43947: NONE => brackify fxy (print_app_expr some_thm vars app) wenzelm@43947: | SOME (Plain_const_syntax (_, s)) => wenzelm@43947: brackify fxy (str s :: map (print_term some_thm vars BR) ts) wenzelm@43947: | SOME (Complex_const_syntax (k, print)) => wenzelm@43947: let wenzelm@43947: fun print' fxy ts = haftmann@48072: print (print_term some_thm) some_thm vars fxy (ts ~~ take k dom); wenzelm@43947: in wenzelm@43947: if k = length ts wenzelm@43947: then print' fxy ts haftmann@28060: else if k < length ts wenzelm@43947: then case chop k ts of (ts1, ts2) => wenzelm@43947: brackify fxy (print' APP ts1 :: map (print_term some_thm vars BR) ts2) wenzelm@43947: else print_term some_thm vars fxy (Code_Thingol.eta_expand k app) wenzelm@43947: end; haftmann@28060: haftmann@33989: fun gen_print_bind print_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@33989: in (print_term thm vars' fxy pat, vars') end; haftmann@28060: haftmann@52070: type tyco_syntax = int * ((fixity -> itype -> Pretty.T) haftmann@52070: -> fixity -> itype list -> Pretty.T); haftmann@52070: haftmann@28060: haftmann@28060: (* mixfix syntax *) haftmann@28060: haftmann@28060: datatype 'a mixfix = haftmann@28060: Arg of fixity haftmann@37881: | String of string haftmann@37881: | Break; haftmann@28060: haftmann@52069: fun printer_of_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@33989: | fillin print (Arg fxy :: mfx) (a :: args) = haftmann@33989: (print fxy o prep_arg) a :: fillin print mfx args haftmann@37881: | fillin print (String s :: mfx) args = haftmann@37881: str s :: fillin print mfx args haftmann@37881: | fillin print (Break :: mfx) args = haftmann@37881: Pretty.brk 1 :: fillin print mfx args; haftmann@28060: in haftmann@33989: (i, fn print => fn fixity_ctxt => fn args => haftmann@33989: gen_brackify (fixity fixity_this fixity_ctxt) (fillin print mfx args)) haftmann@28060: end; haftmann@28060: haftmann@52069: fun read_infix (fixity_this, i) s = haftmann@28060: let haftmann@52069: val l = case fixity_this of L => INFX (i, L) | _ => INFX (i, X); haftmann@52069: val r = case fixity_this of R => INFX (i, R) | _ => INFX (i, X); haftmann@28060: in haftmann@52069: (INFX (i, fixity_this), [Arg l, String " ", String s, Break, Arg r]) haftmann@28060: end; haftmann@28060: haftmann@52069: fun read_mixfix s = haftmann@28060: let haftmann@28060: val sym_any = Scan.one Symbol.is_regular; haftmann@52069: val parse = Scan.optional ($$ "!" >> K NOBR) BR -- Scan.repeat ( haftmann@28060: ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR)) haftmann@28060: || ($$ "_" >> K (Arg BR)) haftmann@37881: || ($$ "/" |-- Scan.repeat ($$ " ") >> (K Break)) haftmann@28060: || (Scan.repeat1 haftmann@28060: ( $$ "'" |-- sym_any haftmann@28060: || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")") haftmann@37881: sym_any) >> (String o implode))); wenzelm@43947: fun err s (_, NONE) = (fn () => "malformed mixfix annotation: " ^ quote s) wenzelm@43947: | err _ (_, SOME msg) = msg; wenzelm@43947: in wenzelm@43947: case Scan.finite Symbol.stopper parse (Symbol.explode s) of haftmann@52069: (fixity_mixfix, []) => fixity_mixfix haftmann@52069: | _ => Scan.!! (err s) Scan.fail () haftmann@28060: end; haftmann@28060: haftmann@52069: val parse_fixity = haftmann@52069: (@{keyword "infix"} >> K X) || (@{keyword "infixl"} >> K L) || (@{keyword "infixr"} >> K R) haftmann@52069: haftmann@52069: val parse_mixfix = haftmann@52069: Parse.string >> read_mixfix haftmann@52069: || parse_fixity -- Parse.nat -- Parse.string haftmann@52069: >> (fn ((fixity, i), s) => read_infix (fixity, i) s); haftmann@28060: haftmann@52069: fun syntax_of_mixfix of_plain of_printer prep_arg (BR, [String s]) = of_plain s haftmann@52069: | syntax_of_mixfix of_plain of_printer prep_arg (fixity, mfx) = haftmann@52069: of_printer (printer_of_mixfix prep_arg (fixity, mfx)); haftmann@37881: haftmann@52069: val parse_tyco_syntax = haftmann@52069: parse_mixfix >> syntax_of_mixfix (fn s => (0, (K o K o K o str) s)) I I; haftmann@52069: haftmann@52069: val parse_const_syntax = haftmann@52069: parse_mixfix >> syntax_of_mixfix plain_const_syntax simple_const_syntax fst; haftmann@37881: haftmann@28060: end; (*struct*)