| author | huffman | 
| Mon, 07 Mar 2005 23:54:01 +0100 | |
| changeset 15587 | f363e6e080e7 | 
| parent 15570 | 8d8c70b41bab | 
| child 15666 | 5c5925dc4921 | 
| permissions | -rw-r--r-- | 
| 9140 | 1 | (* Title: Pure/Isar/isar_output.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | Isar theory output. | |
| 6 | *) | |
| 7 | ||
| 8 | signature ISAR_OUTPUT = | |
| 9 | sig | |
| 14899 | 10 | val display: bool ref | 
| 11 | val quotes: bool ref | |
| 12 | val indent: int ref | |
| 13 | val source: bool ref | |
| 10321 | 14 | val add_commands: (string * (Args.src -> Toplevel.state -> string)) list -> unit | 
| 9140 | 15 | val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit | 
| 9220 | 16 | val print_antiquotations: unit -> unit | 
| 9140 | 17 | val boolean: string -> bool | 
| 18 | val integer: string -> int | |
| 19 | val args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | |
| 10321 | 20 | (Args.src -> Proof.context -> 'a -> string) -> Args.src -> Toplevel.state -> string | 
| 9140 | 21 | datatype markup = Markup | MarkupEnv | Verbatim | 
| 22 | val interest_level: int ref | |
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 23 | val hide_commands: bool ref | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 24 | val add_hidden_commands: string list -> unit | 
| 11714 | 25 | val modes: string list ref | 
| 12939 | 26 | val eval_antiquote: Scan.lexicon -> Toplevel.state -> string * Position.T -> string | 
| 9140 | 27 | val parse_thy: (markup -> OuterLex.token list -> OuterLex.token * OuterLex.token list) -> | 
| 28 | Scan.lexicon -> Toplevel.transition list -> (OuterLex.token, 'a) Source.source -> | |
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 29 | (Toplevel.transition * (Toplevel.state -> Toplevel.state -> Buffer.T -> Buffer.T)) list * Buffer.T | 
| 11716 | 30 | val output_with: (Proof.context -> 'a -> Pretty.T) -> Args.src -> | 
| 31 | Proof.context -> 'a -> string | |
| 9140 | 32 | end; | 
| 33 | ||
| 34 | structure IsarOutput: ISAR_OUTPUT = | |
| 35 | struct | |
| 36 | ||
| 37 | structure T = OuterLex; | |
| 38 | structure P = OuterParse; | |
| 39 | ||
| 40 | ||
| 14899 | 41 | |
| 42 | (** global references -- defaults for options **) | |
| 43 | ||
| 44 | val locale = ref ""; | |
| 45 | val display = ref false; | |
| 46 | val quotes = ref false; | |
| 47 | val indent = ref 0; | |
| 48 | val source = ref false; | |
| 49 | val break = ref false; | |
| 50 | ||
| 51 | ||
| 52 | ||
| 9140 | 53 | (** maintain global commands **) | 
| 54 | ||
| 55 | local | |
| 56 | ||
| 57 | val global_commands = | |
| 10321 | 58 | ref (Symtab.empty: (Args.src -> Toplevel.state -> string) Symtab.table); | 
| 9140 | 59 | |
| 60 | val global_options = | |
| 61 | ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table); | |
| 62 | ||
| 63 | ||
| 64 | fun add_item kind (tab, (name, x)) = | |
| 65 | (if is_none (Symtab.lookup (tab, name)) then () | |
| 66 |   else warning ("Redefined antiquotation output " ^ kind ^ ": " ^ quote name);
 | |
| 67 | Symtab.update ((name, x), tab)); | |
| 68 | ||
| 15570 | 69 | fun add_items kind xs tab = Library.foldl (add_item kind) (tab, xs); | 
| 9140 | 70 | |
| 71 | in | |
| 72 | ||
| 73 | val add_commands = Library.change global_commands o (add_items "command"); | |
| 74 | val add_options = Library.change global_options o (add_items "option"); | |
| 75 | ||
| 76 | fun command src = | |
| 77 | let val ((name, _), pos) = Args.dest_src src in | |
| 78 | (case Symtab.lookup (! global_commands, name) of | |
| 15531 | 79 |       NONE => error ("Unknown antiquotation command: " ^ quote name ^ Position.str_of pos)
 | 
| 80 | | SOME f => transform_failure (curry Antiquote.ANTIQUOTE_FAIL (name, pos)) (f src)) | |
| 9140 | 81 | end; | 
| 82 | ||
| 83 | fun option (name, s) f () = | |
| 84 | (case Symtab.lookup (! global_options, name) of | |
| 15531 | 85 |     NONE => error ("Unknown antiquotation option: " ^ quote name)
 | 
| 86 | | SOME opt => opt s f ()); | |
| 9140 | 87 | |
| 88 | fun options [] f = f | |
| 89 | | options (opt :: opts) f = option opt (options opts f); | |
| 90 | ||
| 9213 | 91 | |
| 9220 | 92 | fun print_antiquotations () = | 
| 93 | [Pretty.big_list "antiquotation commands:" (map Pretty.str (Symtab.keys (! global_commands))), | |
| 94 | Pretty.big_list "antiquotation options:" (map Pretty.str (Symtab.keys (! global_options)))] | |
| 95 | |> Pretty.chunks |> Pretty.writeln; | |
| 9213 | 96 | |
| 9140 | 97 | end; | 
| 98 | ||
| 99 | ||
| 100 | ||
| 101 | (** syntax of antiquotations **) | |
| 102 | ||
| 103 | (* option values *) | |
| 104 | ||
| 105 | fun boolean "" = true | |
| 106 | | boolean "true" = true | |
| 107 | | boolean "false" = false | |
| 108 |   | boolean s = error ("Bad boolean value: " ^ quote s);
 | |
| 109 | ||
| 110 | fun integer s = | |
| 111 | let | |
| 112 | fun int ss = | |
| 14899 | 113 | (case Library.read_int ss of (i, []) => i | 
| 114 |       | _ => error ("Bad integer value: " ^ quote s));
 | |
| 9140 | 115 | in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end; | 
| 116 | ||
| 117 | ||
| 118 | (* args syntax *) | |
| 119 | ||
| 120 | fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = | |
| 121 | Args.syntax "antiquotation" scan; | |
| 122 | ||
| 10321 | 123 | fun args scan f src state : string = | 
| 14899 | 124 | let | 
| 125 | val ctxt0 = | |
| 126 | if ! locale = "" then Toplevel.context_of state | |
| 15531 | 127 | else #3 (Locale.read_context_statement (SOME (! locale)) [] [] | 
| 14899 | 128 | (ProofContext.init (Toplevel.theory_of state))); | 
| 129 | val (ctxt, x) = syntax scan src ctxt0; | |
| 10321 | 130 | in f src ctxt x end; | 
| 9140 | 131 | |
| 132 | ||
| 133 | (* outer syntax *) | |
| 134 | ||
| 135 | local | |
| 136 | ||
| 137 | val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) ""; | |
| 138 | val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) []; | |
| 139 | ||
| 140 | val antiq = P.position P.xname -- properties -- P.arguments --| Scan.ahead P.eof | |
| 141 | >> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos))); | |
| 142 | ||
| 143 | fun antiq_args_aux keyword_lexicon (str, pos) = | |
| 144 | Source.of_string str | |
| 145 | |> Symbol.source false | |
| 146 | |> T.source false (K (keyword_lexicon, Scan.empty_lexicon)) pos | |
| 147 | |> T.source_proper | |
| 15531 | 148 | |> Source.source T.stopper (Scan.error (Scan.bulk (P.!!! antiq))) NONE | 
| 9140 | 149 | |> Source.exhaust; | 
| 150 | ||
| 151 | in | |
| 152 | ||
| 153 | fun antiq_args lex (s, pos) = | |
| 154 | (case antiq_args_aux lex (s, pos) of [x] => x | _ => raise ERROR) handle ERROR => | |
| 155 |     error ("Malformed antiquotation: " ^ quote ("@{" ^ s ^ "}") ^ Position.str_of pos);
 | |
| 156 | ||
| 157 | end; | |
| 158 | ||
| 159 | ||
| 160 | (* eval_antiquote *) | |
| 161 | ||
| 11714 | 162 | val modes = ref ([]: string list); | 
| 163 | ||
| 9140 | 164 | fun eval_antiquote lex state (str, pos) = | 
| 165 | let | |
| 166 | fun expand (Antiquote.Text s) = s | |
| 167 | | expand (Antiquote.Antiq x) = | |
| 168 | let val (opts, src) = antiq_args lex x in | |
| 10739 | 169 | options opts (fn () => command src state) (); (*preview errors!*) | 
| 11714 | 170 | Library.setmp print_mode (! modes @ Latex.modes @ ! print_mode) | 
| 10570 | 171 | (options opts (fn () => command src state)) () | 
| 10321 | 172 | end; | 
| 9140 | 173 | val ants = Antiquote.antiquotes_of (str, pos); | 
| 174 | in | |
| 175 | if Toplevel.is_toplevel state andalso exists Antiquote.is_antiq ants then | |
| 176 |       error ("Cannot expand antiquotations at top-level" ^ Position.str_of pos)
 | |
| 177 | else implode (map expand ants) | |
| 178 | end; | |
| 179 | ||
| 180 | ||
| 181 | ||
| 182 | (** present theory source **) | |
| 183 | ||
| 184 | (* present_tokens *) | |
| 185 | ||
| 186 | val interest_level = ref 0; | |
| 187 | ||
| 15473 | 188 | val hide_commands = ref true; | 
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 189 | val hidden_commands = ref ([] : string list); | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 190 | |
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 191 | fun add_hidden_commands cmds = (hidden_commands := !hidden_commands @ cmds); | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 192 | |
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 193 | fun is_proof state = (case Toplevel.node_of state of | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 194 | Toplevel.Theory _ => false | _ => true) handle Toplevel.UNDEF => false; | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 195 | |
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 196 | fun is_newline (Latex.Basic tk, _) = T.is_newline tk | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 197 | | is_newline _ = false; | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 198 | |
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 199 | fun is_hidden (Latex.Basic tk, _) = | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 200 | T.is_kind T.Command tk andalso T.val_of tk mem !hidden_commands | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 201 | | is_hidden _ = false; | 
| 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 202 | |
| 15570 | 203 | fun filter_newlines toks = (case List.mapPartial | 
| 15531 | 204 | (fn (tk, i) => if is_newline tk then SOME tk else NONE) toks of | 
| 15529 
b86d5c84a9a7
Optimized present_tokens to produce fewer newlines when hiding proofs.
 berghofe parents: 
15473diff
changeset | 205 | [] => [] | [tk] => [tk] | _ :: toks' => toks'); | 
| 
b86d5c84a9a7
Optimized present_tokens to produce fewer newlines when hiding proofs.
 berghofe parents: 
15473diff
changeset | 206 | |
| 15437 | 207 | fun present_tokens lex (flag, toks) state state' = | 
| 15531 | 208 | Buffer.add (case flag of NONE => "" | SOME b => Latex.flag_markup b) o | 
| 15437 | 209 | ((if !hide_commands andalso exists (is_hidden o fst) toks then [] | 
| 15529 
b86d5c84a9a7
Optimized present_tokens to produce fewer newlines when hiding proofs.
 berghofe parents: 
15473diff
changeset | 210 | else if !hide_commands andalso is_proof state then | 
| 
b86d5c84a9a7
Optimized present_tokens to produce fewer newlines when hiding proofs.
 berghofe parents: 
15473diff
changeset | 211 | if is_proof state' then [] else filter_newlines toks | 
| 15570 | 212 | else List.mapPartial (fn (tk, i) => | 
| 15531 | 213 | if i > ! interest_level then NONE else SOME tk) toks) | 
| 15437 | 214 | |> map (apsnd (eval_antiquote lex state')) | 
| 11861 | 215 | |> Latex.output_tokens | 
| 216 | |> Buffer.add); | |
| 9140 | 217 | |
| 218 | ||
| 219 | (* parse_thy *) | |
| 220 | ||
| 221 | datatype markup = Markup | MarkupEnv | Verbatim; | |
| 222 | ||
| 223 | local | |
| 224 | ||
| 225 | val opt_newline = Scan.option (Scan.one T.is_newline); | |
| 226 | ||
| 227 | fun check_level i = | |
| 228 | if i > 0 then Scan.succeed () | |
| 229 | else Scan.fail_with (K "Bad nesting of meta-comments"); | |
| 230 | ||
| 231 | val ignore = | |
| 13775 
a918c547cd4d
corrected swallowing of newlines after end-of-ignore: rollback
 oheimb parents: 
13774diff
changeset | 232 | Scan.depend (fn i => opt_newline |-- P.position (Scan.one T.is_begin_ignore) >> pair (i + 1)) || | 
| 
a918c547cd4d
corrected swallowing of newlines after end-of-ignore: rollback
 oheimb parents: 
13774diff
changeset | 233 | Scan.depend (fn i => P.position (Scan.one T.is_end_ignore) --| | 
| 13774 
77a1e723151d
corrected swallowing of newlines after end-of-ignore (improved)
 oheimb parents: 
13742diff
changeset | 234 | (opt_newline -- check_level i) >> pair (i - 1)); | 
| 9140 | 235 | |
| 236 | val ignore_cmd = Scan.state -- ignore | |
| 15531 | 237 |   >> (fn (i, (x, pos)) => (false, (NONE, ((Latex.Basic x, ("", pos)), i))));
 | 
| 9140 | 238 | |
| 239 | ||
| 240 | val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore); | |
| 241 | val improper = Scan.any is_improper; | |
| 242 | ||
| 243 | val improper_keep_indent = Scan.repeat | |
| 244 | (Scan.unless (Scan.one T.is_indent -- Scan.one T.is_proper) (Scan.one is_improper)); | |
| 245 | ||
| 246 | val improper_end = | |
| 247 | (improper -- P.semicolon) |-- improper_keep_indent || | |
| 248 | improper_keep_indent; | |
| 249 | ||
| 250 | ||
| 251 | val stopper = | |
| 15531 | 252 |   ((false, (NONE, ((Latex.Basic (#1 T.stopper), ("", Position.none)), 0))),
 | 
| 11861 | 253 | fn (_, (_, ((Latex.Basic x, _), _))) => (#2 T.stopper x) | _ => false); | 
| 9140 | 254 | |
| 255 | in | |
| 256 | ||
| 257 | fun parse_thy markup lex trs src = | |
| 258 | let | |
| 259 | val text = P.position P.text; | |
| 9220 | 260 | val token = Scan.depend (fn i => | 
| 9140 | 261 | (improper |-- markup Markup -- P.!!!! (improper |-- text --| improper_end) | 
| 15531 | 262 | >> (fn (x, y) => (true, (SOME true, ((Latex.Markup (T.val_of x), y), i)))) || | 
| 9140 | 263 | improper |-- markup MarkupEnv -- P.!!!! (improper |-- text --| improper_end) | 
| 15531 | 264 | >> (fn (x, y) => (true, (SOME true, ((Latex.MarkupEnv (T.val_of x), y), i)))) || | 
| 9140 | 265 | (P.$$$ "--" |-- P.!!!! (improper |-- text)) | 
| 15531 | 266 | >> (fn y => (false, (NONE, ((Latex.Markup "cmt", y), i)))) || | 
| 9140 | 267 | (improper -- markup Verbatim) |-- P.!!!! (improper |-- text --| improper_end) | 
| 15531 | 268 | >> (fn y => (true, (NONE, ((Latex.Verbatim, y), i)))) || | 
| 9140 | 269 | P.position (Scan.one T.not_eof) | 
| 15531 | 270 |         >> (fn (x, pos) => (T.is_kind T.Command x, (SOME false, ((Latex.Basic x, ("", pos)), i)))))
 | 
| 9220 | 271 | >> pair i); | 
| 9140 | 272 | |
| 273 | val body = Scan.any (not o fst andf not o #2 stopper); | |
| 11861 | 274 | val section = body -- Scan.one fst -- body | 
| 275 | >> (fn ((b1, c), b2) => (#1 (#2 c), map (snd o snd) (b1 @ (c :: b2)))); | |
| 9140 | 276 | |
| 277 | val cmds = | |
| 278 | src | |
| 279 | |> Source.filter (not o T.is_semicolon) | |
| 15531 | 280 | |> Source.source' 0 T.stopper (Scan.error (Scan.bulk (ignore_cmd || token))) NONE | 
| 281 | |> Source.source stopper (Scan.error (Scan.bulk section)) NONE | |
| 9140 | 282 | |> Source.exhaust; | 
| 283 | in | |
| 284 | if length cmds = length trs then | |
| 285 | (trs ~~ map (present_tokens lex) cmds, Buffer.empty) | |
| 11861 | 286 | else error "Messed-up outer syntax for presentation" | 
| 9140 | 287 | end; | 
| 288 | ||
| 289 | end; | |
| 290 | ||
| 291 | ||
| 292 | ||
| 293 | (** setup default output **) | |
| 294 | ||
| 295 | (* options *) | |
| 296 | ||
| 297 | val _ = add_options | |
| 9220 | 298 |  [("show_types", Library.setmp Syntax.show_types o boolean),
 | 
| 299 |   ("show_sorts", Library.setmp Syntax.show_sorts o boolean),
 | |
| 14707 | 300 |   ("show_structs", Library.setmp show_structs o boolean),
 | 
| 15473 | 301 |   ("show_var_qmarks", Library.setmp show_var_qmarks o boolean),
 | 
| 14696 | 302 |   ("long_names", Library.setmp NameSpace.long_names o boolean),
 | 
| 9220 | 303 |   ("eta_contract", Library.setmp Syntax.eta_contract o boolean),
 | 
| 14899 | 304 |   ("locale", Library.setmp locale),
 | 
| 9140 | 305 |   ("display", Library.setmp display o boolean),
 | 
| 13929 
21615e44ba88
Added "break" flag to allow line breaks within \isa{...}
 berghofe parents: 
13775diff
changeset | 306 |   ("break", Library.setmp break o boolean),
 | 
| 9140 | 307 |   ("quotes", Library.setmp quotes o boolean),
 | 
| 308 |   ("mode", fn s => fn f => fn () => Library.setmp print_mode (s :: ! print_mode) f ()),
 | |
| 9732 | 309 |   ("margin", Pretty.setmp_margin o integer),
 | 
| 9750 | 310 |   ("indent", Library.setmp indent o integer),
 | 
| 10321 | 311 |   ("source", Library.setmp source o boolean),
 | 
| 312 |   ("goals_limit", Library.setmp goals_limit o integer)];
 | |
| 9140 | 313 | |
| 314 | ||
| 315 | (* commands *) | |
| 316 | ||
| 317 | fun cond_quote prt = | |
| 318 | if ! quotes then Pretty.quote prt else prt; | |
| 319 | ||
| 320 | fun cond_display prt = | |
| 321 | if ! display then | |
| 14835 | 322 | Output.output (Pretty.string_of (Pretty.indent (! indent) prt)) | 
| 9863 | 323 |     |> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
 | 
| 9140 | 324 | else | 
| 14835 | 325 | Output.output ((if ! break then Pretty.string_of else Pretty.str_of) prt) | 
| 9140 | 326 |     |> enclose "\\isa{" "}";
 | 
| 327 | ||
| 14345 | 328 | fun cond_seq_display prts = | 
| 329 | if ! display then | |
| 14835 | 330 | map (Output.output o Pretty.string_of o Pretty.indent (! indent)) prts | 
| 14345 | 331 | |> separate "\\isasep\\isanewline%\n" | 
| 332 | |> implode | |
| 333 |     |> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
 | |
| 334 | else | |
| 14835 | 335 | map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of)) prts | 
| 14345 | 336 | |> separate "\\isasep\\isanewline%\n" | 
| 337 | |> implode | |
| 338 |     |> enclose "\\isa{" "}";
 | |
| 339 | ||
| 11011 
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
 wenzelm parents: 
10739diff
changeset | 340 | fun tweak_line s = | 
| 
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
 wenzelm parents: 
10739diff
changeset | 341 | if ! display then s else Symbol.strip_blanks s; | 
| 9750 | 342 | |
| 11011 
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
 wenzelm parents: 
10739diff
changeset | 343 | val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines; | 
| 9140 | 344 | |
| 9750 | 345 | val pretty_source = | 
| 346 | pretty_text o space_implode " " o map Args.string_of o #2 o #1 o Args.dest_src; | |
| 9140 | 347 | |
| 12055 | 348 | fun pretty_term ctxt = ProofContext.pretty_term ctxt o ProofContext.extern_skolem ctxt; | 
| 9140 | 349 | |
| 14345 | 350 | fun pretty_thm ctxt = pretty_term ctxt o Thm.prop_of; | 
| 9140 | 351 | |
| 15349 | 352 | fun lhs_of (Const ("==",_) $ t $ _) = t
 | 
| 353 |   | lhs_of (Const ("Trueprop",_) $ t) = lhs_of t
 | |
| 354 |   | lhs_of (Const ("==>",_) $ _ $ t) = lhs_of t
 | |
| 355 | | lhs_of (_ $ t $ _) = t | |
| 356 |   | lhs_of _ = error ("Binary operator expected")
 | |
| 357 | ||
| 358 | fun pretty_lhs ctxt = pretty_term ctxt o lhs_of o Thm.prop_of; | |
| 359 | ||
| 360 | fun rhs_of (Const ("==",_) $ _ $ t) = t
 | |
| 361 |   | rhs_of (Const ("Trueprop",_) $ t) = rhs_of t
 | |
| 362 |   | rhs_of (Const ("==>",_) $ _ $ t) = rhs_of t
 | |
| 363 | | rhs_of (_ $ _ $ t) = t | |
| 364 |   | rhs_of _ = error ("Binary operator expected")
 | |
| 365 | ||
| 366 | fun pretty_rhs ctxt = pretty_term ctxt o rhs_of o Thm.prop_of; | |
| 367 | ||
| 12055 | 368 | fun pretty_prf full ctxt thms = (* FIXME context syntax!? *) | 
| 11524 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11240diff
changeset | 369 | Pretty.chunks (map (ProofSyntax.pretty_proof_of full) thms); | 
| 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11240diff
changeset | 370 | |
| 9750 | 371 | fun output_with pretty src ctxt x = | 
| 372 | let | |
| 12055 | 373 | val prt = pretty ctxt x; (*always pretty in order to catch errors!*) | 
| 9750 | 374 | val prt' = if ! source then pretty_source src else prt; | 
| 375 | in cond_display (cond_quote prt') end; | |
| 9732 | 376 | |
| 14345 | 377 | fun output_seq_with pretty src ctxt xs = | 
| 378 | let | |
| 379 | val prts = map (pretty ctxt) xs; (*always pretty in order to catch errors!*) | |
| 380 | val prts' = if ! source then [pretty_source src] else prts; | |
| 381 | in cond_seq_display (map cond_quote prts') end; | |
| 382 | ||
| 10360 | 383 | fun output_goals main_goal src state = args (Scan.succeed ()) (output_with (fn _ => fn _ => | 
| 384 | Pretty.chunks (Proof.pretty_goals main_goal (Toplevel.proof_of state | |
| 385 | handle Toplevel.UNDEF => error "No proof state")))) src state; | |
| 386 | ||
| 9140 | 387 | val _ = add_commands | 
| 14696 | 388 |  [("thm", args Attrib.local_thmss (output_seq_with pretty_thm)),
 | 
| 15349 | 389 |   ("lhs", args Attrib.local_thmss (output_seq_with pretty_lhs)),
 | 
| 390 |   ("rhs", args Attrib.local_thmss (output_seq_with pretty_rhs)),
 | |
| 9750 | 391 |   ("prop", args Args.local_prop (output_with pretty_term)),
 | 
| 392 |   ("term", args Args.local_term (output_with pretty_term)),
 | |
| 14776 | 393 |   ("typ", args Args.local_typ_raw (output_with ProofContext.pretty_typ)),
 | 
| 14696 | 394 |   ("text", args (Scan.lift Args.name) (output_with (K pretty_text))),
 | 
| 10360 | 395 |   ("goals", output_goals true),
 | 
| 14696 | 396 |   ("subgoals", output_goals false),
 | 
| 397 |   ("prf", args Attrib.local_thmss (output_with (pretty_prf false))),
 | |
| 398 |   ("full_prf", args Attrib.local_thmss (output_with (pretty_prf true)))];
 | |
| 9140 | 399 | |
| 400 | end; |