1(*  Title:      Pure/Thy/thy_output.ML
2    Author:     Makarius
3
4Theory document output.
5*)
6
7signature THY_OUTPUT =
8sig
9  val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text list
10  val check_comments: Proof.context -> Symbol_Pos.T list -> unit
11  val output_token: Proof.context -> Token.T -> Latex.text list
12  val output_source: Proof.context -> string -> Latex.text list
13  type segment = {span: Command_Span.span, command: Toplevel.transition, state: Toplevel.state}
14  val present_thy: Options.T -> theory -> segment list -> Latex.text list
15  val pretty_term: Proof.context -> term -> Pretty.T
16  val pretty_thm: Proof.context -> thm -> Pretty.T
17  val lines: Latex.text list -> Latex.text list
18  val items: Latex.text list -> Latex.text list
19  val isabelle: Proof.context -> Latex.text list -> Latex.text
20  val isabelle_typewriter: Proof.context -> Latex.text list -> Latex.text
21  val typewriter: Proof.context -> string -> Latex.text
22  val verbatim: Proof.context -> string -> Latex.text
23  val source: Proof.context -> Token.src -> Latex.text
24  val pretty: Proof.context -> Pretty.T -> Latex.text
25  val pretty_source: Proof.context -> Token.src -> Pretty.T -> Latex.text
26  val pretty_items: Proof.context -> Pretty.T list -> Latex.text
27  val pretty_items_source: Proof.context -> Token.src -> Pretty.T list -> Latex.text
28  val antiquotation_pretty:
29    binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
30  val antiquotation_pretty_source:
31    binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
32  val antiquotation_raw:
33    binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
34  val antiquotation_verbatim:
35    binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
36end;
37
38structure Thy_Output: THY_OUTPUT =
39struct
40
41(* output document source *)
42
43val output_symbols = single o Latex.symbols_output;
44
45fun output_comment ctxt (kind, syms) =
46  (case kind of
47    Comment.Comment =>
48      Input.cartouche_content syms
49      |> output_document (ctxt |> Config.put Document_Antiquotation.thy_output_display false)
50          {markdown = false}
51      |> Latex.enclose_body "%\n\\isamarkupcmt{" "%\n}"
52  | Comment.Cancel =>
53      Symbol_Pos.cartouche_content syms
54      |> output_symbols
55      |> Latex.enclose_body "%\n\\isamarkupcancel{" "}"
56  | Comment.Latex =>
57      [Latex.symbols (Symbol_Pos.cartouche_content syms)])
58and output_comment_document ctxt (comment, syms) =
59  (case comment of
60    SOME kind => output_comment ctxt (kind, syms)
61  | NONE => [Latex.symbols syms])
62and output_document_text ctxt syms =
63  Comment.read_body syms |> maps (output_comment_document ctxt)
64and output_document ctxt {markdown} source =
65  let
66    val pos = Input.pos_of source;
67    val syms = Input.source_explode source;
68
69    val output_antiquotes =
70      maps (Document_Antiquotation.evaluate (output_document_text ctxt) ctxt);
71
72    fun output_line line =
73      (if Markdown.line_is_item line then [Latex.string "\\item "] else []) @
74        output_antiquotes (Markdown.line_content line);
75
76    fun output_block (Markdown.Par lines) =
77          Latex.block (separate (Latex.string "\n") (map (Latex.block o output_line) lines))
78      | output_block (Markdown.List {kind, body, ...}) =
79          Latex.environment_block (Markdown.print_kind kind) (output_blocks body)
80    and output_blocks blocks = separate (Latex.string "\n\n") (map output_block blocks);
81  in
82    if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then []
83    else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms
84    then
85      let
86        val ants = Antiquote.parse_comments pos syms;
87        val reports = Antiquote.antiq_reports ants;
88        val blocks = Markdown.read_antiquotes ants;
89        val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks);
90      in output_blocks blocks end
91    else
92      let
93        val ants = Antiquote.parse_comments pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms);
94        val reports = Antiquote.antiq_reports ants;
95        val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants);
96      in output_antiquotes ants end
97  end;
98
99
100(* output tokens with formal comments *)
101
102local
103
104val output_symbols_antiq =
105  (fn Antiquote.Text syms => output_symbols syms
106    | Antiquote.Control {name = (name, _), body, ...} =>
107        Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) ::
108          output_symbols body
109    | Antiquote.Antiq {body, ...} =>
110        Latex.enclose_body "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (output_symbols body));
111
112fun output_comment_symbols ctxt {antiq} (comment, syms) =
113  (case (comment, antiq) of
114    (NONE, false) => output_symbols syms
115  | (NONE, true) =>
116      Antiquote.parse_comments (#1 (Symbol_Pos.range syms)) syms
117      |> maps output_symbols_antiq
118  | (SOME comment, _) => output_comment ctxt (comment, syms));
119
120fun output_body ctxt antiq bg en syms =
121  Comment.read_body syms
122  |> maps (output_comment_symbols ctxt {antiq = antiq})
123  |> Latex.enclose_body bg en;
124
125in
126
127fun output_token ctxt tok =
128  let
129    fun output antiq bg en =
130      output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok));
131  in
132    (case Token.kind_of tok of
133      Token.Comment NONE => []
134    | Token.Command => output false "\\isacommand{" "}"
135    | Token.Keyword =>
136        if Symbol.is_ascii_identifier (Token.content_of tok)
137        then output false "\\isakeyword{" "}"
138        else output false "" ""
139    | Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}"
140    | Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}"
141    | Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}"
142    | Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}"
143    | _ => output false "" "")
144  end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok));
145
146fun output_source ctxt s =
147  output_body ctxt false "" "" (Symbol_Pos.explode (s, Position.none));
148
149fun check_comments ctxt =
150  Comment.read_body #> List.app (fn (comment, syms) =>
151    let
152      val pos = #1 (Symbol_Pos.range syms);
153      val _ =
154        comment |> Option.app (fn kind =>
155          Context_Position.reports ctxt (map (pair pos) (Markup.cartouche :: Comment.markups kind)));
156      val _ = output_comment_symbols ctxt {antiq = false} (comment, syms);
157    in if comment = SOME Comment.Comment then check_comments ctxt syms else () end);
158
159end;
160
161
162
163(** present theory source **)
164
165(*NB: arranging white space around command spans is a black art*)
166
167val is_white = Token.is_space orf Token.is_informal_comment;
168val is_black = not o is_white;
169
170val is_white_comment = Token.is_informal_comment;
171val is_black_comment = Token.is_formal_comment;
172
173
174(* presentation tokens *)
175
176datatype token =
177    Ignore_Token
178  | Basic_Token of Token.T
179  | Markup_Token of string * Input.source
180  | Markup_Env_Token of string * Input.source
181  | Raw_Token of Input.source;
182
183fun basic_token pred (Basic_Token tok) = pred tok
184  | basic_token _ _ = false;
185
186val white_token = basic_token is_white;
187val white_comment_token = basic_token is_white_comment;
188val blank_token = basic_token Token.is_blank;
189val newline_token = basic_token Token.is_newline;
190
191fun present_token ctxt tok =
192  (case tok of
193    Ignore_Token => []
194  | Basic_Token tok => output_token ctxt tok
195  | Markup_Token (cmd, source) =>
196      Latex.enclose_body ("%\n\\isamarkup" ^ cmd ^ "{") "%\n}\n"
197        (output_document ctxt {markdown = false} source)
198  | Markup_Env_Token (cmd, source) =>
199      [Latex.environment_block ("isamarkup" ^ cmd) (output_document ctxt {markdown = true} source)]
200  | Raw_Token source =>
201      Latex.string "%\n" :: output_document ctxt {markdown = true} source @ [Latex.string "\n"]);
202
203
204(* command spans *)
205
206type command = string * Position.T * string list;   (*name, position, tags*)
207type source = (token * (string * int)) list;        (*token, markup flag, meta-comment depth*)
208
209datatype span = Span of command * (source * source * source * source) * bool;
210
211fun make_span cmd src =
212  let
213    fun chop_newline (tok :: toks) =
214          if newline_token (fst tok) then ([tok], toks, true)
215          else ([], tok :: toks, false)
216      | chop_newline [] = ([], [], false);
217    val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) =
218      src
219      |> chop_prefix (white_token o fst)
220      ||>> chop_suffix (white_token o fst)
221      ||>> chop_prefix (white_comment_token o fst)
222      ||> chop_newline;
223  in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end;
224
225
226(* present spans *)
227
228local
229
230fun err_bad_nesting pos =
231  error ("Bad nesting of commands in presentation" ^ pos);
232
233fun edge which f (x: string option, y) =
234  if x = y then I
235  else (case which (x, y) of NONE => I | SOME txt => cons (Latex.string (f txt)));
236
237val begin_tag = edge #2 Latex.begin_tag;
238val end_tag = edge #1 Latex.end_tag;
239fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e;
240fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e;
241
242fun read_tag s =
243  (case space_explode "%" s of
244    ["", b] => (SOME b, NONE)
245  | [a, b] => (NONE, SOME (a, b))
246  | _ => error ("Bad document_tags specification: " ^ quote s));
247
248in
249
250fun make_command_tag options keywords =
251  let
252    val document_tags =
253      map read_tag (space_explode "," (Options.string options \<^system_option>\<open>document_tags\<close>));
254    val document_tags_default = map_filter #1 document_tags;
255    val document_tags_command = map_filter #2 document_tags;
256  in
257    fn {cmd_name, cmd_tags, tag, active_tag} => fn state => fn state' =>
258      let
259        val tag' = try hd (fold (update (op =)) cmd_tags (the_list tag));
260
261        val keyword_tags =
262          if cmd_name = "end" andalso Toplevel.is_end_theory state' then ["theory"]
263          else Keyword.command_tags keywords cmd_name;
264        val command_tags =
265          the_list (AList.lookup (op =) document_tags_command cmd_name) @
266          keyword_tags @ document_tags_default;
267
268        val active_tag' =
269          if is_some tag' then tag'
270          else
271            (case command_tags of
272              default_tag :: _ => SOME default_tag
273            | [] =>
274                if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state
275                then active_tag
276                else NONE);
277      in {tag' = tag', active_tag' = active_tag'} end
278  end;
279
280fun present_span thy command_tag span state state'
281    (tag_stack, active_tag, newline, latex, present_cont) =
282  let
283    val ctxt' =
284      Toplevel.presentation_context state'
285        handle Toplevel.UNDEF => Proof_Context.get_global thy Context.PureN;
286    val present = fold (fn (tok, (flag, 0)) =>
287        fold cons (present_token ctxt' tok)
288        #> cons (Latex.string flag)
289      | _ => I);
290
291    val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span;
292
293    val (tag, tags) = tag_stack;
294    val {tag', active_tag'} =
295      command_tag {cmd_name = cmd_name, cmd_tags = cmd_tags, tag = tag, active_tag = active_tag}
296        state state';
297    val edge = (active_tag, active_tag');
298
299    val nesting = Toplevel.level state' - Toplevel.level state;
300
301    val newline' =
302      if is_none active_tag' then span_newline else newline;
303
304    val tag_stack' =
305      if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack
306      else if nesting >= 0 then (tag', replicate nesting tag @ tags)
307      else
308        (case drop (~ nesting - 1) tags of
309          tg :: tgs => (tg, tgs)
310        | [] => err_bad_nesting (Position.here cmd_pos));
311
312    val latex' =
313      latex
314      |> end_tag edge
315      |> close_delim (fst present_cont) edge
316      |> snd present_cont
317      |> open_delim (present (#1 srcs)) edge
318      |> begin_tag edge
319      |> present (#2 srcs);
320    val present_cont' =
321      if newline then (present (#3 srcs), present (#4 srcs))
322      else (I, present (#3 srcs) #> present (#4 srcs));
323  in (tag_stack', active_tag', newline', latex', present_cont') end;
324
325fun present_trailer ((_, tags), active_tag, _, latex, present_cont) =
326  if not (null tags) then err_bad_nesting " at end of theory"
327  else
328    latex
329    |> end_tag (active_tag, NONE)
330    |> close_delim (fst present_cont) (active_tag, NONE)
331    |> snd present_cont;
332
333end;
334
335
336(* present_thy *)
337
338local
339
340val markup_true = "\\isamarkuptrue%\n";
341val markup_false = "\\isamarkupfalse%\n";
342
343val space_proper =
344  Scan.one Token.is_blank -- Scan.many is_white_comment -- Scan.one is_black;
345
346val is_improper = not o (is_black orf Token.is_begin_ignore orf Token.is_end_ignore);
347val improper = Scan.many is_improper;
348val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));
349val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));
350
351val opt_newline = Scan.option (Scan.one Token.is_newline);
352
353val ignore =
354  Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore
355    >> pair (d + 1)) ||
356  Scan.depend (fn d => Scan.one Token.is_end_ignore --|
357    (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline)
358    >> pair (d - 1));
359
360val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end);
361
362val locale =
363  Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |--
364    Parse.!!! (improper |-- Parse.name --| (improper -- Parse.$$$ ")")));
365
366in
367
368type segment = {span: Command_Span.span, command: Toplevel.transition, state: Toplevel.state};
369
370fun present_thy options thy (segments: segment list) =
371  let
372    val keywords = Thy_Header.get_keywords thy;
373
374
375    (* tokens *)
376
377    val ignored = Scan.state --| ignore
378      >> (fn d => (NONE, (Ignore_Token, ("", d))));
379
380    fun markup pred mk flag = Scan.peek (fn d =>
381      improper |--
382        Parse.position (Scan.one (fn tok =>
383          Token.is_command tok andalso pred keywords (Token.content_of tok))) --
384      Scan.repeat tag --
385      Parse.!!!! ((improper -- locale -- improper) |-- Parse.document_source --| improper_end)
386      >> (fn (((tok, pos'), tags), source) =>
387        let val name = Token.content_of tok
388        in (SOME (name, pos', tags), (mk (name, source), (flag, d))) end));
389
390    val command = Scan.peek (fn d =>
391      Scan.optional (Scan.one Token.is_command_modifier ::: improper) [] --
392      Scan.one Token.is_command -- Scan.repeat tag
393      >> (fn ((cmd_mod, cmd), tags) =>
394        map (fn tok => (NONE, (Basic_Token tok, ("", d)))) cmd_mod @
395          [(SOME (Token.content_of cmd, Token.pos_of cmd, tags),
396            (Basic_Token cmd, (markup_false, d)))]));
397
398    val cmt = Scan.peek (fn d =>
399      Scan.one is_black_comment >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));
400
401    val other = Scan.peek (fn d =>
402       Parse.not_eof >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));
403
404    val tokens =
405      (ignored ||
406        markup Keyword.is_document_heading Markup_Token markup_true ||
407        markup Keyword.is_document_body Markup_Env_Token markup_true ||
408        markup Keyword.is_document_raw (Raw_Token o #2) "") >> single ||
409      command ||
410      (cmt || other) >> single;
411
412
413    (* spans *)
414
415    val is_eof = fn (_, (Basic_Token x, _)) => Token.is_eof x | _ => false;
416    val stopper = Scan.stopper (K (NONE, (Basic_Token Token.eof, ("", 0)))) is_eof;
417
418    val cmd = Scan.one (is_some o fst);
419    val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;
420
421    val white_comments = Scan.many (white_comment_token o fst o snd);
422    val blank = Scan.one (blank_token o fst o snd);
423    val newline = Scan.one (newline_token o fst o snd);
424    val before_cmd =
425      Scan.option (newline -- white_comments) --
426      Scan.option (newline -- white_comments) --
427      Scan.option (blank -- white_comments) -- cmd;
428
429    val span =
430      Scan.repeat non_cmd -- cmd --
431        Scan.repeat (Scan.unless before_cmd non_cmd) --
432        Scan.option (newline >> (single o snd))
433      >> (fn (((toks1, (cmd, tok2)), toks3), tok4) =>
434          make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4))));
435
436    val spans = segments
437      |> maps (Command_Span.content o #span)
438      |> drop_suffix Token.is_space
439      |> Source.of_list
440      |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat))
441      |> Source.source stopper (Scan.error (Scan.bulk span))
442      |> Source.exhaust;
443
444    val command_results =
445      segments |> map_filter (fn {command, state, ...} =>
446        if Toplevel.is_ignored command then NONE else SOME (command, state));
447
448
449    (* present commands *)
450
451    val command_tag = make_command_tag options keywords;
452
453    fun present_command tr span st st' =
454      Toplevel.setmp_thread_position tr (present_span thy command_tag span st st');
455
456    fun present _ [] = I
457      | present st ((span, (tr, st')) :: rest) = present_command tr span st st' #> present st' rest;
458  in
459    if length command_results = length spans then
460      ((NONE, []), NONE, true, [], (I, I))
461      |> present Toplevel.toplevel (spans ~~ command_results)
462      |> present_trailer
463      |> rev
464    else error "Messed-up outer syntax for presentation"
465  end;
466
467end;
468
469
470
471(** standard output operations **)
472
473(* pretty printing *)
474
475fun pretty_term ctxt t =
476  Syntax.pretty_term (Variable.auto_fixes t ctxt) t;
477
478fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;
479
480
481(* default output *)
482
483val lines = separate (Latex.string "\\isanewline%\n");
484val items = separate (Latex.string "\\isasep\\isanewline%\n");
485
486fun isabelle ctxt body =
487  if Config.get ctxt Document_Antiquotation.thy_output_display
488  then Latex.environment_block "isabelle" body
489  else Latex.block (Latex.enclose_body "\\isa{" "}" body);
490
491fun isabelle_typewriter ctxt body =
492  if Config.get ctxt Document_Antiquotation.thy_output_display
493  then Latex.environment_block "isabellett" body
494  else Latex.block (Latex.enclose_body "\\isatt{" "}" body);
495
496fun typewriter ctxt s =
497  isabelle_typewriter ctxt [Latex.string (Latex.output_ascii s)];
498
499fun verbatim ctxt =
500  if Config.get ctxt Document_Antiquotation.thy_output_display
501  then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt
502  else split_lines #> map (typewriter ctxt) #> lines #> Latex.block;
503
504fun source ctxt =
505  Token.args_of_src
506  #> map (Token.unparse #> Document_Antiquotation.prepare_lines ctxt)
507  #> space_implode " "
508  #> output_source ctxt
509  #> isabelle ctxt;
510
511fun pretty ctxt =
512  Document_Antiquotation.output ctxt #> Latex.string #> single #> isabelle ctxt;
513
514fun pretty_source ctxt src prt =
515  if Config.get ctxt Document_Antiquotation.thy_output_source
516  then source ctxt src else pretty ctxt prt;
517
518fun pretty_items ctxt =
519  map (Document_Antiquotation.output ctxt #> Latex.string) #> items #> isabelle ctxt;
520
521fun pretty_items_source ctxt src prts =
522  if Config.get ctxt Document_Antiquotation.thy_output_source
523  then source ctxt src else pretty_items ctxt prts;
524
525
526(* antiquotation variants *)
527
528fun antiquotation_pretty name scan f =
529  Document_Antiquotation.setup name scan
530    (fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x));
531
532fun antiquotation_pretty_source name scan f =
533  Document_Antiquotation.setup name scan
534    (fn {context = ctxt, source = src, argument = x} => pretty_source ctxt src (f ctxt x));
535
536fun antiquotation_raw name scan f =
537  Document_Antiquotation.setup name scan
538    (fn {context = ctxt, argument = x, ...} => f ctxt x);
539
540fun antiquotation_verbatim name scan f =
541  antiquotation_raw name scan (fn ctxt => verbatim ctxt o f ctxt);
542
543end;
544