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