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