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