| author | wenzelm | 
| Thu, 05 Aug 2021 20:24:42 +0200 | |
| changeset 74129 | c3794f56a2e2 | 
| parent 73780 | 466fae6bf22e | 
| child 74373 | 6e4093927dbb | 
| permissions | -rw-r--r-- | 
| 73761 | 1 | (* Title: Pure/Thy/document_output.ML | 
| 67386 | 2 | Author: Makarius | 
| 22124 | 3 | |
| 67386 | 4 | Theory document output. | 
| 22124 | 5 | *) | 
| 6 | ||
| 73761 | 7 | signature DOCUMENT_OUTPUT = | 
| 22124 | 8 | sig | 
| 73752 | 9 | val document_reports: Input.source -> Position.report list | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 10 |   val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text
 | 
| 67381 | 11 | val check_comments: Proof.context -> Symbol_Pos.T list -> unit | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 12 | val output_token: Proof.context -> Token.T -> Latex.text | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 13 | val output_source: Proof.context -> string -> Latex.text | 
| 73687 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
71507diff
changeset | 14 | type segment = | 
| 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
71507diff
changeset | 15 |     {span: Command_Span.span, command: Toplevel.transition,
 | 
| 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
71507diff
changeset | 16 | prev_state: Toplevel.state, state: Toplevel.state} | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 17 | val present_thy: Options.T -> theory -> segment list -> Latex.text | 
| 28644 | 18 | val pretty_term: Proof.context -> term -> Pretty.T | 
| 19 | val pretty_thm: Proof.context -> thm -> Pretty.T | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 20 | val isabelle: Proof.context -> Latex.text -> Latex.text | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 21 | val isabelle_typewriter: Proof.context -> Latex.text -> Latex.text | 
| 67463 | 22 | val typewriter: Proof.context -> string -> Latex.text | 
| 23 | val verbatim: Proof.context -> string -> Latex.text | |
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 24 |   val source: Proof.context -> {embedded: bool} -> Token.src -> Latex.text
 | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 25 | val pretty: Proof.context -> Pretty.T -> Latex.text | 
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 26 |   val pretty_source: Proof.context -> {embedded: bool} -> Token.src -> Pretty.T -> Latex.text
 | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 27 | val pretty_items: Proof.context -> Pretty.T list -> Latex.text | 
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 28 |   val pretty_items_source: Proof.context -> {embedded: bool} -> Token.src ->
 | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 29 | Pretty.T list -> Latex.text | 
| 67463 | 30 | val antiquotation_pretty: | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 31 | binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 32 | val antiquotation_pretty_embedded: | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 33 | binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory | 
| 67463 | 34 | val antiquotation_pretty_source: | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 35 | binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 36 | val antiquotation_pretty_source_embedded: | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 37 | binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory | 
| 67463 | 38 | val antiquotation_raw: | 
| 39 | binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 40 | val antiquotation_raw_embedded: | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 41 | binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory | 
| 67463 | 42 | val antiquotation_verbatim: | 
| 43 | binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 44 | val antiquotation_verbatim_embedded: | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 45 | binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory | 
| 22124 | 46 | end; | 
| 47 | ||
| 73761 | 48 | structure Document_Output: DOCUMENT_OUTPUT = | 
| 22124 | 49 | struct | 
| 50 | ||
| 67569 
5d71b114e7a3
avoid proliferation of language_document reports;
 wenzelm parents: 
67567diff
changeset | 51 | (* output document source *) | 
| 56548 | 52 | |
| 73752 | 53 | fun document_reports txt = | 
| 54 | let val pos = Input.pos_of txt in | |
| 55 | [(pos, Markup.language_document (Input.is_delimited txt)), | |
| 56 | (pos, Markup.plain_text)] | |
| 57 | end; | |
| 58 | ||
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 59 | fun output_comment ctxt (kind, syms) = | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 60 | (case kind of | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 61 | Comment.Comment => | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 62 | Input.cartouche_content syms | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 63 | |> output_document (ctxt |> Config.put Document_Antiquotation.thy_output_display false) | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 64 |           {markdown = false}
 | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 65 |       |> Latex.enclose_text "%\n\\isamarkupcmt{" "%\n}"
 | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 66 | | Comment.Cancel => | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 67 | Symbol_Pos.cartouche_content syms | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 68 | |> Latex.symbols_output | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 69 |       |> Latex.enclose_text "%\n\\isamarkupcancel{" "}"
 | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 70 | | Comment.Latex => Latex.symbols (Symbol_Pos.cartouche_content syms) | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 71 | | Comment.Marker => []) | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 72 | and output_comment_document ctxt (comment, syms) = | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 73 | (case comment of | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 74 | SOME kind => output_comment ctxt (kind, syms) | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 75 | | NONE => Latex.symbols syms) | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 76 | and output_document_text ctxt syms = | 
| 67572 | 77 | Comment.read_body syms |> maps (output_comment_document ctxt) | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 78 | and output_document ctxt {markdown} source =
 | 
| 61457 | 79 | let | 
| 80 | val pos = Input.pos_of source; | |
| 81 | val syms = Input.source_explode source; | |
| 82 | ||
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 83 | val output_antiquotes = | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 84 | maps (Document_Antiquotation.evaluate (output_document_text ctxt) ctxt); | 
| 61036 
f6f2959bed67
clarified language context, e.g. relevant for symbols;
 wenzelm parents: 
60100diff
changeset | 85 | |
| 61461 | 86 | fun output_line line = | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 87 | (if Markdown.line_is_item line then Latex.string "\\item " else []) @ | 
| 61461 | 88 | output_antiquotes (Markdown.line_content line); | 
| 89 | ||
| 67194 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 wenzelm parents: 
67191diff
changeset | 90 | fun output_block (Markdown.Par lines) = | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 91 | separate (XML.Text "\n") (map (Latex.block o output_line) lines) | 
| 61459 
5f2ddeb15c06
clarified nesting of paragraphs: indentation is taken into account more uniformly;
 wenzelm parents: 
61458diff
changeset | 92 |       | output_block (Markdown.List {kind, body, ...}) =
 | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 93 | Latex.environment_text (Markdown.print_kind kind) (output_blocks body) | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 94 | and output_blocks blocks = | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 95 | separate (XML.Text "\n\n") (map (Latex.block o output_block) blocks); | 
| 67184 | 96 | in | 
| 67381 | 97 | if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then [] | 
| 67184 | 98 | else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms | 
| 99 | then | |
| 100 | let | |
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 101 | val ants = Antiquote.parse_comments pos syms; | 
| 67184 | 102 | val reports = Antiquote.antiq_reports ants; | 
| 103 | val blocks = Markdown.read_antiquotes ants; | |
| 67381 | 104 | val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks); | 
| 67186 
a58bbe66ac81
avoid excessive whitespace between antiquotations and text;
 wenzelm parents: 
67184diff
changeset | 105 | in output_blocks blocks end | 
| 67184 | 106 | else | 
| 107 | let | |
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 108 | val ants = Antiquote.parse_comments pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms); | 
| 67184 | 109 | val reports = Antiquote.antiq_reports ants; | 
| 67381 | 110 | val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants); | 
| 67186 
a58bbe66ac81
avoid excessive whitespace between antiquotations and text;
 wenzelm parents: 
67184diff
changeset | 111 | in output_antiquotes ants end | 
| 67184 | 112 | end; | 
| 48918 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 wenzelm parents: 
47005diff
changeset | 113 | |
| 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 wenzelm parents: 
47005diff
changeset | 114 | |
| 67425 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 wenzelm parents: 
67421diff
changeset | 115 | (* output tokens with formal comments *) | 
| 22124 | 116 | |
| 67353 | 117 | local | 
| 118 | ||
| 67358 | 119 | val output_symbols_antiq = | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 120 | (fn Antiquote.Text syms => Latex.symbols_output syms | 
| 67353 | 121 |     | Antiquote.Control {name = (name, _), body, ...} =>
 | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 122 | Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) @ | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 123 | Latex.symbols_output body | 
| 67353 | 124 |     | Antiquote.Antiq {body, ...} =>
 | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 125 |         Latex.enclose_text "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (Latex.symbols_output body));
 | 
| 67353 | 126 | |
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 127 | fun output_comment_symbols ctxt {antiq} (comment, syms) =
 | 
| 67425 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 wenzelm parents: 
67421diff
changeset | 128 | (case (comment, antiq) of | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 129 | (NONE, false) => Latex.symbols_output syms | 
| 67425 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 wenzelm parents: 
67421diff
changeset | 130 | | (NONE, true) => | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 131 | Antiquote.parse_comments (#1 (Symbol_Pos.range syms)) syms | 
| 67413 | 132 | |> maps output_symbols_antiq | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 133 | | (SOME comment, _) => output_comment ctxt (comment, syms)); | 
| 67372 | 134 | |
| 67381 | 135 | fun output_body ctxt antiq bg en syms = | 
| 67572 | 136 | Comment.read_body syms | 
| 137 |   |> maps (output_comment_symbols ctxt {antiq = antiq})
 | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 138 | |> Latex.enclose_text bg en; | 
| 68194 | 139 | |
| 140 | in | |
| 141 | ||
| 142 | fun output_token ctxt tok = | |
| 67353 | 143 | let | 
| 67378 | 144 | fun output antiq bg en = | 
| 67381 | 145 | output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok)); | 
| 67378 | 146 | in | 
| 147 | (case Token.kind_of tok of | |
| 67439 
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
 wenzelm parents: 
67429diff
changeset | 148 | Token.Comment NONE => [] | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 149 | | Token.Comment (SOME Comment.Marker) => [] | 
| 67378 | 150 |     | Token.Command => output false "\\isacommand{" "}"
 | 
| 151 | | Token.Keyword => | |
| 152 | if Symbol.is_ascii_identifier (Token.content_of tok) | |
| 153 |         then output false "\\isakeyword{" "}"
 | |
| 154 | else output false "" "" | |
| 155 |     | Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}"
 | |
| 156 |     | Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}"
 | |
| 157 |     | Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}"
 | |
| 158 |     | Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}"
 | |
| 159 | | _ => output false "" "") | |
| 160 | end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok)); | |
| 67353 | 161 | |
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 162 | fun output_source ctxt s = | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 163 | output_body ctxt false "" "" (Symbol_Pos.explode (s, Position.none)); | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 164 | |
| 67381 | 165 | fun check_comments ctxt = | 
| 67572 | 166 | Comment.read_body #> List.app (fn (comment, syms) => | 
| 67388 
5fc0b0c9a735
clarified markup: more like outer syntax side-comment;
 wenzelm parents: 
67386diff
changeset | 167 | let | 
| 
5fc0b0c9a735
clarified markup: more like outer syntax side-comment;
 wenzelm parents: 
67386diff
changeset | 168 | val pos = #1 (Symbol_Pos.range syms); | 
| 
5fc0b0c9a735
clarified markup: more like outer syntax side-comment;
 wenzelm parents: 
67386diff
changeset | 169 | val _ = | 
| 67429 | 170 | comment |> Option.app (fn kind => | 
| 69966 | 171 | Context_Position.reports ctxt (map (pair pos) (Comment.kind_markups kind))); | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67570diff
changeset | 172 |       val _ = output_comment_symbols ctxt {antiq = false} (comment, syms);
 | 
| 67425 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 wenzelm parents: 
67421diff
changeset | 173 | in if comment = SOME Comment.Comment then check_comments ctxt syms else () end); | 
| 67377 
143665524d8e
check formal comments recursively, within arbitrary cartouches (unknown sublanguages);
 wenzelm parents: 
67375diff
changeset | 174 | |
| 67353 | 175 | end; | 
| 176 | ||
| 177 | ||
| 67356 
ba226b87c69e
output token content with formal comments and antiquotations;
 wenzelm parents: 
67354diff
changeset | 178 | |
| 
ba226b87c69e
output token content with formal comments and antiquotations;
 wenzelm parents: 
67354diff
changeset | 179 | (** present theory source **) | 
| 
ba226b87c69e
output token content with formal comments and antiquotations;
 wenzelm parents: 
67354diff
changeset | 180 | |
| 22124 | 181 | (* presentation tokens *) | 
| 182 | ||
| 183 | datatype token = | |
| 71507 | 184 | Ignore | 
| 185 | | Token of Token.T | |
| 186 | | Heading of string * Input.source | |
| 187 | | Body of string * Input.source | |
| 188 | | Raw of Input.source; | |
| 22124 | 189 | |
| 71507 | 190 | fun token_with pred (Token tok) = pred tok | 
| 191 | | token_with _ _ = false; | |
| 22124 | 192 | |
| 71507 | 193 | val white_token = token_with Document_Source.is_white; | 
| 194 | val white_comment_token = token_with Document_Source.is_white_comment; | |
| 195 | val blank_token = token_with Token.is_blank; | |
| 196 | val newline_token = token_with Token.is_newline; | |
| 22124 | 197 | |
| 67381 | 198 | fun present_token ctxt tok = | 
| 61457 | 199 | (case tok of | 
| 71507 | 200 | Ignore => [] | 
| 201 | | Token tok => output_token ctxt tok | |
| 202 | | Heading (cmd, source) => | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 203 |       Latex.enclose_text ("%\n\\isamarkup" ^ cmd ^ "{") "%\n}\n"
 | 
| 67569 
5d71b114e7a3
avoid proliferation of language_document reports;
 wenzelm parents: 
67567diff
changeset | 204 |         (output_document ctxt {markdown = false} source)
 | 
| 71507 | 205 | | Body (cmd, source) => | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 206 |       Latex.environment_text ("isamarkup" ^ cmd) (output_document ctxt {markdown = true} source)
 | 
| 71507 | 207 | | Raw source => | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 208 |       Latex.string "%\n" @ output_document ctxt {markdown = true} source @ Latex.string "\n");
 | 
| 61455 | 209 | |
| 210 | ||
| 22124 | 211 | (* command spans *) | 
| 212 | ||
| 69887 | 213 | type command = string * Position.T; (*name, position*) | 
| 22124 | 214 | type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*) | 
| 215 | ||
| 216 | datatype span = Span of command * (source * source * source * source) * bool; | |
| 217 | ||
| 218 | fun make_span cmd src = | |
| 219 | let | |
| 67522 | 220 | fun chop_newline (tok :: toks) = | 
| 22124 | 221 | if newline_token (fst tok) then ([tok], toks, true) | 
| 222 | else ([], tok :: toks, false) | |
| 67522 | 223 | | chop_newline [] = ([], [], false); | 
| 22124 | 224 | val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) = | 
| 225 | src | |
| 67522 | 226 | |> chop_prefix (white_token o fst) | 
| 227 | ||>> chop_suffix (white_token o fst) | |
| 228 | ||>> chop_prefix (white_comment_token o fst) | |
| 229 | ||> chop_newline; | |
| 22124 | 230 | in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end; | 
| 231 | ||
| 232 | ||
| 233 | (* present spans *) | |
| 234 | ||
| 235 | local | |
| 236 | ||
| 70130 | 237 | fun err_bad_nesting here = | 
| 238 |   error ("Bad nesting of commands in presentation" ^ here);
 | |
| 22124 | 239 | |
| 240 | fun edge which f (x: string option, y) = | |
| 241 | if x = y then I | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 242 | else (case which (x, y) of NONE => I | SOME txt => fold cons (Latex.string (f txt))); | 
| 22124 | 243 | |
| 244 | val begin_tag = edge #2 Latex.begin_tag; | |
| 245 | val end_tag = edge #1 Latex.end_tag; | |
| 246 | fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e; | |
| 247 | fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e; | |
| 248 | ||
| 70134 | 249 | fun document_tag cmd_pos state state' tagging_stack = | 
| 70130 | 250 | let | 
| 251 | val ctxt' = Toplevel.presentation_context state'; | |
| 252 | val nesting = Toplevel.level state' - Toplevel.level state; | |
| 253 | ||
| 70134 | 254 | val (tagging, taggings) = tagging_stack; | 
| 255 | val (tag', tagging') = Document_Source.update_tagging ctxt' tagging; | |
| 256 | ||
| 257 | val tagging_stack' = | |
| 258 | if nesting = 0 andalso not (Toplevel.is_proof state) then tagging_stack | |
| 259 | else if nesting >= 0 then (tagging', replicate nesting tagging @ taggings) | |
| 70130 | 260 | else | 
| 70134 | 261 | (case drop (~ nesting - 1) taggings of | 
| 70130 | 262 | tg :: tgs => (tg, tgs) | 
| 263 | | [] => err_bad_nesting (Position.here cmd_pos)); | |
| 70134 | 264 | in (tag', tagging_stack') end; | 
| 70129 | 265 | |
| 68507 | 266 | fun read_tag s = | 
| 68512 | 267 | (case space_explode "%" s of | 
| 268 | ["", b] => (SOME b, NONE) | |
| 68510 
795f39bfe4e1
simplified: allow only command names, with dummy for default;
 wenzelm parents: 
68509diff
changeset | 269 | | [a, b] => (NONE, SOME (a, b)) | 
| 68507 | 270 |   | _ => error ("Bad document_tags specification: " ^ quote s));
 | 
| 271 | ||
| 22124 | 272 | in | 
| 273 | ||
| 68506 | 274 | fun make_command_tag options keywords = | 
| 275 | let | |
| 68507 | 276 | val document_tags = | 
| 277 | map read_tag (space_explode "," (Options.string options \<^system_option>\<open>document_tags\<close>)); | |
| 278 | val document_tags_default = map_filter #1 document_tags; | |
| 68510 
795f39bfe4e1
simplified: allow only command names, with dummy for default;
 wenzelm parents: 
68509diff
changeset | 279 | val document_tags_command = map_filter #2 document_tags; | 
| 68506 | 280 | in | 
| 70129 | 281 | fn cmd_name => fn state => fn state' => fn active_tag => | 
| 68506 | 282 | let | 
| 68509 | 283 | val keyword_tags = | 
| 284 | if cmd_name = "end" andalso Toplevel.is_end_theory state' then ["theory"] | |
| 285 | else Keyword.command_tags keywords cmd_name; | |
| 68508 | 286 | val command_tags = | 
| 287 | the_list (AList.lookup (op =) document_tags_command cmd_name) @ | |
| 288 | keyword_tags @ document_tags_default; | |
| 68506 | 289 | val active_tag' = | 
| 70129 | 290 | (case command_tags of | 
| 291 | default_tag :: _ => SOME default_tag | |
| 292 | | [] => | |
| 293 | if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state | |
| 294 | then active_tag | |
| 295 | else NONE); | |
| 296 | in active_tag' end | |
| 68506 | 297 | end; | 
| 298 | ||
| 69892 | 299 | fun present_span command_tag span state state' | 
| 70134 | 300 | (tagging_stack, active_tag, newline, latex, present_cont) = | 
| 22124 | 301 | let | 
| 69892 | 302 | val ctxt' = Toplevel.presentation_context state'; | 
| 22124 | 303 | val present = fold (fn (tok, (flag, 0)) => | 
| 67381 | 304 | fold cons (present_token ctxt' tok) | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 305 | #> fold cons (Latex.string flag) | 
| 22124 | 306 | | _ => I); | 
| 307 | ||
| 69887 | 308 | val Span ((cmd_name, cmd_pos), srcs, span_newline) = span; | 
| 22124 | 309 | |
| 70134 | 310 | val (tag', tagging_stack') = document_tag cmd_pos state state' tagging_stack; | 
| 311 | val active_tag' = | |
| 312 | if is_some tag' then Option.map #1 tag' | |
| 313 | else command_tag cmd_name state state' active_tag; | |
| 68506 | 314 | val edge = (active_tag, active_tag'); | 
| 22124 | 315 | |
| 316 | val newline' = | |
| 317 | if is_none active_tag' then span_newline else newline; | |
| 318 | ||
| 67194 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 wenzelm parents: 
67191diff
changeset | 319 | val latex' = | 
| 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 wenzelm parents: 
67191diff
changeset | 320 | latex | 
| 22124 | 321 | |> end_tag edge | 
| 322 | |> close_delim (fst present_cont) edge | |
| 323 | |> snd present_cont | |
| 324 | |> open_delim (present (#1 srcs)) edge | |
| 325 | |> begin_tag edge | |
| 326 | |> present (#2 srcs); | |
| 327 | val present_cont' = | |
| 328 | if newline then (present (#3 srcs), present (#4 srcs)) | |
| 329 | else (I, present (#3 srcs) #> present (#4 srcs)); | |
| 70134 | 330 | in (tagging_stack', active_tag', newline', latex', present_cont') end; | 
| 22124 | 331 | |
| 67194 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 wenzelm parents: 
67191diff
changeset | 332 | fun present_trailer ((_, tags), active_tag, _, latex, present_cont) = | 
| 22124 | 333 | if not (null tags) then err_bad_nesting " at end of theory" | 
| 334 | else | |
| 67194 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 wenzelm parents: 
67191diff
changeset | 335 | latex | 
| 22124 | 336 | |> end_tag (active_tag, NONE) | 
| 337 | |> close_delim (fst present_cont) (active_tag, NONE) | |
| 338 | |> snd present_cont; | |
| 339 | ||
| 340 | end; | |
| 341 | ||
| 342 | ||
| 28427 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 343 | (* present_thy *) | 
| 22124 | 344 | |
| 345 | local | |
| 346 | ||
| 61455 | 347 | val markup_true = "\\isamarkuptrue%\n"; | 
| 348 | val markup_false = "\\isamarkupfalse%\n"; | |
| 349 | ||
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 350 | val opt_newline = Scan.option (Scan.one Token.is_newline); | 
| 22124 | 351 | |
| 352 | val ignore = | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 353 | Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore | 
| 22124 | 354 | >> pair (d + 1)) || | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 355 | Scan.depend (fn d => Scan.one Token.is_end_ignore --| | 
| 43947 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 wenzelm parents: 
43564diff
changeset | 356 | (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline) | 
| 22124 | 357 | >> pair (d - 1)); | 
| 358 | ||
| 359 | val locale = | |
| 69876 | 360 |   Scan.option ((Parse.$$$ "(" -- Document_Source.improper -- Parse.$$$ "in") |--
 | 
| 361 | Parse.!!! | |
| 362 | (Document_Source.improper |-- Parse.name --| (Document_Source.improper -- Parse.$$$ ")"))); | |
| 22124 | 363 | |
| 364 | in | |
| 365 | ||
| 73687 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
71507diff
changeset | 366 | type segment = | 
| 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
71507diff
changeset | 367 |   {span: Command_Span.span, command: Toplevel.transition,
 | 
| 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
71507diff
changeset | 368 | prev_state: Toplevel.state, state: Toplevel.state}; | 
| 68178 | 369 | |
| 68506 | 370 | fun present_thy options thy (segments: segment list) = | 
| 22124 | 371 | let | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 372 | val keywords = Thy_Header.get_keywords thy; | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 373 | |
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 374 | |
| 22124 | 375 | (* tokens *) | 
| 376 | ||
| 377 | val ignored = Scan.state --| ignore | |
| 71507 | 378 |       >> (fn d => (NONE, (Ignore, ("", d))));
 | 
| 22124 | 379 | |
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58992diff
changeset | 380 | fun markup pred mk flag = Scan.peek (fn d => | 
| 69876 | 381 | Document_Source.improper |-- | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58992diff
changeset | 382 | Parse.position (Scan.one (fn tok => | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58992diff
changeset | 383 | Token.is_command tok andalso pred keywords (Token.content_of tok))) -- | 
| 69887 | 384 | (Document_Source.annotation |-- | 
| 385 | Parse.!!!! ((Document_Source.improper -- locale -- Document_Source.improper) |-- | |
| 386 | Parse.document_source --| Document_Source.improper_end)) | |
| 387 | >> (fn ((tok, pos'), source) => | |
| 388 | let val name = Token.content_of tok | |
| 389 | in (SOME (name, pos'), (mk (name, source), (flag, d))) end)); | |
| 22124 | 390 | |
| 391 | val command = Scan.peek (fn d => | |
| 69876 | 392 | Scan.optional (Scan.one Token.is_command_modifier ::: Document_Source.improper) [] -- | 
| 69887 | 393 | Scan.one Token.is_command --| Document_Source.annotation | 
| 394 | >> (fn (cmd_mod, cmd) => | |
| 71507 | 395 |         map (fn tok => (NONE, (Token tok, ("", d)))) cmd_mod @
 | 
| 69887 | 396 | [(SOME (Token.content_of cmd, Token.pos_of cmd), | 
| 71507 | 397 | (Token cmd, (markup_false, d)))])); | 
| 22124 | 398 | |
| 399 | val cmt = Scan.peek (fn d => | |
| 71507 | 400 |       Scan.one Document_Source.is_black_comment >> (fn tok => (NONE, (Token tok, ("", d)))));
 | 
| 22124 | 401 | |
| 402 | val other = Scan.peek (fn d => | |
| 71507 | 403 |        Parse.not_eof >> (fn tok => (NONE, (Token tok, ("", d)))));
 | 
| 22124 | 404 | |
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59917diff
changeset | 405 | val tokens = | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59917diff
changeset | 406 | (ignored || | 
| 71507 | 407 | markup Keyword.is_document_heading Heading markup_true || | 
| 408 | markup Keyword.is_document_body Body markup_true || | |
| 409 | markup Keyword.is_document_raw (Raw o #2) "") >> single || | |
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59917diff
changeset | 410 | command || | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59917diff
changeset | 411 | (cmt || other) >> single; | 
| 22124 | 412 | |
| 413 | ||
| 414 | (* spans *) | |
| 415 | ||
| 71507 | 416 | val is_eof = fn (_, (Token x, _)) => Token.is_eof x | _ => false; | 
| 417 |     val stopper = Scan.stopper (K (NONE, (Token Token.eof, ("", 0)))) is_eof;
 | |
| 22124 | 418 | |
| 419 | val cmd = Scan.one (is_some o fst); | |
| 27732 | 420 | val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2; | 
| 22124 | 421 | |
| 67439 
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
 wenzelm parents: 
67429diff
changeset | 422 | val white_comments = Scan.many (white_comment_token o fst o snd); | 
| 22124 | 423 | val blank = Scan.one (blank_token o fst o snd); | 
| 424 | val newline = Scan.one (newline_token o fst o snd); | |
| 425 | val before_cmd = | |
| 67439 
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
 wenzelm parents: 
67429diff
changeset | 426 | Scan.option (newline -- white_comments) -- | 
| 
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
 wenzelm parents: 
67429diff
changeset | 427 | Scan.option (newline -- white_comments) -- | 
| 
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
 wenzelm parents: 
67429diff
changeset | 428 | Scan.option (blank -- white_comments) -- cmd; | 
| 22124 | 429 | |
| 430 | val span = | |
| 431 | Scan.repeat non_cmd -- cmd -- | |
| 432 | Scan.repeat (Scan.unless before_cmd non_cmd) -- | |
| 433 | Scan.option (newline >> (single o snd)) | |
| 434 | >> (fn (((toks1, (cmd, tok2)), toks3), tok4) => | |
| 435 | make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4)))); | |
| 436 | ||
| 68178 | 437 | val spans = segments | 
| 438 | |> maps (Command_Span.content o #span) | |
| 67522 | 439 | |> drop_suffix Token.is_space | 
| 57080 
0e5fa27d3293
strip trailing white space, to avoid notorious problems of jEdit with last line;
 wenzelm parents: 
56548diff
changeset | 440 | |> Source.of_list | 
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59917diff
changeset | 441 | |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat)) | 
| 58864 | 442 | |> Source.source stopper (Scan.error (Scan.bulk span)) | 
| 22124 | 443 | |> Source.exhaust; | 
| 28427 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 444 | |
| 68178 | 445 | val command_results = | 
| 68182 | 446 |       segments |> map_filter (fn {command, state, ...} =>
 | 
| 447 | if Toplevel.is_ignored command then NONE else SOME (command, state)); | |
| 68178 | 448 | |
| 28427 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 449 | |
| 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 450 | (* present commands *) | 
| 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 451 | |
| 68506 | 452 | val command_tag = make_command_tag options keywords; | 
| 67138 | 453 | |
| 28427 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 454 | fun present_command tr span st st' = | 
| 69892 | 455 | Toplevel.setmp_thread_position tr (present_span command_tag span st st'); | 
| 28427 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 456 | |
| 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 457 | fun present _ [] = I | 
| 68178 | 458 | | present st ((span, (tr, st')) :: rest) = present_command tr span st st' #> present st' rest; | 
| 22124 | 459 | in | 
| 28427 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 wenzelm parents: 
28273diff
changeset | 460 | if length command_results = length spans then | 
| 70134 | 461 | (([], []), NONE, true, [], (I, I)) | 
| 69886 | 462 | |> present (Toplevel.init_toplevel ()) (spans ~~ command_results) | 
| 22124 | 463 | |> present_trailer | 
| 67194 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 wenzelm parents: 
67191diff
changeset | 464 | |> rev | 
| 22124 | 465 | else error "Messed-up outer syntax for presentation" | 
| 466 | end; | |
| 467 | ||
| 468 | end; | |
| 469 | ||
| 470 | ||
| 471 | ||
| 67386 | 472 | (** standard output operations **) | 
| 22124 | 473 | |
| 67463 | 474 | (* pretty printing *) | 
| 22124 | 475 | |
| 67463 | 476 | fun pretty_term ctxt t = | 
| 70308 | 477 | Syntax.pretty_term (Proof_Context.augment t ctxt) t; | 
| 24920 | 478 | |
| 32898 
e871d897969c
term styles also cover antiquotations term_type and typeof
 haftmann parents: 
32890diff
changeset | 479 | fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; | 
| 
e871d897969c
term styles also cover antiquotations term_type and typeof
 haftmann parents: 
32890diff
changeset | 480 | |
| 22124 | 481 | |
| 30390 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 wenzelm parents: 
30381diff
changeset | 482 | (* default output *) | 
| 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 wenzelm parents: 
30381diff
changeset | 483 | |
| 67463 | 484 | fun isabelle ctxt body = | 
| 485 | if Config.get ctxt Document_Antiquotation.thy_output_display | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 486 | then Latex.environment_text "isabelle" body | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 487 |   else Latex.enclose_text "\\isa{" "}" body;
 | 
| 59175 
bf465f335e85
system option "pretty_margin" is superseded by "thy_output_margin";
 wenzelm parents: 
59067diff
changeset | 488 | |
| 67463 | 489 | fun isabelle_typewriter ctxt body = | 
| 490 | if Config.get ctxt Document_Antiquotation.thy_output_display | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 491 | then Latex.environment_text "isabellett" body | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 492 |   else Latex.enclose_text "\\isatt{" "}" body;
 | 
| 67463 | 493 | |
| 494 | fun typewriter ctxt s = | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 495 | isabelle_typewriter ctxt (Latex.string (Latex.output_ascii s)); | 
| 67463 | 496 | |
| 497 | fun verbatim ctxt = | |
| 498 | if Config.get ctxt Document_Antiquotation.thy_output_display | |
| 499 | then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt | |
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 500 | else split_lines #> map (typewriter ctxt #> Latex.block) #> separate (XML.Text "\\isanewline%\n"); | 
| 22124 | 501 | |
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 502 | fun token_source ctxt {embedded} tok =
 | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 503 | if Token.is_kind Token.Cartouche tok andalso embedded andalso | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 504 | not (Config.get ctxt Document_Antiquotation.thy_output_source_cartouche) | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 505 | then Token.content_of tok | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 506 | else Token.unparse tok; | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 507 | |
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 508 | fun is_source ctxt = | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 509 | Config.get ctxt Document_Antiquotation.thy_output_source orelse | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 510 | Config.get ctxt Document_Antiquotation.thy_output_source_cartouche; | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 511 | |
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 512 | fun source ctxt embedded = | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 513 | Token.args_of_src | 
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 514 | #> map (token_source ctxt embedded #> Document_Antiquotation.prepare_lines ctxt) | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 515 | #> space_implode " " | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 516 | #> output_source ctxt | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 517 | #> isabelle ctxt; | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67465diff
changeset | 518 | |
| 67463 | 519 | fun pretty ctxt = | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 520 | Document_Antiquotation.output ctxt #> Latex.string #> isabelle ctxt; | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 521 | |
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 522 | fun pretty_source ctxt embedded src prt = | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 523 | if is_source ctxt then source ctxt embedded src else pretty ctxt prt; | 
| 67463 | 524 | |
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 525 | fun pretty_items ctxt = | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 526 | map (Document_Antiquotation.output ctxt #> XML.Text) #> | 
| 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73761diff
changeset | 527 | separate (XML.Text "\\isasep\\isanewline%\n") #> isabelle ctxt; | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67475diff
changeset | 528 | |
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 529 | fun pretty_items_source ctxt embedded src prts = | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 530 | if is_source ctxt then source ctxt embedded src else pretty_items ctxt prts; | 
| 61473 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 wenzelm parents: 
61471diff
changeset | 531 | |
| 67465 | 532 | |
| 533 | (* antiquotation variants *) | |
| 534 | ||
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 535 | local | 
| 67463 | 536 | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 537 | fun gen_setup name embedded = | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 538 | if embedded | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 539 | then Document_Antiquotation.setup_embedded name | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 540 | else Document_Antiquotation.setup name; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 541 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 542 | fun gen_antiquotation_pretty name embedded scan f = | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 543 |   gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x));
 | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 544 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 545 | fun gen_antiquotation_pretty_source name embedded scan f = | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 546 | gen_setup name embedded scan | 
| 70122 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 547 |     (fn {context = ctxt, source = src, argument = x} =>
 | 
| 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 wenzelm parents: 
69966diff
changeset | 548 |       pretty_source ctxt {embedded = embedded} src (f ctxt x));
 | 
| 67463 | 549 | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 550 | fun gen_antiquotation_raw name embedded scan f = | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 551 |   gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => f ctxt x);
 | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 552 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 553 | fun gen_antiquotation_verbatim name embedded scan f = | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 554 | gen_antiquotation_raw name embedded scan (fn ctxt => verbatim ctxt o f ctxt); | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 555 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 556 | in | 
| 67463 | 557 | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 558 | fun antiquotation_pretty name = gen_antiquotation_pretty name false; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 559 | fun antiquotation_pretty_embedded name = gen_antiquotation_pretty name true; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 560 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 561 | fun antiquotation_pretty_source name = gen_antiquotation_pretty_source name false; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 562 | fun antiquotation_pretty_source_embedded name = gen_antiquotation_pretty_source name true; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 563 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 564 | fun antiquotation_raw name = gen_antiquotation_raw name false; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 565 | fun antiquotation_raw_embedded name = gen_antiquotation_raw name true; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 566 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 567 | fun antiquotation_verbatim name = gen_antiquotation_verbatim name false; | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 568 | fun antiquotation_verbatim_embedded name = gen_antiquotation_verbatim name true; | 
| 61491 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 wenzelm parents: 
61473diff
changeset | 569 | |
| 30390 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 wenzelm parents: 
30381diff
changeset | 570 | end; | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 571 | |
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
68512diff
changeset | 572 | end; |