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