| author | wenzelm | 
| Fri, 05 May 2006 21:59:43 +0200 | |
| changeset 19575 | 2d9940cd52d3 | 
| parent 19058 | 1e65cf5ae9ea | 
| child 19581 | 4ae6a14b742f | 
| permissions | -rw-r--r-- | 
| 9140 | 1 | (* Title: Pure/Isar/isar_output.ML | 
| 2 | ID: $Id$ | |
| 16002 | 3 | Author: Markus Wenzel, TU Muenchen | 
| 9140 | 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 | |
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 14 | val add_commands: (string * (Args.src -> Toplevel.node option -> 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 | |
| 18988 | 19 | val args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> | 
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 20 | (Args.src -> Proof.context -> 'a -> string) -> Args.src -> Toplevel.node option -> string | 
| 9140 | 21 | datatype markup = Markup | MarkupEnv | Verbatim | 
| 11714 | 22 | val modes: string list ref | 
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 23 | val eval_antiquote: Scan.lexicon -> Toplevel.node option -> string * Position.T -> string | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 24 | val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) -> | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 25 | Toplevel.transition list -> (OuterLex.token, 'a) Source.source -> Buffer.T | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 26 | val output_list: (Proof.context -> 'a -> Pretty.T) -> Args.src -> | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 27 | Proof.context -> 'a list -> string | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 28 | val output: (Proof.context -> 'a -> Pretty.T) -> Args.src -> Proof.context -> 'a -> string | 
| 9140 | 29 | end; | 
| 30 | ||
| 31 | structure IsarOutput: ISAR_OUTPUT = | |
| 32 | struct | |
| 33 | ||
| 34 | structure T = OuterLex; | |
| 35 | structure P = OuterParse; | |
| 36 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 37 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 38 | (** global options **) | 
| 14899 | 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 | ||
| 9140 | 49 | (** maintain global commands **) | 
| 50 | ||
| 51 | local | |
| 52 | ||
| 53 | val global_commands = | |
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 54 | ref (Symtab.empty: (Args.src -> Toplevel.node option -> string) Symtab.table); | 
| 9140 | 55 | |
| 56 | val global_options = | |
| 57 | ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table); | |
| 58 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 59 | fun add_item kind (name, x) tab = | 
| 16894 | 60 | (if not (Symtab.defined tab name) then () | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 61 |   else warning ("Redefined antiquotation " ^ kind ^ ": " ^ quote name);
 | 
| 17412 | 62 | Symtab.update (name, x) tab); | 
| 9140 | 63 | |
| 64 | in | |
| 65 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 66 | val add_commands = Library.change global_commands o fold (add_item "command"); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 67 | val add_options = Library.change global_options o fold (add_item "option"); | 
| 9140 | 68 | |
| 69 | fun command src = | |
| 70 | let val ((name, _), pos) = Args.dest_src src in | |
| 17412 | 71 | (case Symtab.lookup (! global_commands) name of | 
| 15531 | 72 |       NONE => error ("Unknown antiquotation command: " ^ quote name ^ Position.str_of pos)
 | 
| 73 | | SOME f => transform_failure (curry Antiquote.ANTIQUOTE_FAIL (name, pos)) (f src)) | |
| 9140 | 74 | end; | 
| 75 | ||
| 76 | fun option (name, s) f () = | |
| 17412 | 77 | (case Symtab.lookup (! global_options) name of | 
| 15531 | 78 |     NONE => error ("Unknown antiquotation option: " ^ quote name)
 | 
| 79 | | SOME opt => opt s f ()); | |
| 9140 | 80 | |
| 81 | fun options [] f = f | |
| 82 | | options (opt :: opts) f = option opt (options opts f); | |
| 83 | ||
| 9213 | 84 | |
| 9220 | 85 | fun print_antiquotations () = | 
| 86 | [Pretty.big_list "antiquotation commands:" (map Pretty.str (Symtab.keys (! global_commands))), | |
| 87 | Pretty.big_list "antiquotation options:" (map Pretty.str (Symtab.keys (! global_options)))] | |
| 88 | |> Pretty.chunks |> Pretty.writeln; | |
| 9213 | 89 | |
| 9140 | 90 | end; | 
| 91 | ||
| 92 | ||
| 93 | ||
| 94 | (** syntax of antiquotations **) | |
| 95 | ||
| 96 | (* option values *) | |
| 97 | ||
| 98 | fun boolean "" = true | |
| 99 | | boolean "true" = true | |
| 100 | | boolean "false" = false | |
| 101 |   | boolean s = error ("Bad boolean value: " ^ quote s);
 | |
| 102 | ||
| 103 | fun integer s = | |
| 104 | let | |
| 105 | fun int ss = | |
| 14899 | 106 | (case Library.read_int ss of (i, []) => i | 
| 107 |       | _ => error ("Bad integer value: " ^ quote s));
 | |
| 9140 | 108 | in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end; | 
| 109 | ||
| 110 | ||
| 111 | (* args syntax *) | |
| 112 | ||
| 18988 | 113 | fun syntax scan = Args.context_syntax "antiquotation" scan; | 
| 9140 | 114 | |
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 115 | fun args scan f src node : string = | 
| 14899 | 116 | let | 
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 117 | val loc = if ! locale = "" then NONE else SOME (! locale); | 
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 118 | val (ctxt, x) = syntax scan src (Toplevel.body_context_node node loc); | 
| 10321 | 119 | in f src ctxt x end; | 
| 9140 | 120 | |
| 121 | ||
| 122 | (* outer syntax *) | |
| 123 | ||
| 124 | local | |
| 125 | ||
| 126 | val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) ""; | |
| 127 | val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) []; | |
| 128 | ||
| 129 | val antiq = P.position P.xname -- properties -- P.arguments --| Scan.ahead P.eof | |
| 130 | >> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos))); | |
| 131 | ||
| 132 | fun antiq_args_aux keyword_lexicon (str, pos) = | |
| 133 | Source.of_string str | |
| 134 | |> Symbol.source false | |
| 135 | |> T.source false (K (keyword_lexicon, Scan.empty_lexicon)) pos | |
| 136 | |> T.source_proper | |
| 15531 | 137 | |> Source.source T.stopper (Scan.error (Scan.bulk (P.!!! antiq))) NONE | 
| 9140 | 138 | |> Source.exhaust; | 
| 139 | ||
| 140 | in | |
| 141 | ||
| 142 | fun antiq_args lex (s, pos) = | |
| 18678 | 143 | let | 
| 144 | fun err msg = cat_error msg | |
| 145 |       ("Malformed antiquotation: " ^ quote ("@{" ^ s ^ "}") ^ Position.str_of pos);
 | |
| 146 | in (case antiq_args_aux lex (s, pos) of [x] => x | _ => err "") handle ERROR msg => err msg end; | |
| 9140 | 147 | |
| 148 | end; | |
| 149 | ||
| 150 | ||
| 151 | (* eval_antiquote *) | |
| 152 | ||
| 11714 | 153 | val modes = ref ([]: string list); | 
| 154 | ||
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 155 | fun eval_antiquote lex node (str, pos) = | 
| 9140 | 156 | let | 
| 157 | fun expand (Antiquote.Text s) = s | |
| 158 | | expand (Antiquote.Antiq x) = | |
| 159 | let val (opts, src) = antiq_args lex x in | |
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 160 | options opts (fn () => command src node) (); (*preview errors!*) | 
| 11714 | 161 | Library.setmp print_mode (! modes @ Latex.modes @ ! print_mode) | 
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 162 | (Output.no_warnings (options opts (fn () => command src node))) () | 
| 10321 | 163 | end; | 
| 9140 | 164 | val ants = Antiquote.antiquotes_of (str, pos); | 
| 165 | in | |
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 166 | if is_none node andalso exists Antiquote.is_antiq ants then | 
| 9140 | 167 |       error ("Cannot expand antiquotations at top-level" ^ Position.str_of pos)
 | 
| 168 | else implode (map expand ants) | |
| 169 | end; | |
| 170 | ||
| 171 | ||
| 172 | ||
| 173 | (** present theory source **) | |
| 174 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 175 | (* presentation tokens *) | 
| 9140 | 176 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 177 | datatype token = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 178 | NoToken | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 179 | | BasicToken of T.token | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 180 | | MarkupToken of string * (string * Position.T) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 181 | | MarkupEnvToken of string * (string * Position.T) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 182 | | VerbatimToken of string * Position.T; | 
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 183 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 184 | fun output_token lex state = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 185 | let | 
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 186 | val eval = eval_antiquote lex (try Toplevel.node_of state) | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 187 | in | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 188 | fn NoToken => "" | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 189 | | BasicToken tok => Latex.output_basic tok | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 190 | | MarkupToken (cmd, txt) => Latex.output_markup cmd (eval txt) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 191 | | MarkupEnvToken (cmd, txt) => Latex.output_markup_env cmd (eval txt) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 192 | | VerbatimToken txt => Latex.output_verbatim (eval txt) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 193 | end; | 
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 194 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 195 | fun basic_token pred (BasicToken tok) = pred tok | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 196 | | basic_token _ _ = false; | 
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 197 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 198 | val improper_token = basic_token (not o T.is_proper); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 199 | val comment_token = basic_token T.is_comment; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 200 | val blank_token = basic_token T.is_blank; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 201 | val newline_token = basic_token T.is_newline; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 202 | |
| 15430 
1e1aeaf1dec3
Implemented hiding of proofs and other commands.
 berghofe parents: 
15349diff
changeset | 203 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 204 | (* command spans *) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 205 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 206 | type command = string * Position.T * string list; (*name, position, tags*) | 
| 17756 | 207 | type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*) | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 208 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 209 | datatype span = Span of command * (source * source * source * source) * bool; | 
| 15529 
b86d5c84a9a7
Optimized present_tokens to produce fewer newlines when hiding proofs.
 berghofe parents: 
15473diff
changeset | 210 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 211 | fun make_span cmd src = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 212 | let | 
| 17756 | 213 | fun take_newline (tok :: toks) = | 
| 214 | if newline_token (fst tok) then ([tok], toks, true) | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 215 | else ([], tok :: toks, false) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 216 | | take_newline [] = ([], [], false); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 217 | val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 218 | src | 
| 17756 | 219 | |> take_prefix (improper_token o fst) | 
| 220 | ||>> take_suffix (improper_token o fst) | |
| 221 | ||>> take_prefix (comment_token o fst) | |
| 17185 
5140808111d1
clarify type tok, do not emit markup flag for suppressed tokens;
 wenzelm parents: 
17151diff
changeset | 222 | ||> take_newline; | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 223 | in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end; | 
| 9140 | 224 | |
| 225 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 226 | (* present spans *) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 227 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 228 | local | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 229 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 230 | fun err_bad_nesting pos = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 231 |   error ("Bad nesting of commands in presentation" ^ pos);
 | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 232 | |
| 17557 | 233 | fun edge1 f (x, y) = the_default I (Option.map (Buffer.add o f) (if x = y then NONE else x)); | 
| 234 | fun edge2 f (x, y) = the_default I (Option.map (Buffer.add o f) (if x = y then NONE else y)); | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 235 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 236 | val begin_tag = edge2 Latex.begin_tag; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 237 | val end_tag = edge1 Latex.end_tag; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 238 | fun open_delim delim e = edge2 Latex.begin_delim e #> delim #> edge2 Latex.end_delim e; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 239 | fun close_delim delim e = edge1 Latex.begin_delim e #> delim #> edge1 Latex.end_delim e; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 240 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 241 | in | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 242 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 243 | fun present_span lex default_tags span state state' | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 244 | (tag_stack, active_tag, newline, buffer, present_cont) = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 245 | let | 
| 17756 | 246 | val present = fold (fn (tok, (flag, 0)) => | 
| 17263 | 247 | Buffer.add (output_token lex state' tok) | 
| 248 | #> Buffer.add flag | |
| 249 | | _ => I); | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 250 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 251 | val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 252 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 253 | val (tag, tags) = tag_stack; | 
| 17557 | 254 | val tag' = try hd (fold OuterKeyword.update_tags cmd_tags (the_list tag)); | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 255 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 256 | val active_tag' = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 257 | if is_some tag' then tag' | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 258 | else try hd (default_tags cmd_name); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 259 | val edge = (active_tag, active_tag'); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 260 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 261 | val newline' = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 262 | if is_none active_tag' then span_newline else newline; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 263 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 264 | val nesting = Toplevel.level state' - Toplevel.level state; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 265 | val tag_stack' = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 266 | if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 267 | else if nesting >= 0 then (tag', replicate nesting tag @ tags) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 268 | else | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 269 | (case Library.drop (~ nesting - 1, tags) of | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 270 | tgs :: tgss => (tgs, tgss) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 271 | | [] => err_bad_nesting (Position.str_of cmd_pos)); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 272 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 273 | val buffer' = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 274 | buffer | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 275 | |> end_tag edge | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 276 | |> close_delim (fst present_cont) edge | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 277 | |> snd present_cont | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 278 | |> open_delim (present (#1 srcs)) edge | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 279 | |> begin_tag edge | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 280 | |> present (#2 srcs); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 281 | val present_cont' = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 282 | if newline then (present (#3 srcs), present (#4 srcs)) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 283 | else (I, present (#3 srcs) #> present (#4 srcs)); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 284 | in (tag_stack', active_tag', newline', buffer', present_cont') end; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 285 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 286 | fun present_trailer ((_, tags), active_tag, _, buffer, present_cont) = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 287 | if not (null tags) then err_bad_nesting " at end of theory" | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 288 | else | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 289 | buffer | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 290 | |> end_tag (active_tag, NONE) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 291 | |> close_delim (fst present_cont) (active_tag, NONE) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 292 | |> snd present_cont; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 293 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 294 | end; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 295 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 296 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 297 | (* present_thy *) | 
| 9140 | 298 | |
| 299 | datatype markup = Markup | MarkupEnv | Verbatim; | |
| 300 | ||
| 301 | local | |
| 302 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 303 | val space_proper = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 304 | Scan.one T.is_blank -- Scan.any T.is_comment -- Scan.one T.is_proper; | 
| 9140 | 305 | |
| 306 | val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore); | |
| 307 | val improper = Scan.any is_improper; | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 308 | val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper)); | 
| 9140 | 309 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 310 | val opt_newline = Scan.option (Scan.one T.is_newline); | 
| 9140 | 311 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 312 | val ignore = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 313 | Scan.depend (fn d => opt_newline |-- Scan.one T.is_begin_ignore | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 314 | >> pair (d + 1)) || | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 315 | Scan.depend (fn d => Scan.one T.is_end_ignore --| | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 316 | (if d = 0 then Scan.fail_with (K "Bad nesting of meta-comments") else opt_newline) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 317 | >> pair (d - 1)); | 
| 9140 | 318 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 319 | val tag = (improper -- P.$$$ "%" -- improper) |-- P.!!! (P.tag_name --| improper_end); | 
| 9140 | 320 | |
| 17263 | 321 | val locale = | 
| 322 |   Scan.option ((P.$$$ "(" -- improper -- P.$$$ "in") |--
 | |
| 323 | P.!!! (improper |-- P.xname --| (improper -- P.$$$ ")"))); | |
| 324 | ||
| 9140 | 325 | in | 
| 326 | ||
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 327 | fun present_thy lex default_tags is_markup trs src = | 
| 9140 | 328 | let | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 329 | (* tokens *) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 330 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 331 | val ignored = Scan.state --| ignore | 
| 17756 | 332 |       >> (fn d => (NONE, (NoToken, ("", d))));
 | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 333 | |
| 17185 
5140808111d1
clarify type tok, do not emit markup flag for suppressed tokens;
 wenzelm parents: 
17151diff
changeset | 334 | fun markup mark mk flag = Scan.peek (fn d => | 
| 17263 | 335 | improper |-- P.position (Scan.one (T.is_kind T.Command andf is_markup mark o T.val_of)) -- | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 336 | Scan.repeat tag -- | 
| 17263 | 337 | P.!!!! ((improper -- locale -- improper) |-- P.position P.text --| improper_end) | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 338 | >> (fn (((tok, pos), tags), txt) => | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 339 | let val name = T.val_of tok | 
| 17756 | 340 | in (SOME (name, pos, tags), (mk (name, txt), (flag, d))) end)); | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 341 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 342 | val command = Scan.peek (fn d => | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 343 | P.position (Scan.one (T.is_kind T.Command)) -- | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 344 | Scan.repeat tag | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 345 | >> (fn ((tok, pos), tags) => | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 346 | let val name = T.val_of tok | 
| 17756 | 347 | in (SOME (name, pos, tags), (BasicToken tok, (Latex.markup_false, d))) end)); | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 348 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 349 | val cmt = Scan.peek (fn d => | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 350 | P.$$$ "--" |-- P.!!!! (improper |-- P.position P.text) | 
| 17756 | 351 |       >> (fn txt => (NONE, (MarkupToken ("cmt", txt), ("", d)))));
 | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 352 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 353 | val other = Scan.peek (fn d => | 
| 17756 | 354 |        Scan.one T.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d)))));
 | 
| 9140 | 355 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 356 | val token = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 357 | ignored || | 
| 17185 
5140808111d1
clarify type tok, do not emit markup flag for suppressed tokens;
 wenzelm parents: 
17151diff
changeset | 358 | markup Markup MarkupToken Latex.markup_true || | 
| 
5140808111d1
clarify type tok, do not emit markup flag for suppressed tokens;
 wenzelm parents: 
17151diff
changeset | 359 | markup MarkupEnv MarkupEnvToken Latex.markup_true || | 
| 
5140808111d1
clarify type tok, do not emit markup flag for suppressed tokens;
 wenzelm parents: 
17151diff
changeset | 360 | markup Verbatim (VerbatimToken o #2) "" || | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 361 | command || cmt || other; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 362 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 363 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 364 | (* spans *) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 365 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 366 | val stopper = | 
| 17756 | 367 |       ((NONE, (BasicToken (#1 T.stopper), ("", 0))),
 | 
| 368 | fn (_, (BasicToken x, _)) => #2 T.stopper x | _ => false); | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 369 | |
| 17756 | 370 | val cmd = Scan.one (is_some o fst); | 
| 371 | val non_cmd = Scan.one (is_none o fst andf not o #2 stopper) >> #2; | |
| 9140 | 372 | |
| 17756 | 373 | val comments = Scan.any (comment_token o fst o snd); | 
| 374 | val blank = Scan.one (blank_token o fst o snd); | |
| 375 | val newline = Scan.one (newline_token o fst o snd); | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 376 | val before_cmd = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 377 | Scan.option (newline -- comments) -- | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 378 | Scan.option (newline -- comments) -- | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 379 | Scan.option (blank -- comments) -- cmd; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 380 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 381 | val span = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 382 | Scan.repeat non_cmd -- cmd -- | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 383 | Scan.repeat (Scan.unless before_cmd non_cmd) -- | 
| 17756 | 384 | Scan.option (newline >> (single o snd)) | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 385 | >> (fn (((toks1, (cmd, tok2)), toks3), tok4) => | 
| 17756 | 386 | make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4)))); | 
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 387 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 388 | val spans = | 
| 9140 | 389 | src | 
| 390 | |> Source.filter (not o T.is_semicolon) | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 391 | |> Source.source' 0 T.stopper (Scan.error (Scan.bulk token)) NONE | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 392 | |> Source.source stopper (Scan.error (Scan.bulk span)) NONE | 
| 9140 | 393 | |> Source.exhaust; | 
| 394 | in | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 395 | if length trs = length spans then | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 396 | ((NONE, []), NONE, true, Buffer.empty, (I, I)) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 397 | |> Toplevel.present_excursion (trs ~~ map (present_span lex default_tags) spans) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 398 | |> present_trailer | 
| 11861 | 399 | else error "Messed-up outer syntax for presentation" | 
| 9140 | 400 | end; | 
| 401 | ||
| 402 | end; | |
| 403 | ||
| 404 | ||
| 405 | ||
| 406 | (** setup default output **) | |
| 407 | ||
| 408 | (* options *) | |
| 409 | ||
| 410 | val _ = add_options | |
| 9220 | 411 |  [("show_types", Library.setmp Syntax.show_types o boolean),
 | 
| 412 |   ("show_sorts", Library.setmp Syntax.show_sorts o boolean),
 | |
| 14707 | 413 |   ("show_structs", Library.setmp show_structs o boolean),
 | 
| 15988 | 414 |   ("show_question_marks", Library.setmp show_question_marks o boolean),
 | 
| 14696 | 415 |   ("long_names", Library.setmp NameSpace.long_names o boolean),
 | 
| 16142 | 416 |   ("short_names", Library.setmp NameSpace.short_names o boolean),
 | 
| 417 |   ("unique_names", Library.setmp NameSpace.unique_names o boolean),
 | |
| 9220 | 418 |   ("eta_contract", Library.setmp Syntax.eta_contract o boolean),
 | 
| 14899 | 419 |   ("locale", Library.setmp locale),
 | 
| 9140 | 420 |   ("display", Library.setmp display o boolean),
 | 
| 13929 
21615e44ba88
Added "break" flag to allow line breaks within \isa{...}
 berghofe parents: 
13775diff
changeset | 421 |   ("break", Library.setmp break o boolean),
 | 
| 9140 | 422 |   ("quotes", Library.setmp quotes o boolean),
 | 
| 423 |   ("mode", fn s => fn f => fn () => Library.setmp print_mode (s :: ! print_mode) f ()),
 | |
| 9732 | 424 |   ("margin", Pretty.setmp_margin o integer),
 | 
| 9750 | 425 |   ("indent", Library.setmp indent o integer),
 | 
| 10321 | 426 |   ("source", Library.setmp source o boolean),
 | 
| 427 |   ("goals_limit", Library.setmp goals_limit o integer)];
 | |
| 9140 | 428 | |
| 16002 | 429 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 430 | (* basic pretty printing *) | 
| 9140 | 431 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 432 | val str_of_source = space_implode " " o map Args.string_of o #2 o #1 o Args.dest_src; | 
| 14345 | 433 | |
| 11011 
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
 wenzelm parents: 
10739diff
changeset | 434 | fun tweak_line s = | 
| 
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
 wenzelm parents: 
10739diff
changeset | 435 | if ! display then s else Symbol.strip_blanks s; | 
| 9750 | 436 | |
| 11011 
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
 wenzelm parents: 
10739diff
changeset | 437 | val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines; | 
| 9140 | 438 | |
| 12055 | 439 | fun pretty_term ctxt = ProofContext.pretty_term ctxt o ProofContext.extern_skolem ctxt; | 
| 9140 | 440 | |
| 16002 | 441 | fun pretty_term_typ ctxt t = | 
| 442 | (Syntax.const Syntax.constrainC $ | |
| 443 | ProofContext.extern_skolem ctxt t $ | |
| 444 | Syntax.term_of_typ (! show_sorts) (Term.fastype_of t)) | |
| 445 | |> ProofContext.pretty_term ctxt; | |
| 15880 
d6aa6c707acf
added antiquotations typeof, const, term_style, thm_style, term_type (something still to be done)
 haftmann parents: 
15666diff
changeset | 446 | |
| 16002 | 447 | fun pretty_term_typeof ctxt = ProofContext.pretty_typ ctxt o Term.fastype_of; | 
| 15880 
d6aa6c707acf
added antiquotations typeof, const, term_style, thm_style, term_type (something still to be done)
 haftmann parents: 
15666diff
changeset | 448 | |
| 16002 | 449 | fun pretty_term_const ctxt t = | 
| 450 | if Term.is_Const t then pretty_term ctxt t | |
| 451 |   else error ("Logical constant expected: " ^ ProofContext.string_of_term ctxt t);
 | |
| 15880 
d6aa6c707acf
added antiquotations typeof, const, term_style, thm_style, term_type (something still to be done)
 haftmann parents: 
15666diff
changeset | 452 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 453 | fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 454 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 455 | fun pretty_term_style ctxt (name, t) = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 456 | pretty_term ctxt (TermStyle.the_style (ProofContext.theory_of ctxt) name ctxt t); | 
| 9140 | 457 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 458 | fun pretty_thm_style ctxt (name, th) = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 459 | pretty_term_style ctxt (name, Thm.full_prop_of th); | 
| 15880 
d6aa6c707acf
added antiquotations typeof, const, term_style, thm_style, term_type (something still to be done)
 haftmann parents: 
15666diff
changeset | 460 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 461 | fun pretty_prf full ctxt thms = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 462 | Pretty.chunks (map (ProofContext.pretty_proof_of ctxt full) thms); | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 463 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 464 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 465 | (* Isar output *) | 
| 15880 
d6aa6c707acf
added antiquotations typeof, const, term_style, thm_style, term_type (something still to be done)
 haftmann parents: 
15666diff
changeset | 466 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 467 | fun output_list pretty src ctxt xs = | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 468 | map (pretty ctxt) xs (*always pretty in order to exhibit errors!*) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 469 | |> (if ! source then K [pretty_text (str_of_source src)] else I) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 470 | |> (if ! quotes then map Pretty.quote else I) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 471 | |> (if ! display then | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 472 | map (Output.output o Pretty.string_of o Pretty.indent (! indent)) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 473 | #> space_implode "\\isasep\\isanewline%\n" | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 474 |     #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
 | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 475 | else | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 476 | map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of)) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 477 | #> space_implode "\\isasep\\isanewline%\n" | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 478 |     #> enclose "\\isa{" "}");
 | 
| 11524 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11240diff
changeset | 479 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 480 | fun output pretty src ctxt = output_list pretty src ctxt o single; | 
| 9732 | 481 | |
| 19058 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 482 | fun proof_state node = | 
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 483 | (case Option.map Toplevel.proof_node node of | 
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 484 | SOME (SOME prf) => ProofHistory.current prf | 
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 485 | | _ => error "No proof state"); | 
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 486 | |
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 487 | fun output_goals main_goal src node = args (Scan.succeed ()) (output (fn _ => fn _ => | 
| 
1e65cf5ae9ea
evaluate antiquotes depending on Toplevel.node option;
 wenzelm parents: 
18988diff
changeset | 488 | Pretty.chunks (Proof.pretty_goals main_goal (proof_state node)))) src node; | 
| 14345 | 489 | |
| 17863 | 490 | fun ml_val txt = "fn _ => (" ^ txt ^ ");";
 | 
| 491 | fun ml_type txt = "val _ = NONE : (" ^ txt ^ ") option;";
 | |
| 492 | fun ml_struct txt = "functor DUMMY_FUNCTOR() = struct structure DUMMY = " ^ txt ^ " end;" | |
| 493 | ||
| 494 | fun output_ml ml src ctxt txt = | |
| 495 | (Context.use_mltext (ml txt) false (SOME (ProofContext.theory_of ctxt)); | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 496 | (if ! source then str_of_source src else txt) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 497 | |> (if ! quotes then quote else I) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 498 |   |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
 | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 499 | else | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 500 | split_lines | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 501 | #> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|") | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 502 | #> space_implode "\\isasep\\isanewline%\n")); | 
| 10360 | 503 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 504 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 505 | (* commands *) | 
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 506 | |
| 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 507 | val _ = add_commands | 
| 18988 | 508 |  [("thm", args Attrib.thms (output_list pretty_thm)),
 | 
| 509 |   ("thm_style", args (Scan.lift Args.liberal_name -- Attrib.thm) (output pretty_thm_style)),
 | |
| 510 |   ("prop", args Args.prop (output pretty_term)),
 | |
| 511 |   ("term", args Args.term (output pretty_term)),
 | |
| 512 |   ("term_style", args (Scan.lift Args.liberal_name -- Args.term) (output pretty_term_style)),
 | |
| 513 |   ("term_type", args Args.term (output pretty_term_typ)),
 | |
| 514 |   ("typeof", args Args.term (output pretty_term_typeof)),
 | |
| 515 |   ("const", args Args.term (output pretty_term_const)),
 | |
| 516 |   ("typ", args Args.typ_abbrev (output ProofContext.pretty_typ)),
 | |
| 17067 
eb07469a4cdd
reimplemented theory presentation, with support for tagged command regions;
 wenzelm parents: 
17030diff
changeset | 517 |   ("text", args (Scan.lift Args.name) (output (K pretty_text))),
 | 
| 10360 | 518 |   ("goals", output_goals true),
 | 
| 14696 | 519 |   ("subgoals", output_goals false),
 | 
| 18988 | 520 |   ("prf", args Attrib.thms (output (pretty_prf false))),
 | 
| 521 |   ("full_prf", args Attrib.thms (output (pretty_prf true))),
 | |
| 17863 | 522 |   ("ML", args (Scan.lift Args.name) (output_ml ml_val)),
 | 
| 523 |   ("ML_type", args (Scan.lift Args.name) (output_ml ml_type)),
 | |
| 524 |   ("ML_struct", args (Scan.lift Args.name) (output_ml ml_struct))];
 | |
| 9140 | 525 | |
| 526 | end; |