| author | haftmann | 
| Tue, 11 Dec 2007 10:23:13 +0100 | |
| changeset 25606 | 23d34f86b88f | 
| parent 25407 | 2859cf34aaf0 | 
| child 26385 | ae7564661e76 | 
| permissions | -rw-r--r-- | 
| 22124 | 1 | (* Title: Pure/Thy/thy_output.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | Theory document output. | |
| 6 | *) | |
| 7 | ||
| 8 | signature THY_OUTPUT = | |
| 9 | sig | |
| 10 | val display: bool ref | |
| 11 | val quotes: bool ref | |
| 12 | val indent: int ref | |
| 13 | val source: bool ref | |
| 14 | val add_commands: (string * (Args.src -> Toplevel.node option -> string)) list -> unit | |
| 15 | val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit | |
| 16 | val print_antiquotations: unit -> unit | |
| 17 | val boolean: string -> bool | |
| 18 | val integer: string -> int | |
| 19 | val args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> | |
| 20 | (Args.src -> Proof.context -> 'a -> string) -> Args.src -> Toplevel.node option -> string | |
| 21 | datatype markup = Markup | MarkupEnv | Verbatim | |
| 22 | val modes: string list ref | |
| 23 | val eval_antiquote: Scan.lexicon -> Toplevel.node option -> string * Position.T -> string | |
| 23863 | 24 | val process_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) -> | 
| 22124 | 25 | Toplevel.transition list -> (OuterLex.token, 'a) Source.source -> Buffer.T | 
| 26 | val output_list: (Proof.context -> 'a -> Pretty.T) -> Args.src -> | |
| 27 | Proof.context -> 'a list -> string | |
| 28 | val output: (Proof.context -> 'a -> Pretty.T) -> Args.src -> Proof.context -> 'a -> string | |
| 29 | end; | |
| 30 | ||
| 31 | structure ThyOutput: THY_OUTPUT = | |
| 32 | struct | |
| 33 | ||
| 34 | structure T = OuterLex; | |
| 35 | structure P = OuterParse; | |
| 36 | ||
| 37 | ||
| 38 | (** global options **) | |
| 39 | ||
| 40 | val locale = ref ""; | |
| 41 | val display = ref false; | |
| 42 | val quotes = ref false; | |
| 43 | val indent = ref 0; | |
| 44 | val source = ref false; | |
| 45 | val break = ref false; | |
| 46 | ||
| 47 | ||
| 48 | ||
| 49 | (** maintain global commands **) | |
| 50 | ||
| 51 | local | |
| 52 | ||
| 53 | val global_commands = | |
| 54 | ref (Symtab.empty: (Args.src -> Toplevel.node option -> string) Symtab.table); | |
| 55 | ||
| 56 | val global_options = | |
| 57 | ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table); | |
| 58 | ||
| 59 | fun add_item kind (name, x) tab = | |
| 60 | (if not (Symtab.defined tab name) then () | |
| 61 |   else warning ("Redefined text antiquotation " ^ kind ^ ": " ^ quote name);
 | |
| 62 | Symtab.update (name, x) tab); | |
| 63 | ||
| 64 | in | |
| 65 | ||
| 23942 | 66 | fun add_commands xs = CRITICAL (fn () => change global_commands (fold (add_item "command") xs)); | 
| 67 | fun add_options xs = CRITICAL (fn () => change global_options (fold (add_item "option") xs)); | |
| 22124 | 68 | |
| 69 | fun command src = | |
| 70 | let val ((name, _), pos) = Args.dest_src src in | |
| 71 | (case Symtab.lookup (! global_commands) name of | |
| 72 |       NONE => error ("Unknown text antiquotation command: " ^ quote name ^ Position.str_of pos)
 | |
| 23942 | 73 | | SOME f => f src) | 
| 22124 | 74 | end; | 
| 75 | ||
| 76 | fun option (name, s) f () = | |
| 77 | (case Symtab.lookup (! global_options) name of | |
| 78 |     NONE => error ("Unknown text antiquotation option: " ^ quote name)
 | |
| 79 | | SOME opt => opt s f ()); | |
| 80 | ||
| 81 | fun options [] f = f | |
| 82 | | options (opt :: opts) f = option opt (options opts f); | |
| 83 | ||
| 84 | ||
| 85 | fun print_antiquotations () = | |
| 86 | [Pretty.big_list "text antiquotation commands:" | |
| 23272 | 87 | (map Pretty.str (sort_strings (Symtab.keys (! global_commands)))), | 
| 88 | Pretty.big_list "text antiquotation options:" | |
| 89 | (map Pretty.str (sort_strings (Symtab.keys (! global_options))))] | |
| 22124 | 90 | |> Pretty.chunks |> Pretty.writeln; | 
| 91 | ||
| 92 | end; | |
| 93 | ||
| 94 | ||
| 95 | ||
| 96 | (** syntax of antiquotations **) | |
| 97 | ||
| 98 | (* option values *) | |
| 99 | ||
| 100 | fun boolean "" = true | |
| 101 | | boolean "true" = true | |
| 102 | | boolean "false" = false | |
| 103 |   | boolean s = error ("Bad boolean value: " ^ quote s);
 | |
| 104 | ||
| 105 | fun integer s = | |
| 106 | let | |
| 107 | fun int ss = | |
| 108 | (case Library.read_int ss of (i, []) => i | |
| 109 |       | _ => error ("Bad integer value: " ^ quote s));
 | |
| 110 | in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end; | |
| 111 | ||
| 112 | ||
| 113 | (* args syntax *) | |
| 114 | ||
| 115 | fun syntax scan = Args.context_syntax "text antiquotation" scan; | |
| 116 | ||
| 117 | fun args scan f src node : string = | |
| 118 | let | |
| 119 | val loc = if ! locale = "" then NONE else SOME (! locale); | |
| 120 | val (x, ctxt) = syntax scan src (Toplevel.presentation_context node loc); | |
| 121 | in f src ctxt x end; | |
| 122 | ||
| 123 | ||
| 124 | (* outer syntax *) | |
| 125 | ||
| 126 | local | |
| 127 | ||
| 128 | val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) ""; | |
| 129 | val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) []; | |
| 130 | ||
| 131 | in | |
| 132 | ||
| 133 | val antiq = P.!!! (P.position P.xname -- properties -- P.arguments --| Scan.ahead P.eof) | |
| 134 | >> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos))); | |
| 135 | ||
| 136 | end; | |
| 137 | ||
| 138 | ||
| 139 | (* eval_antiquote *) | |
| 140 | ||
| 141 | val modes = ref ([]: string list); | |
| 142 | ||
| 143 | fun eval_antiquote lex node (str, pos) = | |
| 144 | let | |
| 145 | fun expand (Antiquote.Text s) = s | |
| 146 | | expand (Antiquote.Antiq x) = | |
| 147 | let val (opts, src) = Antiquote.scan_arguments lex antiq x in | |
| 148 | options opts (fn () => command src node) (); (*preview errors!*) | |
| 23935 | 149 | PrintMode.with_modes (! modes @ Latex.modes) | 
| 22124 | 150 | (Output.no_warnings (options opts (fn () => command src node))) () | 
| 151 | end; | |
| 152 | val ants = Antiquote.scan_antiquotes (str, pos); | |
| 153 | in | |
| 154 | if is_none node andalso exists Antiquote.is_antiq ants then | |
| 155 |       error ("Cannot expand text antiquotations at top-level" ^ Position.str_of pos)
 | |
| 156 | else implode (map expand ants) | |
| 157 | end; | |
| 158 | ||
| 159 | ||
| 160 | ||
| 161 | (** present theory source **) | |
| 162 | ||
| 163 | (*NB: arranging white space around command spans is a black art.*) | |
| 164 | ||
| 165 | (* presentation tokens *) | |
| 166 | ||
| 167 | datatype token = | |
| 168 | NoToken | |
| 169 | | BasicToken of T.token | |
| 170 | | MarkupToken of string * (string * Position.T) | |
| 171 | | MarkupEnvToken of string * (string * Position.T) | |
| 172 | | VerbatimToken of string * Position.T; | |
| 173 | ||
| 174 | fun output_token lex state = | |
| 175 | let | |
| 176 | val eval = eval_antiquote lex (try Toplevel.node_of state) | |
| 177 | in | |
| 178 | fn NoToken => "" | |
| 179 | | BasicToken tok => Latex.output_basic tok | |
| 180 | | MarkupToken (cmd, txt) => Latex.output_markup cmd (eval txt) | |
| 181 | | MarkupEnvToken (cmd, txt) => Latex.output_markup_env cmd (eval txt) | |
| 182 | | VerbatimToken txt => Latex.output_verbatim (eval txt) | |
| 183 | end; | |
| 184 | ||
| 185 | fun basic_token pred (BasicToken tok) = pred tok | |
| 186 | | basic_token _ _ = false; | |
| 187 | ||
| 188 | val improper_token = basic_token (not o T.is_proper); | |
| 189 | val comment_token = basic_token T.is_comment; | |
| 190 | val blank_token = basic_token T.is_blank; | |
| 191 | val newline_token = basic_token T.is_newline; | |
| 192 | ||
| 193 | ||
| 194 | (* command spans *) | |
| 195 | ||
| 196 | type command = string * Position.T * string list; (*name, position, tags*) | |
| 197 | type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*) | |
| 198 | ||
| 199 | datatype span = Span of command * (source * source * source * source) * bool; | |
| 200 | ||
| 201 | fun make_span cmd src = | |
| 202 | let | |
| 203 | fun take_newline (tok :: toks) = | |
| 204 | if newline_token (fst tok) then ([tok], toks, true) | |
| 205 | else ([], tok :: toks, false) | |
| 206 | | take_newline [] = ([], [], false); | |
| 207 | val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) = | |
| 208 | src | |
| 209 | |> take_prefix (improper_token o fst) | |
| 210 | ||>> take_suffix (improper_token o fst) | |
| 211 | ||>> take_prefix (comment_token o fst) | |
| 212 | ||> take_newline; | |
| 213 | in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end; | |
| 214 | ||
| 215 | ||
| 216 | (* present spans *) | |
| 217 | ||
| 218 | local | |
| 219 | ||
| 220 | fun err_bad_nesting pos = | |
| 221 |   error ("Bad nesting of commands in presentation" ^ pos);
 | |
| 222 | ||
| 223 | fun edge which f (x: string option, y) = | |
| 224 | if x = y then I | |
| 225 | else (case which (x, y) of NONE => I | SOME txt => Buffer.add (f txt)); | |
| 226 | ||
| 227 | val begin_tag = edge #2 Latex.begin_tag; | |
| 228 | val end_tag = edge #1 Latex.end_tag; | |
| 229 | fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e; | |
| 230 | fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e; | |
| 231 | ||
| 232 | in | |
| 233 | ||
| 234 | fun present_span lex default_tags span state state' | |
| 235 | (tag_stack, active_tag, newline, buffer, present_cont) = | |
| 236 | let | |
| 237 | val present = fold (fn (tok, (flag, 0)) => | |
| 238 | Buffer.add (output_token lex state' tok) | |
| 239 | #> Buffer.add flag | |
| 240 | | _ => I); | |
| 241 | ||
| 242 | val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span; | |
| 243 | ||
| 244 | val (tag, tags) = tag_stack; | |
| 245 | val tag' = try hd (fold OuterKeyword.update_tags cmd_tags (the_list tag)); | |
| 246 | ||
| 247 | val active_tag' = | |
| 248 | if is_some tag' then tag' | |
| 249 | else if cmd_name = "end" andalso not (Toplevel.is_toplevel state') then NONE | |
| 250 | else try hd (default_tags cmd_name); | |
| 251 | val edge = (active_tag, active_tag'); | |
| 252 | ||
| 253 | val newline' = | |
| 254 | if is_none active_tag' then span_newline else newline; | |
| 255 | ||
| 256 | val nesting = Toplevel.level state' - Toplevel.level state; | |
| 257 | val tag_stack' = | |
| 258 | if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack | |
| 259 | else if nesting >= 0 then (tag', replicate nesting tag @ tags) | |
| 260 | else | |
| 261 | (case Library.drop (~ nesting - 1, tags) of | |
| 262 | tgs :: tgss => (tgs, tgss) | |
| 263 | | [] => err_bad_nesting (Position.str_of cmd_pos)); | |
| 264 | ||
| 265 | val buffer' = | |
| 266 | buffer | |
| 267 | |> end_tag edge | |
| 268 | |> close_delim (fst present_cont) edge | |
| 269 | |> snd present_cont | |
| 270 | |> open_delim (present (#1 srcs)) edge | |
| 271 | |> begin_tag edge | |
| 272 | |> present (#2 srcs); | |
| 273 | val present_cont' = | |
| 274 | if newline then (present (#3 srcs), present (#4 srcs)) | |
| 275 | else (I, present (#3 srcs) #> present (#4 srcs)); | |
| 276 | in (tag_stack', active_tag', newline', buffer', present_cont') end; | |
| 277 | ||
| 278 | fun present_trailer ((_, tags), active_tag, _, buffer, present_cont) = | |
| 279 | if not (null tags) then err_bad_nesting " at end of theory" | |
| 280 | else | |
| 281 | buffer | |
| 282 | |> end_tag (active_tag, NONE) | |
| 283 | |> close_delim (fst present_cont) (active_tag, NONE) | |
| 284 | |> snd present_cont; | |
| 285 | ||
| 286 | end; | |
| 287 | ||
| 288 | ||
| 23863 | 289 | (* process_thy *) | 
| 22124 | 290 | |
| 291 | datatype markup = Markup | MarkupEnv | Verbatim; | |
| 292 | ||
| 293 | local | |
| 294 | ||
| 295 | val space_proper = | |
| 296 | Scan.one T.is_blank -- Scan.many T.is_comment -- Scan.one T.is_proper; | |
| 297 | ||
| 298 | val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore); | |
| 299 | val improper = Scan.many is_improper; | |
| 300 | val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper)); | |
| 301 | val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one T.is_blank)); | |
| 302 | ||
| 303 | val opt_newline = Scan.option (Scan.one T.is_newline); | |
| 304 | ||
| 305 | val ignore = | |
| 306 | Scan.depend (fn d => opt_newline |-- Scan.one T.is_begin_ignore | |
| 307 | >> pair (d + 1)) || | |
| 308 | Scan.depend (fn d => Scan.one T.is_end_ignore --| | |
| 309 | (if d = 0 then Scan.fail_with (K "Bad nesting of meta-comments") else opt_newline) | |
| 310 | >> pair (d - 1)); | |
| 311 | ||
| 312 | val tag = (improper -- P.$$$ "%" -- improper) |-- P.!!! (P.tag_name --| blank_end); | |
| 313 | ||
| 314 | val locale = | |
| 315 |   Scan.option ((P.$$$ "(" -- improper -- P.$$$ "in") |--
 | |
| 316 | P.!!! (improper |-- P.xname --| (improper -- P.$$$ ")"))); | |
| 317 | ||
| 318 | in | |
| 319 | ||
| 23863 | 320 | fun process_thy lex default_tags is_markup trs src = | 
| 22124 | 321 | let | 
| 322 | (* tokens *) | |
| 323 | ||
| 324 | val ignored = Scan.state --| ignore | |
| 325 |       >> (fn d => (NONE, (NoToken, ("", d))));
 | |
| 326 | ||
| 327 | fun markup mark mk flag = Scan.peek (fn d => | |
| 328 | improper |-- P.position (Scan.one (T.is_kind T.Command andf is_markup mark o T.val_of)) -- | |
| 329 | Scan.repeat tag -- | |
| 330 | P.!!!! ((improper -- locale -- improper) |-- P.position P.text --| improper_end) | |
| 331 | >> (fn (((tok, pos), tags), txt) => | |
| 332 | let val name = T.val_of tok | |
| 333 | in (SOME (name, pos, tags), (mk (name, txt), (flag, d))) end)); | |
| 334 | ||
| 335 | val command = Scan.peek (fn d => | |
| 336 | P.position (Scan.one (T.is_kind T.Command)) -- | |
| 337 | Scan.repeat tag | |
| 338 | >> (fn ((tok, pos), tags) => | |
| 339 | let val name = T.val_of tok | |
| 340 | in (SOME (name, pos, tags), (BasicToken tok, (Latex.markup_false, d))) end)); | |
| 341 | ||
| 342 | val cmt = Scan.peek (fn d => | |
| 343 | P.$$$ "--" |-- P.!!!! (improper |-- P.position P.text) | |
| 344 |       >> (fn txt => (NONE, (MarkupToken ("cmt", txt), ("", d)))));
 | |
| 345 | ||
| 346 | val other = Scan.peek (fn d => | |
| 23725 | 347 |        P.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d)))));
 | 
| 22124 | 348 | |
| 349 | val token = | |
| 350 | ignored || | |
| 351 | markup Markup MarkupToken Latex.markup_true || | |
| 352 | markup MarkupEnv MarkupEnvToken Latex.markup_true || | |
| 353 | markup Verbatim (VerbatimToken o #2) "" || | |
| 354 | command || cmt || other; | |
| 355 | ||
| 356 | ||
| 357 | (* spans *) | |
| 358 | ||
| 359 | val stopper = | |
| 360 |       ((NONE, (BasicToken (#1 T.stopper), ("", 0))),
 | |
| 361 | fn (_, (BasicToken x, _)) => #2 T.stopper x | _ => false); | |
| 362 | ||
| 363 | val cmd = Scan.one (is_some o fst); | |
| 364 | val non_cmd = Scan.one (is_none o fst andf not o #2 stopper) >> #2; | |
| 365 | ||
| 366 | val comments = Scan.many (comment_token o fst o snd); | |
| 367 | val blank = Scan.one (blank_token o fst o snd); | |
| 368 | val newline = Scan.one (newline_token o fst o snd); | |
| 369 | val before_cmd = | |
| 370 | Scan.option (newline -- comments) -- | |
| 371 | Scan.option (newline -- comments) -- | |
| 372 | Scan.option (blank -- comments) -- cmd; | |
| 373 | ||
| 374 | val span = | |
| 375 | Scan.repeat non_cmd -- cmd -- | |
| 376 | Scan.repeat (Scan.unless before_cmd non_cmd) -- | |
| 377 | Scan.option (newline >> (single o snd)) | |
| 378 | >> (fn (((toks1, (cmd, tok2)), toks3), tok4) => | |
| 379 | make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4)))); | |
| 380 | ||
| 381 | val spans = | |
| 382 | src | |
| 383 | |> Source.filter (not o T.is_semicolon) | |
| 384 | |> Source.source' 0 T.stopper (Scan.error (Scan.bulk token)) NONE | |
| 385 | |> Source.source stopper (Scan.error (Scan.bulk span)) NONE | |
| 386 | |> Source.exhaust; | |
| 387 | in | |
| 388 | if length trs = length spans then | |
| 389 | ((NONE, []), NONE, true, Buffer.empty, (I, I)) | |
| 390 | |> Toplevel.present_excursion (trs ~~ map (present_span lex default_tags) spans) | |
| 391 | |> present_trailer | |
| 392 | else error "Messed-up outer syntax for presentation" | |
| 393 | end; | |
| 394 | ||
| 395 | end; | |
| 396 | ||
| 397 | ||
| 398 | ||
| 399 | (** setup default output **) | |
| 400 | ||
| 401 | (* options *) | |
| 402 | ||
| 403 | val _ = add_options | |
| 404 |  [("show_types", Library.setmp Syntax.show_types o boolean),
 | |
| 405 |   ("show_sorts", Library.setmp Syntax.show_sorts o boolean),
 | |
| 406 |   ("show_structs", Library.setmp show_structs o boolean),
 | |
| 407 |   ("show_question_marks", Library.setmp show_question_marks o boolean),
 | |
| 408 |   ("long_names", Library.setmp NameSpace.long_names o boolean),
 | |
| 409 |   ("short_names", Library.setmp NameSpace.short_names o boolean),
 | |
| 410 |   ("unique_names", Library.setmp NameSpace.unique_names o boolean),
 | |
| 411 |   ("eta_contract", Library.setmp Syntax.eta_contract o boolean),
 | |
| 412 |   ("locale", Library.setmp locale),
 | |
| 413 |   ("display", Library.setmp display o boolean),
 | |
| 414 |   ("break", Library.setmp break o boolean),
 | |
| 415 |   ("quotes", Library.setmp quotes o boolean),
 | |
| 23935 | 416 |   ("mode", fn s => fn f => PrintMode.with_modes [s] f),
 | 
| 22124 | 417 |   ("margin", Pretty.setmp_margin o integer),
 | 
| 418 |   ("indent", Library.setmp indent o integer),
 | |
| 419 |   ("source", Library.setmp source o boolean),
 | |
| 420 |   ("goals_limit", Library.setmp goals_limit o integer)];
 | |
| 421 | ||
| 422 | ||
| 423 | (* basic pretty printing *) | |
| 424 | ||
| 425 | val str_of_source = space_implode " " o map Args.string_of o #2 o #1 o Args.dest_src; | |
| 426 | ||
| 427 | fun tweak_line s = | |
| 428 | if ! display then s else Symbol.strip_blanks s; | |
| 429 | ||
| 430 | val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines; | |
| 431 | ||
| 24920 | 432 | fun pretty_term ctxt = Syntax.pretty_term ctxt o ProofContext.revert_skolems ctxt; | 
| 433 | ||
| 22124 | 434 | fun pretty_term_abbrev ctxt = | 
| 435 | ProofContext.pretty_term_abbrev ctxt o ProofContext.revert_skolems ctxt; | |
| 436 | ||
| 437 | fun pretty_term_typ ctxt t = | |
| 24920 | 438 | Syntax.pretty_term ctxt (TypeInfer.constrain (Term.fastype_of t) t); | 
| 22124 | 439 | |
| 24920 | 440 | fun pretty_term_typeof ctxt = Syntax.pretty_typ ctxt o Term.fastype_of; | 
| 22124 | 441 | |
| 25373 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 442 | fun pretty_const ctxt c = | 
| 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 443 | let | 
| 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 444 | val t = Const (c, Consts.type_scheme (ProofContext.consts_of ctxt) c) | 
| 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 445 | handle TYPE (msg, _, _) => error msg; | 
| 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 446 | val ([t'], _) = Variable.import_terms true [t] ctxt; | 
| 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 447 | in pretty_term ctxt t' end; | 
| 22124 | 448 | |
| 25407 
2859cf34aaf0
abbrev: bypass full term check via ProofContext.standard_infer_types (prevents forced expansion);
 wenzelm parents: 
25373diff
changeset | 449 | fun pretty_abbrev ctxt s = | 
| 22124 | 450 | let | 
| 25407 
2859cf34aaf0
abbrev: bypass full term check via ProofContext.standard_infer_types (prevents forced expansion);
 wenzelm parents: 
25373diff
changeset | 451 | val t = Syntax.parse_term ctxt s |> singleton (ProofContext.standard_infer_types ctxt); | 
| 24920 | 452 |     fun err () = error ("Abbreviated constant expected: " ^ Syntax.string_of_term ctxt t);
 | 
| 22124 | 453 | val (head, args) = Term.strip_comb t; | 
| 454 | val (c, T) = Term.dest_Const head handle TERM _ => err (); | |
| 25054 | 455 | val (U, u) = Consts.the_abbreviation (ProofContext.consts_of ctxt) c | 
| 22124 | 456 | handle TYPE _ => err (); | 
| 457 | val t' = Term.betapplys (Envir.expand_atom T (U, u), args); | |
| 458 | in pretty_term_abbrev ctxt (Logic.mk_equals (t, t')) end; | |
| 459 | ||
| 460 | fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; | |
| 461 | ||
| 462 | fun pretty_term_style ctxt (name, t) = | |
| 463 | pretty_term ctxt (TermStyle.the_style (ProofContext.theory_of ctxt) name ctxt t); | |
| 464 | ||
| 465 | fun pretty_thm_style ctxt (name, th) = | |
| 466 | pretty_term_style ctxt (name, Thm.full_prop_of th); | |
| 467 | ||
| 468 | fun pretty_prf full ctxt thms = | |
| 469 | Pretty.chunks (map (ProofContext.pretty_proof_of ctxt full) thms); | |
| 470 | ||
| 471 | fun pretty_theory ctxt name = | |
| 472 | (Theory.requires (ProofContext.theory_of ctxt) name "presentation"; Pretty.str name); | |
| 473 | ||
| 474 | ||
| 475 | (* Isar output *) | |
| 476 | ||
| 477 | fun output_list pretty src ctxt xs = | |
| 478 | map (pretty ctxt) xs (*always pretty in order to exhibit errors!*) | |
| 479 | |> (if ! source then K [pretty_text (str_of_source src)] else I) | |
| 480 | |> (if ! quotes then map Pretty.quote else I) | |
| 481 | |> (if ! display then | |
| 482 | map (Output.output o Pretty.string_of o Pretty.indent (! indent)) | |
| 483 | #> space_implode "\\isasep\\isanewline%\n" | |
| 484 |     #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
 | |
| 485 | else | |
| 486 | map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of)) | |
| 487 | #> space_implode "\\isasep\\isanewline%\n" | |
| 488 |     #> enclose "\\isa{" "}");
 | |
| 489 | ||
| 490 | fun output pretty src ctxt = output_list pretty src ctxt o single; | |
| 491 | ||
| 492 | fun proof_state node = | |
| 493 | (case Option.map Toplevel.proof_node node of | |
| 494 | SOME (SOME prf) => ProofHistory.current prf | |
| 495 | | _ => error "No proof state"); | |
| 496 | ||
| 497 | fun output_goals main_goal src node = args (Scan.succeed ()) (output (fn _ => fn _ => | |
| 498 | Pretty.chunks (Proof.pretty_goals main_goal (proof_state node)))) src node; | |
| 499 | ||
| 500 | fun ml_val txt = "fn _ => (" ^ txt ^ ");";
 | |
| 501 | fun ml_type txt = "val _ = NONE : (" ^ txt ^ ") option;";
 | |
| 502 | fun ml_struct txt = "functor DUMMY_FUNCTOR() = struct structure DUMMY = " ^ txt ^ " end;" | |
| 503 | ||
| 504 | fun output_ml ml src ctxt txt = | |
| 505 | (ML_Context.use_mltext (ml txt) false (SOME (Context.Proof ctxt)); | |
| 506 | (if ! source then str_of_source src else txt) | |
| 507 | |> (if ! quotes then quote else I) | |
| 508 |   |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
 | |
| 509 | else | |
| 510 | split_lines | |
| 511 | #> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|") | |
| 512 | #> space_implode "\\isasep\\isanewline%\n")); | |
| 513 | ||
| 514 | ||
| 515 | (* commands *) | |
| 516 | ||
| 517 | val _ = add_commands | |
| 518 |  [("thm", args Attrib.thms (output_list pretty_thm)),
 | |
| 519 |   ("thm_style", args (Scan.lift Args.liberal_name -- Attrib.thm) (output pretty_thm_style)),
 | |
| 520 |   ("prop", args Args.prop (output pretty_term)),
 | |
| 521 |   ("term", args Args.term (output pretty_term)),
 | |
| 522 |   ("term_style", args (Scan.lift Args.liberal_name -- Args.term) (output pretty_term_style)),
 | |
| 523 |   ("term_type", args Args.term (output pretty_term_typ)),
 | |
| 524 |   ("typeof", args Args.term (output pretty_term_typeof)),
 | |
| 25373 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 wenzelm parents: 
25241diff
changeset | 525 |   ("const", args Args.const_proper (output pretty_const)),
 | 
| 25407 
2859cf34aaf0
abbrev: bypass full term check via ProofContext.standard_infer_types (prevents forced expansion);
 wenzelm parents: 
25373diff
changeset | 526 |   ("abbrev", args (Scan.lift Args.name) (output pretty_abbrev)),
 | 
| 24920 | 527 |   ("typ", args Args.typ_abbrev (output Syntax.pretty_typ)),
 | 
| 22124 | 528 |   ("text", args (Scan.lift Args.name) (output (K pretty_text))),
 | 
| 529 |   ("goals", output_goals true),
 | |
| 530 |   ("subgoals", output_goals false),
 | |
| 531 |   ("prf", args Attrib.thms (output (pretty_prf false))),
 | |
| 532 |   ("full_prf", args Attrib.thms (output (pretty_prf true))),
 | |
| 533 |   ("theory", args (Scan.lift Args.name) (output pretty_theory)),
 | |
| 534 |   ("ML", args (Scan.lift Args.name) (output_ml ml_val)),
 | |
| 535 |   ("ML_type", args (Scan.lift Args.name) (output_ml ml_type)),
 | |
| 536 |   ("ML_struct", args (Scan.lift Args.name) (output_ml ml_struct))];
 | |
| 537 | ||
| 538 | end; |