1 (* Title: Pure/Thy/thy_output.ML
2 Author: Markus Wenzel, TU Muenchen
4 Theory document output with antiquotations.
9 val display: bool Unsynchronized.ref
10 val quotes: bool Unsynchronized.ref
11 val indent: int Unsynchronized.ref
12 val source: bool Unsynchronized.ref
13 val break: bool Unsynchronized.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 -> 'a context_parser ->
22 ({state: Toplevel.state, source: Args.src, context: Proof.context} -> 'a -> string) -> unit
23 datatype markup = Markup | MarkupEnv | Verbatim
24 val modes: string list Unsynchronized.ref
25 val eval_antiquote: Scan.lexicon -> Toplevel.state -> Symbol_Pos.text * Position.T -> string
26 val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) ->
27 (Toplevel.transition * Toplevel.state) list -> (Token.T, 'a) Source.source -> Buffer.T
28 val pretty_text: string -> Pretty.T
29 val pretty_term: Proof.context -> term -> Pretty.T
30 val pretty_thm: Proof.context -> thm -> Pretty.T
31 val str_of_source: Args.src -> string
32 val maybe_pretty_source: ('a -> Pretty.T) -> Args.src -> 'a list -> Pretty.T list
33 val output: Pretty.T list -> string
36 structure Thy_Output: THY_OUTPUT =
39 (** global options **)
41 val display = Unsynchronized.ref false;
42 val quotes = Unsynchronized.ref false;
43 val indent = Unsynchronized.ref 0;
44 val source = Unsynchronized.ref false;
45 val break = Unsynchronized.ref false;
49 (** maintain global antiquotations **)
54 Unsynchronized.ref (Symtab.empty: (Args.src -> Toplevel.state -> string) Symtab.table);
57 Unsynchronized.ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table);
59 fun add_item kind (name, x) tab =
60 (if not (Symtab.defined tab name) then ()
61 else warning ("Redefined document antiquotation " ^ kind ^ ": " ^ quote name);
62 Symtab.update (name, x) tab);
67 CRITICAL (fn () => Unsynchronized.change global_commands (fold (add_item "command") xs));
69 CRITICAL (fn () => Unsynchronized.change global_options (fold (add_item "option") xs));
71 fun defined_command name = Symtab.defined (! global_commands) name;
72 fun defined_option name = Symtab.defined (! global_options) name;
75 let val ((name, _), pos) = Args.dest_src src in
76 (case Symtab.lookup (! global_commands) name of
77 NONE => error ("Unknown document antiquotation: " ^ quote name ^ Position.str_of pos)
79 (Position.report (Markup.doc_antiq name) pos;
80 (fn state => f src state handle ERROR msg =>
81 cat_error msg ("The error(s) above occurred in document antiquotation: " ^
82 quote name ^ Position.str_of pos))))
85 fun option (name, s) f () =
86 (case Symtab.lookup (! global_options) name of
87 NONE => error ("Unknown document antiquotation option: " ^ quote name)
88 | SOME opt => opt s f ());
91 | options (opt :: opts) f = option opt (options opts f);
94 fun print_antiquotations () =
95 [Pretty.big_list "document antiquotation commands:"
96 (map Pretty.str (sort_strings (Symtab.keys (! global_commands)))),
97 Pretty.big_list "document antiquotation options:"
98 (map Pretty.str (sort_strings (Symtab.keys (! global_options))))]
99 |> Pretty.chunks |> Pretty.writeln;
103 fun antiquotation name scan out = add_commands [(name, fn src => fn state =>
104 let val (x, ctxt) = Args.context_syntax "document antiquotation"
105 scan src (Toplevel.presentation_context_of state)
106 in out {source = src, state = state, context = ctxt} x end)];
110 (** syntax of antiquotations **)
114 fun boolean "" = true
115 | boolean "true" = true
116 | boolean "false" = false
117 | boolean s = error ("Bad boolean value: " ^ quote s);
122 (case Library.read_int ss of (i, []) => i
123 | _ => error ("Bad integer value: " ^ quote s));
124 in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end;
132 Parse.xname -- Scan.optional (Parse.$$$ "=" |-- Parse.!!! Parse.xname) "";
135 Scan.optional (Parse.$$$ "[" |-- Parse.!!! (Parse.enum "," property --| Parse.$$$ "]")) [];
140 Parse.!!! (Parse.position Parse.xname -- properties -- Args.parse --| Scan.ahead Parse.eof)
141 >> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos)));
148 val modes = Unsynchronized.ref ([]: string list);
150 fun eval_antiquote lex state (txt, pos) =
152 fun expand (Antiquote.Text ss) = Symbol_Pos.content ss
153 | expand (Antiquote.Antiq (ss, (pos, _))) =
154 let val (opts, src) = Token.read_antiq lex antiq (ss, pos) in
155 options opts (fn () => command src state) (); (*preview errors!*)
156 Print_Mode.with_modes (! modes @ Latex.modes)
157 (Output.no_warnings_CRITICAL (options opts (fn () => command src state))) ()
159 | expand (Antiquote.Open _) = ""
160 | expand (Antiquote.Close _) = "";
161 val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);
163 if Toplevel.is_toplevel state andalso not (forall Antiquote.is_text ants) then
164 error ("Unknown context -- cannot expand document antiquotations" ^ Position.str_of pos)
165 else implode (map expand ants)
170 (** present theory source **)
172 (*NB: arranging white space around command spans is a black art.*)
174 (* presentation tokens *)
178 | BasicToken of Token.T
179 | MarkupToken of string * (string * Position.T)
180 | MarkupEnvToken of string * (string * Position.T)
181 | VerbatimToken of string * Position.T;
183 fun output_token lex state =
184 let val eval = eval_antiquote lex state in
186 | BasicToken tok => Latex.output_basic tok
187 | MarkupToken (cmd, txt) => Latex.output_markup cmd (eval txt)
188 | MarkupEnvToken (cmd, txt) => Latex.output_markup_env cmd (eval txt)
189 | VerbatimToken txt => Latex.output_verbatim (eval txt)
192 fun basic_token pred (BasicToken tok) = pred tok
193 | basic_token _ _ = false;
195 val improper_token = basic_token (not o Token.is_proper);
196 val comment_token = basic_token Token.is_comment;
197 val blank_token = basic_token Token.is_blank;
198 val newline_token = basic_token Token.is_newline;
203 type command = string * Position.T * string list; (*name, position, tags*)
204 type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*)
206 datatype span = Span of command * (source * source * source * source) * bool;
208 fun make_span cmd src =
210 fun take_newline (tok :: toks) =
211 if newline_token (fst tok) then ([tok], toks, true)
212 else ([], tok :: toks, false)
213 | take_newline [] = ([], [], false);
214 val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) =
216 |> take_prefix (improper_token o fst)
217 ||>> take_suffix (improper_token o fst)
218 ||>> take_prefix (comment_token o fst)
220 in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end;
227 fun err_bad_nesting pos =
228 error ("Bad nesting of commands in presentation" ^ pos);
230 fun edge which f (x: string option, y) =
232 else (case which (x, y) of NONE => I | SOME txt => Buffer.add (f txt));
234 val begin_tag = edge #2 Latex.begin_tag;
235 val end_tag = edge #1 Latex.end_tag;
236 fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e;
237 fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e;
241 fun present_span lex default_tags span state state'
242 (tag_stack, active_tag, newline, buffer, present_cont) =
244 val present = fold (fn (tok, (flag, 0)) =>
245 Buffer.add (output_token lex state' tok)
249 val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span;
251 val (tag, tags) = tag_stack;
252 val tag' = try hd (fold Keyword.update_tags cmd_tags (the_list tag));
255 if is_some tag' then tag'
256 else if cmd_name = "end" andalso not (Toplevel.is_toplevel state') then NONE
257 else try hd (default_tags cmd_name);
258 val edge = (active_tag, active_tag');
261 if is_none active_tag' then span_newline else newline;
263 val nesting = Toplevel.level state' - Toplevel.level state;
265 if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack
266 else if nesting >= 0 then (tag', replicate nesting tag @ tags)
268 (case drop (~ nesting - 1) tags of
269 tgs :: tgss => (tgs, tgss)
270 | [] => err_bad_nesting (Position.str_of cmd_pos));
275 |> close_delim (fst present_cont) edge
277 |> open_delim (present (#1 srcs)) edge
279 |> present (#2 srcs);
281 if newline then (present (#3 srcs), present (#4 srcs))
282 else (I, present (#3 srcs) #> present (#4 srcs));
283 in (tag_stack', active_tag', newline', buffer', present_cont') end;
285 fun present_trailer ((_, tags), active_tag, _, buffer, present_cont) =
286 if not (null tags) then err_bad_nesting " at end of theory"
289 |> end_tag (active_tag, NONE)
290 |> close_delim (fst present_cont) (active_tag, NONE)
298 datatype markup = Markup | MarkupEnv | Verbatim;
303 Scan.one Token.is_blank -- Scan.many Token.is_comment -- Scan.one Token.is_proper;
305 val is_improper = not o (Token.is_proper orf Token.is_begin_ignore orf Token.is_end_ignore);
306 val improper = Scan.many is_improper;
307 val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));
308 val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));
310 val opt_newline = Scan.option (Scan.one Token.is_newline);
313 Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore
315 Scan.depend (fn d => Scan.one Token.is_end_ignore --|
316 (if d = 0 then Scan.fail_with (K "Bad nesting of meta-comments") else opt_newline)
319 val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end);
322 Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |--
323 Parse.!!! (improper |-- Parse.xname --| (improper -- Parse.$$$ ")")));
327 fun present_thy lex default_tags is_markup command_results src =
331 val ignored = Scan.state --| ignore
332 >> (fn d => (NONE, (NoToken, ("", d))));
334 fun markup mark mk flag = Scan.peek (fn d =>
337 (Scan.one (Token.is_kind Token.Command andf is_markup mark o Token.content_of)) --
339 Parse.!!!! ((improper -- locale -- improper) |-- Parse.doc_source --| improper_end)
340 >> (fn (((tok, pos), tags), txt) =>
341 let val name = Token.content_of tok
342 in (SOME (name, pos, tags), (mk (name, txt), (flag, d))) end));
344 val command = Scan.peek (fn d =>
345 Parse.position (Scan.one (Token.is_kind Token.Command)) --
347 >> (fn ((tok, pos), tags) =>
348 let val name = Token.content_of tok
349 in (SOME (name, pos, tags), (BasicToken tok, (Latex.markup_false, d))) end));
351 val cmt = Scan.peek (fn d =>
352 Parse.$$$ "--" |-- Parse.!!!! (improper |-- Parse.doc_source)
353 >> (fn txt => (NONE, (MarkupToken ("cmt", txt), ("", d)))));
355 val other = Scan.peek (fn d =>
356 Parse.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d)))));
360 markup Markup MarkupToken Latex.markup_true ||
361 markup MarkupEnv MarkupEnvToken Latex.markup_true ||
362 markup Verbatim (VerbatimToken o #2) "" ||
363 command || cmt || other;
368 val is_eof = fn (_, (BasicToken x, _)) => Token.is_eof x | _ => false;
369 val stopper = Scan.stopper (K (NONE, (BasicToken Token.eof, ("", 0)))) is_eof;
371 val cmd = Scan.one (is_some o fst);
372 val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;
374 val comments = Scan.many (comment_token o fst o snd);
375 val blank = Scan.one (blank_token o fst o snd);
376 val newline = Scan.one (newline_token o fst o snd);
378 Scan.option (newline -- comments) --
379 Scan.option (newline -- comments) --
380 Scan.option (blank -- comments) -- cmd;
383 Scan.repeat non_cmd -- cmd --
384 Scan.repeat (Scan.unless before_cmd non_cmd) --
385 Scan.option (newline >> (single o snd))
386 >> (fn (((toks1, (cmd, tok2)), toks3), tok4) =>
387 make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4))));
391 |> Source.filter (not o Token.is_semicolon)
392 |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk token)) NONE
393 |> Source.source stopper (Scan.error (Scan.bulk span)) NONE
397 (* present commands *)
399 fun present_command tr span st st' =
400 Toplevel.setmp_thread_position tr (present_span lex default_tags span st st');
403 | present st (((tr, st'), span) :: rest) = present_command tr span st st' #> present st' rest;
405 if length command_results = length spans then
406 ((NONE, []), NONE, true, Buffer.empty, (I, I))
407 |> present Toplevel.toplevel (command_results ~~ spans)
409 else error "Messed-up outer syntax for presentation"
416 (** setup default output **)
421 [("show_types", setmp_CRITICAL Syntax.show_types o boolean),
422 ("show_sorts", setmp_CRITICAL Syntax.show_sorts o boolean),
423 ("show_structs", setmp_CRITICAL show_structs o boolean),
424 ("show_question_marks", setmp_CRITICAL show_question_marks o boolean),
425 ("long_names", setmp_CRITICAL Name_Space.long_names o boolean),
426 ("short_names", setmp_CRITICAL Name_Space.short_names o boolean),
427 ("unique_names", setmp_CRITICAL Name_Space.unique_names o boolean),
428 ("eta_contract", setmp_CRITICAL Syntax.eta_contract o boolean),
429 ("display", setmp_CRITICAL display o boolean),
430 ("break", setmp_CRITICAL break o boolean),
431 ("quotes", setmp_CRITICAL quotes o boolean),
432 ("mode", fn s => fn f => Print_Mode.with_modes [s] f),
433 ("margin", setmp_CRITICAL Pretty.margin_default o integer),
434 ("indent", setmp_CRITICAL indent o integer),
435 ("source", setmp_CRITICAL source o boolean),
436 ("goals_limit", setmp_CRITICAL goals_limit o integer)];
439 (* basic pretty printing *)
442 if ! display then s else Symbol.strip_blanks s;
444 val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines;
446 fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;
448 fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;
450 fun pretty_term_style ctxt (style, t) =
451 pretty_term ctxt (style t);
453 fun pretty_thm_style ctxt (style, th) =
454 pretty_term ctxt (style (Thm.full_prop_of th));
456 fun pretty_term_typ ctxt (style, t) =
458 in pretty_term ctxt (Type_Infer.constrain (Term.fastype_of t') t') end;
460 fun pretty_term_typeof ctxt (style, t) =
461 Syntax.pretty_typ ctxt (Term.fastype_of (style t));
463 fun pretty_const ctxt c =
465 val t = Const (c, Consts.type_scheme (ProofContext.consts_of ctxt) c)
466 handle TYPE (msg, _, _) => error msg;
467 val ([t'], _) = Variable.import_terms true [t] ctxt;
468 in pretty_term ctxt t' end;
470 fun pretty_abbrev ctxt s =
472 val t = Syntax.parse_term ctxt s |> singleton (ProofContext.standard_infer_types ctxt);
473 fun err () = error ("Abbreviated constant expected: " ^ Syntax.string_of_term ctxt t);
474 val (head, args) = Term.strip_comb t;
475 val (c, T) = Term.dest_Const head handle TERM _ => err ();
476 val (U, u) = Consts.the_abbreviation (ProofContext.consts_of ctxt) c
477 handle TYPE _ => err ();
478 val t' = Term.betapplys (Envir.expand_atom T (U, u), args);
479 val eq = Logic.mk_equals (t, t');
480 val ctxt' = Variable.auto_fixes eq ctxt;
481 in ProofContext.pretty_term_abbrev ctxt' eq end;
483 fun pretty_prf full ctxt = Proof_Syntax.pretty_proof_of ctxt full;
485 fun pretty_theory ctxt name =
486 (Theory.requires (ProofContext.theory_of ctxt) name "presentation"; Pretty.str name);
491 val str_of_source = space_implode " " o map Token.unparse o #2 o #1 o Args.dest_src;
493 fun maybe_pretty_source pretty src xs =
494 map pretty xs (*always pretty in order to exhibit errors!*)
495 |> (if ! source then K [pretty_text (str_of_source src)] else I);
499 |> (if ! quotes then map Pretty.quote else I)
500 |> (if ! display then
501 map (Output.output o Pretty.string_of o Pretty.indent (! indent))
502 #> space_implode "\\isasep\\isanewline%\n"
503 #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
505 map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of))
506 #> space_implode "\\isasep\\isanewline%\n"
507 #> enclose "\\isa{" "}");
511 (** concrete antiquotations **)
517 fun basic_entities name scan pretty = antiquotation name scan
518 (fn {source, context, ...} => output o maybe_pretty_source (pretty context) source);
520 fun basic_entities_style name scan pretty = antiquotation name scan
521 (fn {source, context, ...} => fn (style, xs) =>
522 output (maybe_pretty_source (fn x => pretty context (style, x)) source xs));
524 fun basic_entity name scan = basic_entities name (scan >> single);
528 val _ = basic_entities_style "thm" (Term_Style.parse -- Attrib.thms) pretty_thm_style;
529 val _ = basic_entity "prop" (Term_Style.parse -- Args.prop) pretty_term_style;
530 val _ = basic_entity "term" (Term_Style.parse -- Args.term) pretty_term_style;
531 val _ = basic_entity "term_type" (Term_Style.parse -- Args.term) pretty_term_typ;
532 val _ = basic_entity "typeof" (Term_Style.parse -- Args.term) pretty_term_typeof;
533 val _ = basic_entity "const" (Args.const_proper false) pretty_const;
534 val _ = basic_entity "abbrev" (Scan.lift Args.name_source) pretty_abbrev;
535 val _ = basic_entity "typ" Args.typ_abbrev Syntax.pretty_typ;
536 val _ = basic_entity "text" (Scan.lift Args.name) (K pretty_text);
537 val _ = basic_entities "prf" Attrib.thms (pretty_prf false);
538 val _ = basic_entities "full_prf" Attrib.thms (pretty_prf true);
539 val _ = basic_entity "theory" (Scan.lift Args.name) pretty_theory;
541 val _ = basic_entities_style "thm_style" (Term_Style.parse_bare -- Attrib.thms) pretty_thm_style;
542 val _ = basic_entity "term_style" (Term_Style.parse_bare -- Args.term) pretty_term_style;
551 fun proof_state state =
552 (case try Toplevel.proof_of state of
554 | _ => error "No proof state");
556 fun goal_state name main_goal = antiquotation name (Scan.succeed ())
557 (fn {state, ...} => fn () => output
558 [Pretty.chunks (Proof.pretty_goals main_goal (proof_state state))]);
562 val _ = goal_state "goals" true;
563 val _ = goal_state "subgoals" false;
570 val _ = Keyword.keyword "by";
572 val _ = antiquotation "lemma"
573 (Args.prop -- Scan.lift (Args.$$$ "by" |-- Method.parse -- Scan.option Method.parse))
574 (fn {source, context, ...} => fn (prop, methods) =>
577 (case Args.dest_src source of ((a, arg :: _), pos) => Args.src ((a, [arg]), pos));
579 |> Proof.theorem NONE (K I) [[(prop, [])]]
580 |> Proof.global_terminal_proof methods;
581 in output (maybe_pretty_source (pretty_term context) prop_src [prop]) end);
588 fun ml_text name ml = antiquotation name (Scan.lift Args.name_source_position)
589 (fn {context, ...} => fn (txt, pos) =>
590 (ML_Context.eval_in (SOME context) false pos (ml pos txt);
591 Symbol_Pos.content (Symbol_Pos.explode (txt, pos))
592 |> (if ! quotes then quote else I)
593 |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
596 #> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|")
597 #> space_implode "\\isasep\\isanewline%\n")));
599 fun ml_enclose bg en pos txt =
600 ML_Lex.read Position.none bg @ ML_Lex.read pos txt @ ML_Lex.read Position.none en;
604 val _ = ml_text "ML" (ml_enclose "fn _ => (" ");");
605 val _ = ml_text "ML_type" (ml_enclose "val _ = NONE : (" ") option;");
606 val _ = ml_text "ML_struct" (ml_enclose "functor XXX() = struct structure XX = " " end;");
608 val _ = ml_text "ML_functor" (* FIXME formal treatment of functor name (!?) *)
610 ML_Lex.read Position.none ("ML_Env.check_functor " ^
611 ML_Syntax.print_string (Symbol_Pos.content (Symbol_Pos.explode (txt, pos)))));
613 val _ = ml_text "ML_text" (K (K []));