wenzelm@22124: (* Title: Pure/Thy/thy_output.ML wenzelm@22124: ID: $Id$ wenzelm@22124: Author: Markus Wenzel, TU Muenchen wenzelm@22124: wenzelm@22124: Theory document output. wenzelm@22124: *) wenzelm@22124: wenzelm@22124: signature THY_OUTPUT = wenzelm@22124: sig wenzelm@22124: val display: bool ref wenzelm@22124: val quotes: bool ref wenzelm@22124: val indent: int ref wenzelm@22124: val source: bool ref wenzelm@22124: val add_commands: (string * (Args.src -> Toplevel.node option -> string)) list -> unit wenzelm@22124: val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit wenzelm@26893: val defined_command: string -> bool wenzelm@26893: val defined_option: string -> bool wenzelm@22124: val print_antiquotations: unit -> unit wenzelm@22124: val boolean: string -> bool wenzelm@22124: val integer: string -> int wenzelm@22124: val args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> wenzelm@22124: (Args.src -> Proof.context -> 'a -> string) -> Args.src -> Toplevel.node option -> string wenzelm@22124: datatype markup = Markup | MarkupEnv | Verbatim wenzelm@22124: val modes: string list ref wenzelm@22124: val eval_antiquote: Scan.lexicon -> Toplevel.node option -> string * Position.T -> string wenzelm@23863: val process_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) -> wenzelm@22124: Toplevel.transition list -> (OuterLex.token, 'a) Source.source -> Buffer.T wenzelm@22124: val output_list: (Proof.context -> 'a -> Pretty.T) -> Args.src -> wenzelm@22124: Proof.context -> 'a list -> string wenzelm@22124: val output: (Proof.context -> 'a -> Pretty.T) -> Args.src -> Proof.context -> 'a -> string wenzelm@22124: end; wenzelm@22124: wenzelm@22124: structure ThyOutput: THY_OUTPUT = wenzelm@22124: struct wenzelm@22124: wenzelm@22124: structure T = OuterLex; wenzelm@22124: structure P = OuterParse; wenzelm@22124: wenzelm@22124: wenzelm@22124: (** global options **) wenzelm@22124: wenzelm@22124: val locale = ref ""; wenzelm@22124: val display = ref false; wenzelm@22124: val quotes = ref false; wenzelm@22124: val indent = ref 0; wenzelm@22124: val source = ref false; wenzelm@22124: val break = ref false; wenzelm@22124: wenzelm@22124: wenzelm@22124: wenzelm@22124: (** maintain global commands **) wenzelm@22124: wenzelm@22124: local wenzelm@22124: wenzelm@22124: val global_commands = wenzelm@22124: ref (Symtab.empty: (Args.src -> Toplevel.node option -> string) Symtab.table); wenzelm@22124: wenzelm@22124: val global_options = wenzelm@22124: ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table); wenzelm@22124: wenzelm@22124: fun add_item kind (name, x) tab = wenzelm@22124: (if not (Symtab.defined tab name) then () wenzelm@22124: else warning ("Redefined text antiquotation " ^ kind ^ ": " ^ quote name); wenzelm@22124: Symtab.update (name, x) tab); wenzelm@22124: wenzelm@22124: in wenzelm@22124: wenzelm@23942: fun add_commands xs = CRITICAL (fn () => change global_commands (fold (add_item "command") xs)); wenzelm@23942: fun add_options xs = CRITICAL (fn () => change global_options (fold (add_item "option") xs)); wenzelm@22124: wenzelm@26893: fun defined_command name = Symtab.defined (! global_commands) name; wenzelm@26893: fun defined_option name = Symtab.defined (! global_options) name; wenzelm@26893: wenzelm@22124: fun command src = wenzelm@22124: let val ((name, _), pos) = Args.dest_src src in wenzelm@22124: (case Symtab.lookup (! global_commands) name of wenzelm@22124: NONE => error ("Unknown text antiquotation command: " ^ quote name ^ Position.str_of pos) wenzelm@23942: | SOME f => f src) wenzelm@22124: end; wenzelm@22124: wenzelm@22124: fun option (name, s) f () = wenzelm@22124: (case Symtab.lookup (! global_options) name of wenzelm@22124: NONE => error ("Unknown text antiquotation option: " ^ quote name) wenzelm@22124: | SOME opt => opt s f ()); wenzelm@22124: wenzelm@22124: fun options [] f = f wenzelm@22124: | options (opt :: opts) f = option opt (options opts f); wenzelm@22124: wenzelm@22124: wenzelm@22124: fun print_antiquotations () = wenzelm@22124: [Pretty.big_list "text antiquotation commands:" wenzelm@23272: (map Pretty.str (sort_strings (Symtab.keys (! global_commands)))), wenzelm@23272: Pretty.big_list "text antiquotation options:" wenzelm@23272: (map Pretty.str (sort_strings (Symtab.keys (! global_options))))] wenzelm@22124: |> Pretty.chunks |> Pretty.writeln; wenzelm@22124: wenzelm@22124: end; wenzelm@22124: wenzelm@22124: wenzelm@22124: wenzelm@22124: (** syntax of antiquotations **) wenzelm@22124: wenzelm@22124: (* option values *) wenzelm@22124: wenzelm@22124: fun boolean "" = true wenzelm@22124: | boolean "true" = true wenzelm@22124: | boolean "false" = false wenzelm@22124: | boolean s = error ("Bad boolean value: " ^ quote s); wenzelm@22124: wenzelm@22124: fun integer s = wenzelm@22124: let wenzelm@22124: fun int ss = wenzelm@22124: (case Library.read_int ss of (i, []) => i wenzelm@22124: | _ => error ("Bad integer value: " ^ quote s)); wenzelm@22124: in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* args syntax *) wenzelm@22124: wenzelm@22124: fun syntax scan = Args.context_syntax "text antiquotation" scan; wenzelm@22124: wenzelm@22124: fun args scan f src node : string = wenzelm@22124: let wenzelm@22124: val loc = if ! locale = "" then NONE else SOME (! locale); wenzelm@22124: val (x, ctxt) = syntax scan src (Toplevel.presentation_context node loc); wenzelm@22124: in f src ctxt x end; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* outer syntax *) wenzelm@22124: wenzelm@22124: local wenzelm@22124: wenzelm@22124: val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) ""; wenzelm@22124: val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) []; wenzelm@22124: wenzelm@22124: in wenzelm@22124: wenzelm@22124: val antiq = P.!!! (P.position P.xname -- properties -- P.arguments --| Scan.ahead P.eof) wenzelm@22124: >> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos))); wenzelm@22124: wenzelm@22124: end; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* eval_antiquote *) wenzelm@22124: wenzelm@22124: val modes = ref ([]: string list); wenzelm@22124: wenzelm@22124: fun eval_antiquote lex node (str, pos) = wenzelm@22124: let wenzelm@22124: fun expand (Antiquote.Text s) = s wenzelm@22124: | expand (Antiquote.Antiq x) = wenzelm@22124: let val (opts, src) = Antiquote.scan_arguments lex antiq x in wenzelm@22124: options opts (fn () => command src node) (); (*preview errors!*) wenzelm@23935: PrintMode.with_modes (! modes @ Latex.modes) wenzelm@22124: (Output.no_warnings (options opts (fn () => command src node))) () wenzelm@27344: end wenzelm@27344: | expand (Antiquote.Open _) = "" wenzelm@27344: | expand (Antiquote.Close _) = ""; wenzelm@22124: val ants = Antiquote.scan_antiquotes (str, pos); wenzelm@22124: in wenzelm@22124: if is_none node andalso exists Antiquote.is_antiq ants then wenzelm@27344: error ("Unknown context -- cannot expand text antiquotations" ^ Position.str_of pos) wenzelm@22124: else implode (map expand ants) wenzelm@22124: end; wenzelm@22124: wenzelm@22124: wenzelm@22124: wenzelm@22124: (** present theory source **) wenzelm@22124: wenzelm@22124: (*NB: arranging white space around command spans is a black art.*) wenzelm@22124: wenzelm@22124: (* presentation tokens *) wenzelm@22124: wenzelm@22124: datatype token = wenzelm@22124: NoToken wenzelm@22124: | BasicToken of T.token wenzelm@22124: | MarkupToken of string * (string * Position.T) wenzelm@22124: | MarkupEnvToken of string * (string * Position.T) wenzelm@22124: | VerbatimToken of string * Position.T; wenzelm@22124: wenzelm@22124: fun output_token lex state = wenzelm@22124: let wenzelm@22124: val eval = eval_antiquote lex (try Toplevel.node_of state) wenzelm@22124: in wenzelm@22124: fn NoToken => "" wenzelm@22124: | BasicToken tok => Latex.output_basic tok wenzelm@22124: | MarkupToken (cmd, txt) => Latex.output_markup cmd (eval txt) wenzelm@22124: | MarkupEnvToken (cmd, txt) => Latex.output_markup_env cmd (eval txt) wenzelm@22124: | VerbatimToken txt => Latex.output_verbatim (eval txt) wenzelm@22124: end; wenzelm@22124: wenzelm@22124: fun basic_token pred (BasicToken tok) = pred tok wenzelm@22124: | basic_token _ _ = false; wenzelm@22124: wenzelm@22124: val improper_token = basic_token (not o T.is_proper); wenzelm@22124: val comment_token = basic_token T.is_comment; wenzelm@22124: val blank_token = basic_token T.is_blank; wenzelm@22124: val newline_token = basic_token T.is_newline; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* command spans *) wenzelm@22124: wenzelm@22124: type command = string * Position.T * string list; (*name, position, tags*) wenzelm@22124: type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*) wenzelm@22124: wenzelm@22124: datatype span = Span of command * (source * source * source * source) * bool; wenzelm@22124: wenzelm@22124: fun make_span cmd src = wenzelm@22124: let wenzelm@22124: fun take_newline (tok :: toks) = wenzelm@22124: if newline_token (fst tok) then ([tok], toks, true) wenzelm@22124: else ([], tok :: toks, false) wenzelm@22124: | take_newline [] = ([], [], false); wenzelm@22124: val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) = wenzelm@22124: src wenzelm@22124: |> take_prefix (improper_token o fst) wenzelm@22124: ||>> take_suffix (improper_token o fst) wenzelm@22124: ||>> take_prefix (comment_token o fst) wenzelm@22124: ||> take_newline; wenzelm@22124: in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* present spans *) wenzelm@22124: wenzelm@22124: local wenzelm@22124: wenzelm@22124: fun err_bad_nesting pos = wenzelm@22124: error ("Bad nesting of commands in presentation" ^ pos); wenzelm@22124: wenzelm@22124: fun edge which f (x: string option, y) = wenzelm@22124: if x = y then I wenzelm@22124: else (case which (x, y) of NONE => I | SOME txt => Buffer.add (f txt)); wenzelm@22124: wenzelm@22124: val begin_tag = edge #2 Latex.begin_tag; wenzelm@22124: val end_tag = edge #1 Latex.end_tag; wenzelm@22124: fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e; wenzelm@22124: fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e; wenzelm@22124: wenzelm@22124: in wenzelm@22124: wenzelm@22124: fun present_span lex default_tags span state state' wenzelm@22124: (tag_stack, active_tag, newline, buffer, present_cont) = wenzelm@22124: let wenzelm@22124: val present = fold (fn (tok, (flag, 0)) => wenzelm@22124: Buffer.add (output_token lex state' tok) wenzelm@22124: #> Buffer.add flag wenzelm@22124: | _ => I); wenzelm@22124: wenzelm@22124: val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span; wenzelm@22124: wenzelm@22124: val (tag, tags) = tag_stack; wenzelm@22124: val tag' = try hd (fold OuterKeyword.update_tags cmd_tags (the_list tag)); wenzelm@22124: wenzelm@22124: val active_tag' = wenzelm@22124: if is_some tag' then tag' wenzelm@22124: else if cmd_name = "end" andalso not (Toplevel.is_toplevel state') then NONE wenzelm@22124: else try hd (default_tags cmd_name); wenzelm@22124: val edge = (active_tag, active_tag'); wenzelm@22124: wenzelm@22124: val newline' = wenzelm@22124: if is_none active_tag' then span_newline else newline; wenzelm@22124: wenzelm@22124: val nesting = Toplevel.level state' - Toplevel.level state; wenzelm@22124: val tag_stack' = wenzelm@22124: if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack wenzelm@22124: else if nesting >= 0 then (tag', replicate nesting tag @ tags) wenzelm@22124: else wenzelm@22124: (case Library.drop (~ nesting - 1, tags) of wenzelm@22124: tgs :: tgss => (tgs, tgss) wenzelm@22124: | [] => err_bad_nesting (Position.str_of cmd_pos)); wenzelm@22124: wenzelm@22124: val buffer' = wenzelm@22124: buffer wenzelm@22124: |> end_tag edge wenzelm@22124: |> close_delim (fst present_cont) edge wenzelm@22124: |> snd present_cont wenzelm@22124: |> open_delim (present (#1 srcs)) edge wenzelm@22124: |> begin_tag edge wenzelm@22124: |> present (#2 srcs); wenzelm@22124: val present_cont' = wenzelm@22124: if newline then (present (#3 srcs), present (#4 srcs)) wenzelm@22124: else (I, present (#3 srcs) #> present (#4 srcs)); wenzelm@22124: in (tag_stack', active_tag', newline', buffer', present_cont') end; wenzelm@22124: wenzelm@22124: fun present_trailer ((_, tags), active_tag, _, buffer, present_cont) = wenzelm@22124: if not (null tags) then err_bad_nesting " at end of theory" wenzelm@22124: else wenzelm@22124: buffer wenzelm@22124: |> end_tag (active_tag, NONE) wenzelm@22124: |> close_delim (fst present_cont) (active_tag, NONE) wenzelm@22124: |> snd present_cont; wenzelm@22124: wenzelm@22124: end; wenzelm@22124: wenzelm@22124: wenzelm@23863: (* process_thy *) wenzelm@22124: wenzelm@22124: datatype markup = Markup | MarkupEnv | Verbatim; wenzelm@22124: wenzelm@22124: local wenzelm@22124: wenzelm@22124: val space_proper = wenzelm@22124: Scan.one T.is_blank -- Scan.many T.is_comment -- Scan.one T.is_proper; wenzelm@22124: wenzelm@22124: val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore); wenzelm@22124: val improper = Scan.many is_improper; wenzelm@22124: val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper)); wenzelm@22124: val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one T.is_blank)); wenzelm@22124: wenzelm@22124: val opt_newline = Scan.option (Scan.one T.is_newline); wenzelm@22124: wenzelm@22124: val ignore = wenzelm@22124: Scan.depend (fn d => opt_newline |-- Scan.one T.is_begin_ignore wenzelm@22124: >> pair (d + 1)) || wenzelm@22124: Scan.depend (fn d => Scan.one T.is_end_ignore --| wenzelm@22124: (if d = 0 then Scan.fail_with (K "Bad nesting of meta-comments") else opt_newline) wenzelm@22124: >> pair (d - 1)); wenzelm@22124: wenzelm@22124: val tag = (improper -- P.$$$ "%" -- improper) |-- P.!!! (P.tag_name --| blank_end); wenzelm@22124: wenzelm@22124: val locale = wenzelm@22124: Scan.option ((P.$$$ "(" -- improper -- P.$$$ "in") |-- wenzelm@22124: P.!!! (improper |-- P.xname --| (improper -- P.$$$ ")"))); wenzelm@22124: wenzelm@22124: in wenzelm@22124: wenzelm@23863: fun process_thy lex default_tags is_markup trs src = wenzelm@22124: let wenzelm@22124: (* tokens *) wenzelm@22124: wenzelm@22124: val ignored = Scan.state --| ignore wenzelm@22124: >> (fn d => (NONE, (NoToken, ("", d)))); wenzelm@22124: wenzelm@22124: fun markup mark mk flag = Scan.peek (fn d => wenzelm@22124: improper |-- P.position (Scan.one (T.is_kind T.Command andf is_markup mark o T.val_of)) -- wenzelm@22124: Scan.repeat tag -- wenzelm@22124: P.!!!! ((improper -- locale -- improper) |-- P.position P.text --| improper_end) wenzelm@22124: >> (fn (((tok, pos), tags), txt) => wenzelm@22124: let val name = T.val_of tok wenzelm@22124: in (SOME (name, pos, tags), (mk (name, txt), (flag, d))) end)); wenzelm@22124: wenzelm@22124: val command = Scan.peek (fn d => wenzelm@22124: P.position (Scan.one (T.is_kind T.Command)) -- wenzelm@22124: Scan.repeat tag wenzelm@22124: >> (fn ((tok, pos), tags) => wenzelm@22124: let val name = T.val_of tok wenzelm@22124: in (SOME (name, pos, tags), (BasicToken tok, (Latex.markup_false, d))) end)); wenzelm@22124: wenzelm@22124: val cmt = Scan.peek (fn d => wenzelm@22124: P.$$$ "--" |-- P.!!!! (improper |-- P.position P.text) wenzelm@22124: >> (fn txt => (NONE, (MarkupToken ("cmt", txt), ("", d))))); wenzelm@22124: wenzelm@22124: val other = Scan.peek (fn d => wenzelm@23725: P.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d))))); wenzelm@22124: wenzelm@22124: val token = wenzelm@22124: ignored || wenzelm@22124: markup Markup MarkupToken Latex.markup_true || wenzelm@22124: markup MarkupEnv MarkupEnvToken Latex.markup_true || wenzelm@22124: markup Verbatim (VerbatimToken o #2) "" || wenzelm@22124: command || cmt || other; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* spans *) wenzelm@22124: wenzelm@22124: val stopper = wenzelm@22124: ((NONE, (BasicToken (#1 T.stopper), ("", 0))), wenzelm@22124: fn (_, (BasicToken x, _)) => #2 T.stopper x | _ => false); wenzelm@22124: wenzelm@22124: val cmd = Scan.one (is_some o fst); wenzelm@22124: val non_cmd = Scan.one (is_none o fst andf not o #2 stopper) >> #2; wenzelm@22124: wenzelm@22124: val comments = Scan.many (comment_token o fst o snd); wenzelm@22124: val blank = Scan.one (blank_token o fst o snd); wenzelm@22124: val newline = Scan.one (newline_token o fst o snd); wenzelm@22124: val before_cmd = wenzelm@22124: Scan.option (newline -- comments) -- wenzelm@22124: Scan.option (newline -- comments) -- wenzelm@22124: Scan.option (blank -- comments) -- cmd; wenzelm@22124: wenzelm@22124: val span = wenzelm@22124: Scan.repeat non_cmd -- cmd -- wenzelm@22124: Scan.repeat (Scan.unless before_cmd non_cmd) -- wenzelm@22124: Scan.option (newline >> (single o snd)) wenzelm@22124: >> (fn (((toks1, (cmd, tok2)), toks3), tok4) => wenzelm@22124: make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4)))); wenzelm@22124: wenzelm@22124: val spans = wenzelm@22124: src wenzelm@22124: |> Source.filter (not o T.is_semicolon) wenzelm@22124: |> Source.source' 0 T.stopper (Scan.error (Scan.bulk token)) NONE wenzelm@22124: |> Source.source stopper (Scan.error (Scan.bulk span)) NONE wenzelm@22124: |> Source.exhaust; wenzelm@22124: in wenzelm@22124: if length trs = length spans then wenzelm@22124: ((NONE, []), NONE, true, Buffer.empty, (I, I)) wenzelm@22124: |> Toplevel.present_excursion (trs ~~ map (present_span lex default_tags) spans) wenzelm@22124: |> present_trailer wenzelm@22124: else error "Messed-up outer syntax for presentation" wenzelm@22124: end; wenzelm@22124: wenzelm@22124: end; wenzelm@22124: wenzelm@22124: wenzelm@22124: wenzelm@22124: (** setup default output **) wenzelm@22124: wenzelm@22124: (* options *) wenzelm@22124: wenzelm@22124: val _ = add_options wenzelm@22124: [("show_types", Library.setmp Syntax.show_types o boolean), wenzelm@22124: ("show_sorts", Library.setmp Syntax.show_sorts o boolean), wenzelm@22124: ("show_structs", Library.setmp show_structs o boolean), wenzelm@22124: ("show_question_marks", Library.setmp show_question_marks o boolean), wenzelm@22124: ("long_names", Library.setmp NameSpace.long_names o boolean), wenzelm@22124: ("short_names", Library.setmp NameSpace.short_names o boolean), wenzelm@22124: ("unique_names", Library.setmp NameSpace.unique_names o boolean), wenzelm@22124: ("eta_contract", Library.setmp Syntax.eta_contract o boolean), wenzelm@22124: ("locale", Library.setmp locale), wenzelm@22124: ("display", Library.setmp display o boolean), wenzelm@22124: ("break", Library.setmp break o boolean), wenzelm@22124: ("quotes", Library.setmp quotes o boolean), wenzelm@23935: ("mode", fn s => fn f => PrintMode.with_modes [s] f), wenzelm@22124: ("margin", Pretty.setmp_margin o integer), wenzelm@22124: ("indent", Library.setmp indent o integer), wenzelm@22124: ("source", Library.setmp source o boolean), wenzelm@22124: ("goals_limit", Library.setmp goals_limit o integer)]; wenzelm@22124: wenzelm@22124: wenzelm@22124: (* basic pretty printing *) wenzelm@22124: wenzelm@22124: val str_of_source = space_implode " " o map Args.string_of o #2 o #1 o Args.dest_src; wenzelm@22124: wenzelm@22124: fun tweak_line s = wenzelm@22124: if ! display then s else Symbol.strip_blanks s; wenzelm@22124: wenzelm@22124: val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines; wenzelm@22124: wenzelm@26710: fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t; wenzelm@24920: wenzelm@26710: fun pretty_term_abbrev ctxt t = ProofContext.pretty_term_abbrev (Variable.auto_fixes t ctxt) t; wenzelm@22124: wenzelm@22124: fun pretty_term_typ ctxt t = wenzelm@24920: Syntax.pretty_term ctxt (TypeInfer.constrain (Term.fastype_of t) t); wenzelm@22124: wenzelm@24920: fun pretty_term_typeof ctxt = Syntax.pretty_typ ctxt o Term.fastype_of; wenzelm@22124: wenzelm@25373: fun pretty_const ctxt c = wenzelm@25373: let wenzelm@25373: val t = Const (c, Consts.type_scheme (ProofContext.consts_of ctxt) c) wenzelm@25373: handle TYPE (msg, _, _) => error msg; wenzelm@25373: val ([t'], _) = Variable.import_terms true [t] ctxt; wenzelm@25373: in pretty_term ctxt t' end; wenzelm@22124: wenzelm@25407: fun pretty_abbrev ctxt s = wenzelm@22124: let wenzelm@25407: val t = Syntax.parse_term ctxt s |> singleton (ProofContext.standard_infer_types ctxt); wenzelm@24920: fun err () = error ("Abbreviated constant expected: " ^ Syntax.string_of_term ctxt t); wenzelm@22124: val (head, args) = Term.strip_comb t; wenzelm@22124: val (c, T) = Term.dest_Const head handle TERM _ => err (); wenzelm@25054: val (U, u) = Consts.the_abbreviation (ProofContext.consts_of ctxt) c wenzelm@22124: handle TYPE _ => err (); wenzelm@22124: val t' = Term.betapplys (Envir.expand_atom T (U, u), args); wenzelm@22124: in pretty_term_abbrev ctxt (Logic.mk_equals (t, t')) end; wenzelm@22124: wenzelm@22124: fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; wenzelm@22124: wenzelm@22124: fun pretty_term_style ctxt (name, t) = wenzelm@22124: pretty_term ctxt (TermStyle.the_style (ProofContext.theory_of ctxt) name ctxt t); wenzelm@22124: wenzelm@22124: fun pretty_thm_style ctxt (name, th) = wenzelm@22124: pretty_term_style ctxt (name, Thm.full_prop_of th); wenzelm@22124: wenzelm@22124: fun pretty_prf full ctxt thms = wenzelm@27258: Pretty.chunks (map (ProofSyntax.pretty_proof_of ctxt full) thms); wenzelm@22124: wenzelm@22124: fun pretty_theory ctxt name = wenzelm@22124: (Theory.requires (ProofContext.theory_of ctxt) name "presentation"; Pretty.str name); wenzelm@22124: wenzelm@22124: wenzelm@22124: (* Isar output *) wenzelm@22124: wenzelm@22124: fun output_list pretty src ctxt xs = wenzelm@22124: map (pretty ctxt) xs (*always pretty in order to exhibit errors!*) wenzelm@22124: |> (if ! source then K [pretty_text (str_of_source src)] else I) wenzelm@22124: |> (if ! quotes then map Pretty.quote else I) wenzelm@22124: |> (if ! display then wenzelm@22124: map (Output.output o Pretty.string_of o Pretty.indent (! indent)) wenzelm@22124: #> space_implode "\\isasep\\isanewline%\n" wenzelm@22124: #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}" wenzelm@22124: else wenzelm@22124: map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of)) wenzelm@22124: #> space_implode "\\isasep\\isanewline%\n" wenzelm@22124: #> enclose "\\isa{" "}"); wenzelm@22124: wenzelm@22124: fun output pretty src ctxt = output_list pretty src ctxt o single; wenzelm@22124: wenzelm@22124: fun proof_state node = wenzelm@22124: (case Option.map Toplevel.proof_node node of wenzelm@27566: SOME (SOME prf) => ProofNode.current prf wenzelm@22124: | _ => error "No proof state"); wenzelm@22124: wenzelm@22124: fun output_goals main_goal src node = args (Scan.succeed ()) (output (fn _ => fn _ => wenzelm@22124: Pretty.chunks (Proof.pretty_goals main_goal (proof_state node)))) src node; wenzelm@22124: wenzelm@22124: fun ml_val txt = "fn _ => (" ^ txt ^ ");"; wenzelm@22124: fun ml_type txt = "val _ = NONE : (" ^ txt ^ ") option;"; wenzelm@22124: fun ml_struct txt = "functor DUMMY_FUNCTOR() = struct structure DUMMY = " ^ txt ^ " end;" wenzelm@22124: wenzelm@26385: fun output_ml ml src ctxt (txt, pos) = wenzelm@26455: (ML_Context.eval_in (SOME (Context.Proof ctxt)) false pos (ml txt); wenzelm@22124: (if ! source then str_of_source src else txt) wenzelm@22124: |> (if ! quotes then quote else I) wenzelm@22124: |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}" wenzelm@22124: else wenzelm@22124: split_lines wenzelm@22124: #> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|") wenzelm@22124: #> space_implode "\\isasep\\isanewline%\n")); wenzelm@22124: wenzelm@22124: wenzelm@27522: (* embedded lemmas *) wenzelm@27522: wenzelm@27522: fun pretty_lemma ctxt (prop, methods) = wenzelm@27522: let wenzelm@27522: val _ = ctxt wenzelm@27522: |> Proof.theorem_i NONE (K I) [[(prop, [])]] wenzelm@27522: |> Proof.global_terminal_proof methods; wenzelm@27522: in pretty_term ctxt prop end; wenzelm@22124: wenzelm@27381: val embedded_lemma = wenzelm@27522: args (Args.prop -- Scan.lift (Args.$$$ "by" |-- Method.parse_args -- Scan.option Method.parse_args)) wenzelm@27381: (output pretty_lemma o (fn ((a, arg :: _), p) => (Args.src ((a, [arg]), p))) o Args.dest_src); wenzelm@27381: wenzelm@27522: wenzelm@27522: (* commands *) wenzelm@27522: wenzelm@27381: val _ = OuterKeyword.keyword "by"; wenzelm@27381: wenzelm@22124: val _ = add_commands wenzelm@22124: [("thm", args Attrib.thms (output_list pretty_thm)), wenzelm@22124: ("thm_style", args (Scan.lift Args.liberal_name -- Attrib.thm) (output pretty_thm_style)), wenzelm@22124: ("prop", args Args.prop (output pretty_term)), wenzelm@27381: ("lemma", embedded_lemma), wenzelm@22124: ("term", args Args.term (output pretty_term)), wenzelm@22124: ("term_style", args (Scan.lift Args.liberal_name -- Args.term) (output pretty_term_style)), wenzelm@22124: ("term_type", args Args.term (output pretty_term_typ)), wenzelm@22124: ("typeof", args Args.term (output pretty_term_typeof)), wenzelm@25373: ("const", args Args.const_proper (output pretty_const)), wenzelm@25407: ("abbrev", args (Scan.lift Args.name) (output pretty_abbrev)), wenzelm@24920: ("typ", args Args.typ_abbrev (output Syntax.pretty_typ)), wenzelm@22124: ("text", args (Scan.lift Args.name) (output (K pretty_text))), wenzelm@22124: ("goals", output_goals true), wenzelm@22124: ("subgoals", output_goals false), wenzelm@22124: ("prf", args Attrib.thms (output (pretty_prf false))), wenzelm@22124: ("full_prf", args Attrib.thms (output (pretty_prf true))), wenzelm@22124: ("theory", args (Scan.lift Args.name) (output pretty_theory)), wenzelm@26385: ("ML", args (Scan.lift (Args.position Args.name)) (output_ml ml_val)), wenzelm@26385: ("ML_type", args (Scan.lift (Args.position Args.name)) (output_ml ml_type)), wenzelm@26385: ("ML_struct", args (Scan.lift (Args.position Args.name)) (output_ml ml_struct))]; wenzelm@22124: wenzelm@22124: end;