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