author | wenzelm |
Sat, 03 Feb 2018 14:39:17 +0100 | |
changeset 67569 | 5d71b114e7a3 |
parent 67567 | 52e6ffa61e9c |
child 67570 | c1fe89e9a00b |
permissions | -rw-r--r-- |
22124 | 1 |
(* Title: Pure/Thy/thy_output.ML |
67386 | 2 |
Author: Makarius |
22124 | 3 |
|
67386 | 4 |
Theory document output. |
22124 | 5 |
*) |
6 |
||
7 |
signature THY_OUTPUT = |
|
8 |
sig |
|
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
9 |
val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text list |
67381 | 10 |
val check_comments: Proof.context -> Symbol_Pos.T list -> unit |
11 |
val check_token_comments: Proof.context -> Token.T -> unit |
|
12 |
val output_token: Proof.context -> Token.T -> Latex.text list |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
13 |
val output_source: Proof.context -> string -> Latex.text list |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
14 |
val present_thy: theory -> (Toplevel.transition * Toplevel.state) list -> |
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
15 |
Token.T list -> Latex.text list |
28644 | 16 |
val pretty_term: Proof.context -> term -> Pretty.T |
17 |
val pretty_thm: Proof.context -> thm -> Pretty.T |
|
67508 | 18 |
val lines: Latex.text list -> Latex.text list |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
19 |
val items: Latex.text list -> Latex.text list |
67463 | 20 |
val isabelle: Proof.context -> Latex.text list -> Latex.text |
21 |
val isabelle_typewriter: Proof.context -> Latex.text list -> Latex.text |
|
22 |
val typewriter: Proof.context -> string -> Latex.text |
|
23 |
val verbatim: Proof.context -> string -> Latex.text |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
24 |
val source: Proof.context -> Token.src -> Latex.text |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
25 |
val pretty: Proof.context -> Pretty.T -> Latex.text |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
26 |
val pretty_source: Proof.context -> Token.src -> Pretty.T -> Latex.text |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
27 |
val pretty_items: Proof.context -> Pretty.T list -> Latex.text |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
28 |
val pretty_items_source: Proof.context -> Token.src -> Pretty.T list -> Latex.text |
67463 | 29 |
val antiquotation_pretty: |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
30 |
binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory |
67463 | 31 |
val antiquotation_pretty_source: |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
32 |
binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory |
67463 | 33 |
val antiquotation_raw: |
34 |
binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory |
|
35 |
val antiquotation_verbatim: |
|
36 |
binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory |
|
22124 | 37 |
end; |
38 |
||
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
39 |
structure Thy_Output: THY_OUTPUT = |
22124 | 40 |
struct |
41 |
||
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
42 |
(* output document source *) |
56548 | 43 |
|
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
44 |
fun output_document ctxt {markdown} source = |
61457 | 45 |
let |
46 |
val pos = Input.pos_of source; |
|
47 |
val syms = Input.source_explode source; |
|
48 |
||
67463 | 49 |
val output_antiquotes = maps (Document_Antiquotation.evaluate ctxt); |
61036
f6f2959bed67
clarified language context, e.g. relevant for symbols;
wenzelm
parents:
60100
diff
changeset
|
50 |
|
61461 | 51 |
fun output_line line = |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
52 |
(if Markdown.line_is_item line then [Latex.string "\\item "] else []) @ |
61461 | 53 |
output_antiquotes (Markdown.line_content line); |
54 |
||
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
55 |
fun output_block (Markdown.Par lines) = |
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
56 |
Latex.block (separate (Latex.string "\n") (map (Latex.block o output_line) lines)) |
61459
5f2ddeb15c06
clarified nesting of paragraphs: indentation is taken into account more uniformly;
wenzelm
parents:
61458
diff
changeset
|
57 |
| output_block (Markdown.List {kind, body, ...}) = |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
58 |
Latex.environment_block (Markdown.print_kind kind) (output_blocks body) |
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
59 |
and output_blocks blocks = separate (Latex.string "\n\n") (map output_block blocks); |
67184 | 60 |
in |
67381 | 61 |
if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then [] |
67184 | 62 |
else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms |
63 |
then |
|
64 |
let |
|
65 |
val ants = Antiquote.parse pos syms; |
|
66 |
val reports = Antiquote.antiq_reports ants; |
|
67 |
val blocks = Markdown.read_antiquotes ants; |
|
67381 | 68 |
val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks); |
67186
a58bbe66ac81
avoid excessive whitespace between antiquotations and text;
wenzelm
parents:
67184
diff
changeset
|
69 |
in output_blocks blocks end |
67184 | 70 |
else |
71 |
let |
|
67463 | 72 |
val ants = Antiquote.parse pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms); |
67184 | 73 |
val reports = Antiquote.antiq_reports ants; |
67381 | 74 |
val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants); |
67186
a58bbe66ac81
avoid excessive whitespace between antiquotations and text;
wenzelm
parents:
67184
diff
changeset
|
75 |
in output_antiquotes ants end |
67184 | 76 |
end; |
48918
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
wenzelm
parents:
47005
diff
changeset
|
77 |
|
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
wenzelm
parents:
47005
diff
changeset
|
78 |
|
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
79 |
(* output tokens with formal comments *) |
22124 | 80 |
|
67353 | 81 |
local |
82 |
||
67461 | 83 |
val output_symbols = single o Latex.symbols_output; |
67356
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
84 |
|
67358 | 85 |
val output_symbols_antiq = |
86 |
(fn Antiquote.Text syms => output_symbols syms |
|
67353 | 87 |
| Antiquote.Control {name = (name, _), body, ...} => |
67374
5a049cf98438
clarified output (see also 909dcdec2122, 34d1913f0b20);
wenzelm
parents:
67372
diff
changeset
|
88 |
Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) :: |
67358 | 89 |
output_symbols body |
67353 | 90 |
| Antiquote.Antiq {body, ...} => |
67358 | 91 |
Latex.enclose_body "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (output_symbols body)); |
67353 | 92 |
|
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
93 |
fun output_symbols_comment ctxt {antiq} (comment, syms) = |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
94 |
(case (comment, antiq) of |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
95 |
(NONE, false) => output_symbols syms |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
96 |
| (NONE, true) => |
67413 | 97 |
Antiquote.parse (#1 (Symbol_Pos.range syms)) syms |
98 |
|> maps output_symbols_antiq |
|
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
99 |
| (SOME Comment.Comment, _) => |
67413 | 100 |
let |
67567 | 101 |
val source = Input.cartouche_content syms; |
67475
1a279ad4c6a4
disable "display" style in marginal (line) comment;
wenzelm
parents:
67474
diff
changeset
|
102 |
val ctxt' = ctxt |
1a279ad4c6a4
disable "display" style in marginal (line) comment;
wenzelm
parents:
67474
diff
changeset
|
103 |
|> Config.put Document_Antiquotation.thy_output_display false; |
67413 | 104 |
in |
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
105 |
output_document ctxt' {markdown = false} source |
67421
c4a10621c72e
allow TeX comment % in formal comment body, but avoid extra space (cf. d7c6054b2ab1);
wenzelm
parents:
67413
diff
changeset
|
106 |
|> Latex.enclose_body "%\n\\isamarkupcmt{" "%\n}" |
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
107 |
end |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
108 |
| (SOME Comment.Cancel, _) => |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
109 |
output_symbols (Symbol_Pos.cartouche_content syms) |
67429 | 110 |
|> Latex.enclose_body "%\n\\isamarkupcancel{" "}" |
111 |
| (SOME Comment.Latex, _) => |
|
67461 | 112 |
[Latex.symbols (Symbol_Pos.cartouche_content syms)]); |
67372 | 113 |
|
67353 | 114 |
in |
115 |
||
67381 | 116 |
fun output_body ctxt antiq bg en syms = |
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
117 |
(case Comment.read_body syms of |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
118 |
SOME res => maps (output_symbols_comment ctxt {antiq = antiq}) res |
67375 | 119 |
| NONE => output_symbols syms) |> Latex.enclose_body bg en |
67381 | 120 |
and output_token ctxt tok = |
67353 | 121 |
let |
67378 | 122 |
fun output antiq bg en = |
67381 | 123 |
output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok)); |
67378 | 124 |
in |
125 |
(case Token.kind_of tok of |
|
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
126 |
Token.Comment NONE => [] |
67378 | 127 |
| Token.Command => output false "\\isacommand{" "}" |
128 |
| Token.Keyword => |
|
129 |
if Symbol.is_ascii_identifier (Token.content_of tok) |
|
130 |
then output false "\\isakeyword{" "}" |
|
131 |
else output false "" "" |
|
132 |
| Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}" |
|
133 |
| Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}" |
|
134 |
| Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}" |
|
135 |
| Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}" |
|
136 |
| _ => output false "" "") |
|
137 |
end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok)); |
|
67353 | 138 |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
139 |
fun output_source ctxt s = |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
140 |
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:
67465
diff
changeset
|
141 |
|
67381 | 142 |
fun check_comments ctxt = |
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
143 |
Comment.read_body #> (Option.app o List.app) (fn (comment, syms) => |
67388
5fc0b0c9a735
clarified markup: more like outer syntax side-comment;
wenzelm
parents:
67386
diff
changeset
|
144 |
let |
5fc0b0c9a735
clarified markup: more like outer syntax side-comment;
wenzelm
parents:
67386
diff
changeset
|
145 |
val pos = #1 (Symbol_Pos.range syms); |
5fc0b0c9a735
clarified markup: more like outer syntax side-comment;
wenzelm
parents:
67386
diff
changeset
|
146 |
val _ = |
67429 | 147 |
comment |> Option.app (fn kind => |
67506
30233285270a
more markup: disable spell-checker for raw latex;
wenzelm
parents:
67505
diff
changeset
|
148 |
Context_Position.reports ctxt (map (pair pos) (Markup.cartouche :: Comment.markups kind))); |
67425
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
149 |
val _ = output_symbols_comment ctxt {antiq = false} (comment, syms); |
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
wenzelm
parents:
67421
diff
changeset
|
150 |
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:
67375
diff
changeset
|
151 |
|
67381 | 152 |
fun check_token_comments ctxt tok = |
153 |
check_comments ctxt (Input.source_explode (Token.input_of tok)); |
|
67377
143665524d8e
check formal comments recursively, within arbitrary cartouches (unknown sublanguages);
wenzelm
parents:
67375
diff
changeset
|
154 |
|
67353 | 155 |
end; |
156 |
||
157 |
||
67356
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
158 |
|
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
159 |
(** present theory source **) |
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
160 |
|
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
161 |
(*NB: arranging white space around command spans is a black art*) |
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
162 |
|
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
163 |
val is_white = Token.is_space orf Token.is_informal_comment; |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
164 |
val is_black = not o is_white; |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
165 |
|
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
166 |
val is_white_comment = Token.is_informal_comment; |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
167 |
val is_black_comment = Token.is_formal_comment; |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
168 |
|
67356
ba226b87c69e
output token content with formal comments and antiquotations;
wenzelm
parents:
67354
diff
changeset
|
169 |
|
22124 | 170 |
(* presentation tokens *) |
171 |
||
172 |
datatype token = |
|
67360 | 173 |
Ignore_Token |
59065 | 174 |
| Basic_Token of Token.T |
175 |
| Markup_Token of string * Input.source |
|
176 |
| Markup_Env_Token of string * Input.source |
|
61456 | 177 |
| Raw_Token of Input.source; |
22124 | 178 |
|
59065 | 179 |
fun basic_token pred (Basic_Token tok) = pred tok |
22124 | 180 |
| basic_token _ _ = false; |
181 |
||
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
182 |
val white_token = basic_token is_white; |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
183 |
val white_comment_token = basic_token is_white_comment; |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
184 |
val blank_token = basic_token Token.is_blank; |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
185 |
val newline_token = basic_token Token.is_newline; |
22124 | 186 |
|
67381 | 187 |
fun present_token ctxt tok = |
61457 | 188 |
(case tok of |
67360 | 189 |
Ignore_Token => [] |
67381 | 190 |
| Basic_Token tok => output_token ctxt tok |
67357 | 191 |
| Markup_Token (cmd, source) => |
192 |
Latex.enclose_body ("%\n\\isamarkup" ^ cmd ^ "{") "%\n}\n" |
|
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
193 |
(output_document ctxt {markdown = false} source) |
61457 | 194 |
| Markup_Env_Token (cmd, source) => |
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
195 |
[Latex.environment_block ("isamarkup" ^ cmd) (output_document ctxt {markdown = true} source)] |
67186
a58bbe66ac81
avoid excessive whitespace between antiquotations and text;
wenzelm
parents:
67184
diff
changeset
|
196 |
| Raw_Token source => |
67569
5d71b114e7a3
avoid proliferation of language_document reports;
wenzelm
parents:
67567
diff
changeset
|
197 |
Latex.string "%\n" :: output_document ctxt {markdown = true} source @ [Latex.string "\n"]); |
61455 | 198 |
|
199 |
||
22124 | 200 |
(* command spans *) |
201 |
||
202 |
type command = string * Position.T * string list; (*name, position, tags*) |
|
203 |
type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*) |
|
204 |
||
205 |
datatype span = Span of command * (source * source * source * source) * bool; |
|
206 |
||
207 |
fun make_span cmd src = |
|
208 |
let |
|
67522 | 209 |
fun chop_newline (tok :: toks) = |
22124 | 210 |
if newline_token (fst tok) then ([tok], toks, true) |
211 |
else ([], tok :: toks, false) |
|
67522 | 212 |
| chop_newline [] = ([], [], false); |
22124 | 213 |
val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) = |
214 |
src |
|
67522 | 215 |
|> chop_prefix (white_token o fst) |
216 |
||>> chop_suffix (white_token o fst) |
|
217 |
||>> chop_prefix (white_comment_token o fst) |
|
218 |
||> chop_newline; |
|
22124 | 219 |
in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end; |
220 |
||
221 |
||
222 |
(* present spans *) |
|
223 |
||
224 |
local |
|
225 |
||
226 |
fun err_bad_nesting pos = |
|
227 |
error ("Bad nesting of commands in presentation" ^ pos); |
|
228 |
||
229 |
fun edge which f (x: string option, y) = |
|
230 |
if x = y then I |
|
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
231 |
else (case which (x, y) of NONE => I | SOME txt => cons (Latex.string (f txt))); |
22124 | 232 |
|
233 |
val begin_tag = edge #2 Latex.begin_tag; |
|
234 |
val end_tag = edge #1 Latex.end_tag; |
|
235 |
fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e; |
|
236 |
fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e; |
|
237 |
||
238 |
in |
|
239 |
||
67381 | 240 |
fun present_span thy keywords document_tags span state state' |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
241 |
(tag_stack, active_tag, newline, latex, present_cont) = |
22124 | 242 |
let |
67381 | 243 |
val ctxt' = |
244 |
Toplevel.presentation_context state' |
|
245 |
handle Toplevel.UNDEF => Proof_Context.init_global (Context.get_theory thy Context.PureN); |
|
22124 | 246 |
val present = fold (fn (tok, (flag, 0)) => |
67381 | 247 |
fold cons (present_token ctxt' tok) |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
248 |
#> cons (Latex.string flag) |
22124 | 249 |
| _ => I); |
250 |
||
251 |
val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span; |
|
252 |
||
253 |
val (tag, tags) = tag_stack; |
|
46924 | 254 |
val tag' = try hd (fold (update (op =)) cmd_tags (the_list tag)); |
22124 | 255 |
|
58999
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
256 |
val nesting = Toplevel.level state' - Toplevel.level state; |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
257 |
|
22124 | 258 |
val active_tag' = |
259 |
if is_some tag' then tag' |
|
260 |
else if cmd_name = "end" andalso not (Toplevel.is_toplevel state') then NONE |
|
58999
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
261 |
else |
67138 | 262 |
(case Keyword.command_tags keywords cmd_name @ document_tags of |
58999
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
263 |
default_tag :: _ => SOME default_tag |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
264 |
| [] => |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
265 |
if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
266 |
then active_tag |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
267 |
else NONE); |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
268 |
|
22124 | 269 |
val edge = (active_tag, active_tag'); |
270 |
||
271 |
val newline' = |
|
272 |
if is_none active_tag' then span_newline else newline; |
|
273 |
||
274 |
val tag_stack' = |
|
275 |
if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack |
|
276 |
else if nesting >= 0 then (tag', replicate nesting tag @ tags) |
|
277 |
else |
|
33957 | 278 |
(case drop (~ nesting - 1) tags of |
58999
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
279 |
tg :: tgs => (tg, tgs) |
48992 | 280 |
| [] => err_bad_nesting (Position.here cmd_pos)); |
22124 | 281 |
|
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
282 |
val latex' = |
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
283 |
latex |
22124 | 284 |
|> end_tag edge |
285 |
|> close_delim (fst present_cont) edge |
|
286 |
|> snd present_cont |
|
287 |
|> open_delim (present (#1 srcs)) edge |
|
288 |
|> begin_tag edge |
|
289 |
|> present (#2 srcs); |
|
290 |
val present_cont' = |
|
291 |
if newline then (present (#3 srcs), present (#4 srcs)) |
|
292 |
else (I, present (#3 srcs) #> present (#4 srcs)); |
|
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
293 |
in (tag_stack', active_tag', newline', latex', present_cont') end; |
22124 | 294 |
|
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
295 |
fun present_trailer ((_, tags), active_tag, _, latex, present_cont) = |
22124 | 296 |
if not (null tags) then err_bad_nesting " at end of theory" |
297 |
else |
|
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
298 |
latex |
22124 | 299 |
|> end_tag (active_tag, NONE) |
300 |
|> close_delim (fst present_cont) (active_tag, NONE) |
|
301 |
|> snd present_cont; |
|
302 |
||
303 |
end; |
|
304 |
||
305 |
||
28427
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
306 |
(* present_thy *) |
22124 | 307 |
|
308 |
local |
|
309 |
||
61455 | 310 |
val markup_true = "\\isamarkuptrue%\n"; |
311 |
val markup_false = "\\isamarkupfalse%\n"; |
|
312 |
||
22124 | 313 |
val space_proper = |
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
314 |
Scan.one Token.is_blank -- Scan.many is_white_comment -- Scan.one is_black; |
22124 | 315 |
|
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
316 |
val is_improper = not o (is_black orf Token.is_begin_ignore orf Token.is_end_ignore); |
22124 | 317 |
val improper = Scan.many is_improper; |
318 |
val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper)); |
|
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
319 |
val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank)); |
22124 | 320 |
|
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
321 |
val opt_newline = Scan.option (Scan.one Token.is_newline); |
22124 | 322 |
|
323 |
val ignore = |
|
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
324 |
Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore |
22124 | 325 |
>> pair (d + 1)) || |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
326 |
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:
43564
diff
changeset
|
327 |
(if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline) |
22124 | 328 |
>> pair (d - 1)); |
329 |
||
36950 | 330 |
val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end); |
22124 | 331 |
|
332 |
val locale = |
|
36950 | 333 |
Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |-- |
62969 | 334 |
Parse.!!! (improper |-- Parse.name --| (improper -- Parse.$$$ ")"))); |
22124 | 335 |
|
336 |
in |
|
337 |
||
58928
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58923
diff
changeset
|
338 |
fun present_thy thy command_results toks = |
22124 | 339 |
let |
58928
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58923
diff
changeset
|
340 |
val keywords = Thy_Header.get_keywords thy; |
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58923
diff
changeset
|
341 |
|
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58923
diff
changeset
|
342 |
|
22124 | 343 |
(* tokens *) |
344 |
||
345 |
val ignored = Scan.state --| ignore |
|
67360 | 346 |
>> (fn d => (NONE, (Ignore_Token, ("", d)))); |
22124 | 347 |
|
58999
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
348 |
fun markup pred mk flag = Scan.peek (fn d => |
36950 | 349 |
improper |-- |
58999
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
350 |
Parse.position (Scan.one (fn tok => |
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents:
58992
diff
changeset
|
351 |
Token.is_command tok andalso pred keywords (Token.content_of tok))) -- |
22124 | 352 |
Scan.repeat tag -- |
51627
589daaf48dba
tuned signature -- agree with markup terminology;
wenzelm
parents:
51626
diff
changeset
|
353 |
Parse.!!!! ((improper -- locale -- improper) |-- Parse.document_source --| improper_end) |
59064 | 354 |
>> (fn (((tok, pos'), tags), source) => |
59065 | 355 |
let val name = Token.content_of tok |
356 |
in (SOME (name, pos', tags), (mk (name, source), (flag, d))) end)); |
|
22124 | 357 |
|
358 |
val command = Scan.peek (fn d => |
|
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
61748
diff
changeset
|
359 |
Scan.optional (Scan.one Token.is_command_modifier ::: improper) [] -- |
59924
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
360 |
Scan.one Token.is_command -- Scan.repeat tag |
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59924
diff
changeset
|
361 |
>> (fn ((cmd_mod, cmd), tags) => |
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59924
diff
changeset
|
362 |
map (fn tok => (NONE, (Basic_Token tok, ("", d)))) cmd_mod @ |
59924
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
363 |
[(SOME (Token.content_of cmd, Token.pos_of cmd, tags), |
61455 | 364 |
(Basic_Token cmd, (markup_false, d)))])); |
22124 | 365 |
|
366 |
val cmt = Scan.peek (fn d => |
|
67446 | 367 |
Scan.one is_black_comment >> (fn tok => (NONE, (Basic_Token tok, ("", d))))); |
22124 | 368 |
|
369 |
val other = Scan.peek (fn d => |
|
59065 | 370 |
Parse.not_eof >> (fn tok => (NONE, (Basic_Token tok, ("", d))))); |
22124 | 371 |
|
59924
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
372 |
val tokens = |
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
373 |
(ignored || |
61455 | 374 |
markup Keyword.is_document_heading Markup_Token markup_true || |
375 |
markup Keyword.is_document_body Markup_Env_Token markup_true || |
|
61456 | 376 |
markup Keyword.is_document_raw (Raw_Token o #2) "") >> single || |
59924
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
377 |
command || |
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
378 |
(cmt || other) >> single; |
22124 | 379 |
|
380 |
||
381 |
(* spans *) |
|
382 |
||
59065 | 383 |
val is_eof = fn (_, (Basic_Token x, _)) => Token.is_eof x | _ => false; |
384 |
val stopper = Scan.stopper (K (NONE, (Basic_Token Token.eof, ("", 0)))) is_eof; |
|
22124 | 385 |
|
386 |
val cmd = Scan.one (is_some o fst); |
|
27732 | 387 |
val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2; |
22124 | 388 |
|
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
389 |
val white_comments = Scan.many (white_comment_token o fst o snd); |
22124 | 390 |
val blank = Scan.one (blank_token o fst o snd); |
391 |
val newline = Scan.one (newline_token o fst o snd); |
|
392 |
val before_cmd = |
|
67439
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
393 |
Scan.option (newline -- white_comments) -- |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
394 |
Scan.option (newline -- white_comments) -- |
78759a7bd874
more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents:
67429
diff
changeset
|
395 |
Scan.option (blank -- white_comments) -- cmd; |
22124 | 396 |
|
397 |
val span = |
|
398 |
Scan.repeat non_cmd -- cmd -- |
|
399 |
Scan.repeat (Scan.unless before_cmd non_cmd) -- |
|
400 |
Scan.option (newline >> (single o snd)) |
|
401 |
>> (fn (((toks1, (cmd, tok2)), toks3), tok4) => |
|
402 |
make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4)))); |
|
403 |
||
57080
0e5fa27d3293
strip trailing white space, to avoid notorious problems of jEdit with last line;
wenzelm
parents:
56548
diff
changeset
|
404 |
val spans = toks |
67522 | 405 |
|> drop_suffix Token.is_space |
57080
0e5fa27d3293
strip trailing white space, to avoid notorious problems of jEdit with last line;
wenzelm
parents:
56548
diff
changeset
|
406 |
|> Source.of_list |
59924
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
wenzelm
parents:
59917
diff
changeset
|
407 |
|> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat)) |
58864 | 408 |
|> Source.source stopper (Scan.error (Scan.bulk span)) |
22124 | 409 |
|> Source.exhaust; |
28427
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
410 |
|
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
411 |
|
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
412 |
(* present commands *) |
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
413 |
|
67147 | 414 |
val document_tags = space_explode "," (Options.default_string \<^system_option>\<open>document_tags\<close>); |
67138 | 415 |
|
28427
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
416 |
fun present_command tr span st st' = |
67191 | 417 |
Toplevel.setmp_thread_position tr |
67381 | 418 |
(present_span thy keywords document_tags span st st'); |
28427
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
419 |
|
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
420 |
fun present _ [] = I |
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
421 |
| present st (((tr, st'), span) :: rest) = present_command tr span st st' #> present st' rest; |
22124 | 422 |
in |
28427
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
423 |
if length command_results = length spans then |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
424 |
((NONE, []), NONE, true, [], (I, I)) |
28427
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents:
28273
diff
changeset
|
425 |
|> present Toplevel.toplevel (command_results ~~ spans) |
22124 | 426 |
|> present_trailer |
67194
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents:
67191
diff
changeset
|
427 |
|> rev |
22124 | 428 |
else error "Messed-up outer syntax for presentation" |
429 |
end; |
|
430 |
||
431 |
end; |
|
432 |
||
433 |
||
434 |
||
67386 | 435 |
(** standard output operations **) |
22124 | 436 |
|
67463 | 437 |
(* pretty printing *) |
22124 | 438 |
|
67463 | 439 |
fun pretty_term ctxt t = |
440 |
Syntax.pretty_term (Variable.auto_fixes t ctxt) t; |
|
24920 | 441 |
|
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32890
diff
changeset
|
442 |
fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; |
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32890
diff
changeset
|
443 |
|
22124 | 444 |
|
30390
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
wenzelm
parents:
30381
diff
changeset
|
445 |
(* default output *) |
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
wenzelm
parents:
30381
diff
changeset
|
446 |
|
67508 | 447 |
val lines = separate (Latex.string "\\isanewline%\n"); |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
448 |
val items = separate (Latex.string "\\isasep\\isanewline%\n"); |
22124 | 449 |
|
67463 | 450 |
fun isabelle ctxt body = |
451 |
if Config.get ctxt Document_Antiquotation.thy_output_display |
|
452 |
then Latex.environment_block "isabelle" body |
|
453 |
else Latex.block (Latex.enclose_body "\\isa{" "}" body); |
|
59175
bf465f335e85
system option "pretty_margin" is superseded by "thy_output_margin";
wenzelm
parents:
59067
diff
changeset
|
454 |
|
67463 | 455 |
fun isabelle_typewriter ctxt body = |
456 |
if Config.get ctxt Document_Antiquotation.thy_output_display |
|
457 |
then Latex.environment_block "isabellett" body |
|
458 |
else Latex.block (Latex.enclose_body "\\isatt{" "}" body); |
|
459 |
||
460 |
fun typewriter ctxt s = |
|
461 |
isabelle_typewriter ctxt [Latex.string (Latex.output_ascii s)]; |
|
462 |
||
463 |
fun verbatim ctxt = |
|
464 |
if Config.get ctxt Document_Antiquotation.thy_output_display |
|
465 |
then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt |
|
67508 | 466 |
else split_lines #> map (typewriter ctxt) #> lines #> Latex.block; |
22124 | 467 |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
468 |
fun source ctxt = |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
469 |
Token.args_of_src |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
470 |
#> map (Token.unparse #> Document_Antiquotation.prepare_lines ctxt) |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
471 |
#> space_implode " " |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
472 |
#> output_source ctxt |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
473 |
#> isabelle ctxt; |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
474 |
|
67463 | 475 |
fun pretty ctxt = |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
476 |
Document_Antiquotation.output ctxt #> Latex.string #> single #> isabelle ctxt; |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
477 |
|
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
478 |
fun pretty_source ctxt src prt = |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
479 |
if Config.get ctxt Document_Antiquotation.thy_output_source |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
480 |
then source ctxt src else pretty ctxt prt; |
67463 | 481 |
|
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
482 |
fun pretty_items ctxt = |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
483 |
map (Document_Antiquotation.output ctxt #> Latex.string) #> items #> isabelle ctxt; |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
484 |
|
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
485 |
fun pretty_items_source ctxt src prts = |
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67465
diff
changeset
|
486 |
if Config.get ctxt Document_Antiquotation.thy_output_source |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67475
diff
changeset
|
487 |
then source ctxt src else pretty_items ctxt prts; |
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61471
diff
changeset
|
488 |
|
67465 | 489 |
|
490 |
(* antiquotation variants *) |
|
491 |
||
67463 | 492 |
fun antiquotation_pretty name scan f = |
493 |
Document_Antiquotation.setup name scan |
|
494 |
(fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x)); |
|
495 |
||
496 |
fun antiquotation_pretty_source name scan f = |
|
497 |
Document_Antiquotation.setup name scan |
|
498 |
(fn {context = ctxt, source = src, argument = x} => pretty_source ctxt src (f ctxt x)); |
|
499 |
||
500 |
fun antiquotation_raw name scan f = |
|
501 |
Document_Antiquotation.setup name scan |
|
502 |
(fn {context = ctxt, argument = x, ...} => f ctxt x); |
|
503 |
||
504 |
fun antiquotation_verbatim name scan f = |
|
505 |
antiquotation_raw name scan (fn ctxt => verbatim ctxt o f ctxt); |
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61473
diff
changeset
|
506 |
|
30390
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
wenzelm
parents:
30381
diff
changeset
|
507 |
end; |