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