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