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