| author | wenzelm | 
| Mon, 15 Nov 2021 17:26:31 +0100 | |
| changeset 74790 | 3ce6fb9db485 | 
| parent 74784 | d2522bb4db1b | 
| child 74882 | 947bb3e09a88 | 
| permissions | -rw-r--r-- | 
| 55030 | 1  | 
(* Title: Pure/Tools/rail.ML  | 
| 42504 | 2  | 
Author: Michael Kerscher, TU München  | 
3  | 
Author: Makarius  | 
|
4  | 
||
5  | 
Railroad diagrams in LaTeX.  | 
|
6  | 
*)  | 
|
7  | 
||
| 62748 | 8  | 
signature RAIL =  | 
9  | 
sig  | 
|
10  | 
datatype rails =  | 
|
11  | 
Cat of int * rail list  | 
|
12  | 
and rail =  | 
|
13  | 
Bar of rails list |  | 
|
14  | 
Plus of rails * rails |  | 
|
15  | 
Newline of int |  | 
|
16  | 
Nonterminal of string |  | 
|
17  | 
Terminal of bool * string |  | 
|
18  | 
Antiquote of bool * Antiquote.antiq  | 
|
19  | 
val read: Proof.context -> Input.source -> (string Antiquote.antiquote * rail) list  | 
|
| 67463 | 20  | 
val output_rules: Proof.context -> (string Antiquote.antiquote * rail) list -> Latex.text  | 
| 62748 | 21  | 
end;  | 
22  | 
||
23  | 
structure Rail: RAIL =  | 
|
| 42504 | 24  | 
struct  | 
25  | 
||
26  | 
(** lexical syntax **)  | 
|
27  | 
||
| 56165 | 28  | 
(* singleton keywords *)  | 
29  | 
||
30  | 
val keywords =  | 
|
31  | 
Symtab.make [  | 
|
32  | 
    ("|", Markup.keyword3),
 | 
|
33  | 
    ("*", Markup.keyword3),
 | 
|
34  | 
    ("+", Markup.keyword3),
 | 
|
35  | 
    ("?", Markup.keyword3),
 | 
|
36  | 
    ("(", Markup.empty),
 | 
|
37  | 
    (")", Markup.empty),
 | 
|
38  | 
    ("\<newline>", Markup.keyword2),
 | 
|
39  | 
    (";", Markup.keyword2),
 | 
|
40  | 
    (":", Markup.keyword2),
 | 
|
41  | 
    ("@", Markup.keyword1)];
 | 
|
42  | 
||
43  | 
||
| 42504 | 44  | 
(* datatype token *)  | 
45  | 
||
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
46  | 
datatype kind =  | 
| 
67425
 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 
wenzelm 
parents: 
67388 
diff
changeset
 | 
47  | 
Keyword | Ident | String | Space | Comment of Comment.kind | Antiq of Antiquote.antiq | EOF;  | 
| 42504 | 48  | 
|
49  | 
datatype token = Token of Position.range * (kind * string);  | 
|
50  | 
||
51  | 
fun pos_of (Token ((pos, _), _)) = pos;  | 
|
52  | 
fun end_pos_of (Token ((_, pos), _)) = pos;  | 
|
53  | 
||
| 58465 | 54  | 
fun range_of (toks as tok :: _) =  | 
55  | 
let val pos' = end_pos_of (List.last toks)  | 
|
| 62797 | 56  | 
in Position.range (pos_of tok, pos') end  | 
| 58465 | 57  | 
| range_of [] = Position.no_range;  | 
58  | 
||
| 42504 | 59  | 
fun kind_of (Token (_, (k, _))) = k;  | 
60  | 
fun content_of (Token (_, (_, x))) = x;  | 
|
61  | 
||
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
62  | 
fun is_proper (Token (_, (Space, _))) = false  | 
| 
67425
 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 
wenzelm 
parents: 
67388 
diff
changeset
 | 
63  | 
| is_proper (Token (_, (Comment _, _))) = false  | 
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
64  | 
| is_proper _ = true;  | 
| 
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
65  | 
|
| 42504 | 66  | 
|
67  | 
(* diagnostics *)  | 
|
68  | 
||
69  | 
val print_kind =  | 
|
70  | 
fn Keyword => "rail keyword"  | 
|
71  | 
| Ident => "identifier"  | 
|
72  | 
| String => "single-quoted string"  | 
|
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
73  | 
| Space => "white space"  | 
| 
67425
 
7d4a088dbc0e
clarified modules: uniform notion of formal comments;
 
wenzelm 
parents: 
67388 
diff
changeset
 | 
74  | 
| Comment _ => "formal comment"  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
75  | 
| Antiq _ => "antiquotation"  | 
| 
48911
 
5debc3e4fa81
tuned messages: end-of-input rarely means physical end-of-file from the past;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
76  | 
| EOF => "end-of-input";  | 
| 42504 | 77  | 
|
78  | 
fun print (Token ((pos, _), (k, x))) =  | 
|
79  | 
(if k = EOF then print_kind k else print_kind k ^ " " ^ quote x) ^  | 
|
| 48992 | 80  | 
Position.here pos;  | 
| 42504 | 81  | 
|
82  | 
fun print_keyword x = print_kind Keyword ^ " " ^ quote x;  | 
|
83  | 
||
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
84  | 
fun reports_of_token (Token ((pos, _), (Keyword, x))) =  | 
| 56165 | 85  | 
map (pair pos) (the_list (Symtab.lookup keywords x) @ Completion.suppress_abbrevs x)  | 
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
86  | 
| reports_of_token (Token ((pos, _), (String, _))) = [(pos, Markup.inner_string)]  | 
| 61457 | 87  | 
| reports_of_token (Token (_, (Antiq antiq, _))) = Antiquote.antiq_reports [Antiquote.Antiq antiq]  | 
| 55613 | 88  | 
| reports_of_token _ = [];  | 
89  | 
||
| 42504 | 90  | 
|
91  | 
(* stopper *)  | 
|
92  | 
||
93  | 
fun mk_eof pos = Token ((pos, Position.none), (EOF, ""));  | 
|
94  | 
val eof = mk_eof Position.none;  | 
|
95  | 
||
96  | 
fun is_eof (Token (_, (EOF, _))) = true  | 
|
97  | 
| is_eof _ = false;  | 
|
98  | 
||
99  | 
val stopper =  | 
|
100  | 
Scan.stopper (fn [] => eof | toks => mk_eof (end_pos_of (List.last toks))) is_eof;  | 
|
101  | 
||
102  | 
||
103  | 
(* tokenize *)  | 
|
104  | 
||
105  | 
local  | 
|
106  | 
||
107  | 
fun token k ss = [Token (Symbol_Pos.range ss, (k, Symbol_Pos.content ss))];  | 
|
108  | 
||
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61462 
diff
changeset
 | 
109  | 
fun antiq_token antiq =  | 
| 
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61462 
diff
changeset
 | 
110  | 
[Token (#range antiq, (Antiq antiq, Symbol_Pos.content (#body antiq)))];  | 
| 58465 | 111  | 
|
| 42504 | 112  | 
val scan_space = Scan.many1 (Symbol.is_blank o Symbol_Pos.symbol);  | 
113  | 
||
114  | 
val scan_keyword =  | 
|
| 56165 | 115  | 
Scan.one (Symtab.defined keywords o Symbol_Pos.symbol);  | 
| 42504 | 116  | 
|
| 48764 | 117  | 
val err_prefix = "Rail lexical error: ";  | 
118  | 
||
| 42504 | 119  | 
val scan_token =  | 
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
120  | 
scan_space >> token Space ||  | 
| 
69891
 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 
wenzelm 
parents: 
69592 
diff
changeset
 | 
121  | 
Comment.scan_inner >> (fn (kind, ss) => token (Comment kind) ss) ||  | 
| 58465 | 122  | 
Antiquote.scan_antiq >> antiq_token ||  | 
| 42504 | 123  | 
scan_keyword >> (token Keyword o single) ||  | 
124  | 
Lexicon.scan_id >> token Ident ||  | 
|
| 55613 | 125  | 
Symbol_Pos.scan_string_q err_prefix >> (fn (pos1, (ss, pos2)) =>  | 
| 62797 | 126  | 
[Token (Position.range (pos1, pos2), (String, Symbol_Pos.content ss))]);  | 
| 42504 | 127  | 
|
| 
42506
 
876887b07e8d
more robust error handling (NB: Source.source requires total scanner or recover);
 
wenzelm 
parents: 
42504 
diff
changeset
 | 
128  | 
val scan =  | 
| 61476 | 129  | 
Scan.repeats scan_token --|  | 
| 48764 | 130  | 
Symbol_Pos.!!! (fn () => err_prefix ^ "bad input")  | 
| 
42506
 
876887b07e8d
more robust error handling (NB: Source.source requires total scanner or recover);
 
wenzelm 
parents: 
42504 
diff
changeset
 | 
131  | 
(Scan.ahead (Scan.one Symbol_Pos.is_eof));  | 
| 
 
876887b07e8d
more robust error handling (NB: Source.source requires total scanner or recover);
 
wenzelm 
parents: 
42504 
diff
changeset
 | 
132  | 
|
| 42504 | 133  | 
in  | 
134  | 
||
| 55613 | 135  | 
val tokenize = #1 o Scan.error (Scan.finite Symbol_Pos.stopper scan);  | 
| 42504 | 136  | 
|
137  | 
end;  | 
|
138  | 
||
139  | 
||
140  | 
||
141  | 
(** parsing **)  | 
|
142  | 
||
| 58465 | 143  | 
(* parser combinators *)  | 
144  | 
||
| 42504 | 145  | 
fun !!! scan =  | 
146  | 
let  | 
|
147  | 
val prefix = "Rail syntax error";  | 
|
148  | 
||
| 
48911
 
5debc3e4fa81
tuned messages: end-of-input rarely means physical end-of-file from the past;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
149  | 
fun get_pos [] = " (end-of-input)"  | 
| 48992 | 150  | 
| get_pos (tok :: _) = Position.here (pos_of tok);  | 
| 42504 | 151  | 
|
| 
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
 | 
152  | 
fun err (toks, NONE) = (fn () => prefix ^ get_pos toks)  | 
| 42504 | 153  | 
| err (toks, SOME msg) =  | 
| 
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
 | 
154  | 
(fn () =>  | 
| 
 
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
 | 
155  | 
let val s = msg () in  | 
| 
 
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
 | 
156  | 
if String.isPrefix prefix s then s  | 
| 
 
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
 | 
157  | 
else prefix ^ get_pos toks ^ ": " ^ s  | 
| 
 
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
 | 
158  | 
end);  | 
| 42504 | 159  | 
in Scan.!! err scan end;  | 
160  | 
||
161  | 
fun $$$ x =  | 
|
162  | 
Scan.one (fn tok => kind_of tok = Keyword andalso content_of tok = x) ||  | 
|
163  | 
Scan.fail_with  | 
|
| 
48911
 
5debc3e4fa81
tuned messages: end-of-input rarely means physical end-of-file from the past;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
164  | 
(fn [] => (fn () => print_keyword x ^ " expected,\nbut end-of-input was found")  | 
| 
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
 | 
165  | 
| tok :: _ => (fn () => print_keyword x ^ " expected,\nbut " ^ print tok ^ " was found"));  | 
| 42504 | 166  | 
|
167  | 
fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan);  | 
|
168  | 
fun enum sep scan = enum1 sep scan || Scan.succeed [];  | 
|
169  | 
||
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
170  | 
val ident = Scan.some (fn tok => if kind_of tok = Ident then SOME (content_of tok) else NONE);  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
171  | 
val string = Scan.some (fn tok => if kind_of tok = String then SOME (content_of tok) else NONE);  | 
| 42504 | 172  | 
|
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
173  | 
val antiq = Scan.some (fn tok => (case kind_of tok of Antiq a => SOME a | _ => NONE));  | 
| 42504 | 174  | 
|
| 58465 | 175  | 
fun RANGE scan = Scan.trace scan >> apsnd range_of;  | 
176  | 
fun RANGE_APP scan = RANGE scan >> (fn (f, r) => f r);  | 
|
177  | 
||
178  | 
||
179  | 
(* parse trees *)  | 
|
180  | 
||
181  | 
datatype trees =  | 
|
182  | 
CAT of tree list * Position.range  | 
|
183  | 
and tree =  | 
|
184  | 
BAR of trees list * Position.range |  | 
|
185  | 
STAR of (trees * trees) * Position.range |  | 
|
186  | 
PLUS of (trees * trees) * Position.range |  | 
|
187  | 
MAYBE of tree * Position.range |  | 
|
188  | 
NEWLINE of Position.range |  | 
|
189  | 
NONTERMINAL of string * Position.range |  | 
|
190  | 
TERMINAL of (bool * string) * Position.range |  | 
|
191  | 
ANTIQUOTE of (bool * Antiquote.antiq) * Position.range;  | 
|
192  | 
||
193  | 
fun reports_of_tree ctxt =  | 
|
| 71675 | 194  | 
if Context_Position.reports_enabled ctxt then  | 
| 58465 | 195  | 
let  | 
196  | 
fun reports r =  | 
|
197  | 
if r = Position.no_range then []  | 
|
| 62806 | 198  | 
else [(Position.range_position r, Markup.expression "")];  | 
| 58465 | 199  | 
fun trees (CAT (ts, r)) = reports r @ maps tree ts  | 
200  | 
and tree (BAR (Ts, r)) = reports r @ maps trees Ts  | 
|
201  | 
| tree (STAR ((T1, T2), r)) = reports r @ trees T1 @ trees T2  | 
|
202  | 
| tree (PLUS ((T1, T2), r)) = reports r @ trees T1 @ trees T2  | 
|
203  | 
| tree (MAYBE (t, r)) = reports r @ tree t  | 
|
204  | 
| tree (NEWLINE r) = reports r  | 
|
205  | 
| tree (NONTERMINAL (_, r)) = reports r  | 
|
206  | 
| tree (TERMINAL (_, r)) = reports r  | 
|
207  | 
| tree (ANTIQUOTE (_, r)) = reports r;  | 
|
208  | 
in distinct (op =) o tree end  | 
|
209  | 
else K [];  | 
|
210  | 
||
211  | 
local  | 
|
212  | 
||
213  | 
val at_mode = Scan.option ($$$ "@") >> (fn NONE => false | _ => true);  | 
|
214  | 
||
215  | 
fun body x = (RANGE (enum1 "|" body1) >> BAR) x  | 
|
216  | 
and body0 x = (RANGE (enum "|" body1) >> BAR) x  | 
|
217  | 
and body1 x =  | 
|
218  | 
(RANGE_APP (body2 :|-- (fn a =>  | 
|
219  | 
$$$ "*" |-- !!! body4e >> (fn b => fn r => CAT ([STAR ((a, b), r)], r)) ||  | 
|
220  | 
$$$ "+" |-- !!! body4e >> (fn b => fn r => CAT ([PLUS ((a, b), r)], r)) ||  | 
|
221  | 
Scan.succeed (K a)))) x  | 
|
222  | 
and body2 x = (RANGE (Scan.repeat1 body3) >> CAT) x  | 
|
223  | 
and body3 x =  | 
|
224  | 
(RANGE_APP (body4 :|-- (fn a =>  | 
|
225  | 
$$$ "?" >> K (curry MAYBE a) ||  | 
|
226  | 
Scan.succeed (K a)))) x  | 
|
227  | 
and body4 x =  | 
|
228  | 
 ($$$ "(" |-- !!! (body0 --| $$$ ")") ||
 | 
|
229  | 
RANGE_APP  | 
|
230  | 
($$$ "\<newline>" >> K NEWLINE ||  | 
|
231  | 
ident >> curry NONTERMINAL ||  | 
|
232  | 
at_mode -- string >> curry TERMINAL ||  | 
|
233  | 
at_mode -- antiq >> curry ANTIQUOTE)) x  | 
|
234  | 
and body4e x =  | 
|
235  | 
(RANGE (Scan.option body4) >> (fn (a, r) => CAT (the_list a, r))) x;  | 
|
236  | 
||
237  | 
val rule_name = ident >> Antiquote.Text || antiq >> Antiquote.Antiq;  | 
|
238  | 
val rule = rule_name -- ($$$ ":" |-- !!! body) || body >> pair (Antiquote.Text "");  | 
|
239  | 
val rules = enum1 ";" (Scan.option rule) >> map_filter I;  | 
|
240  | 
||
241  | 
in  | 
|
242  | 
||
243  | 
fun parse_rules toks =  | 
|
244  | 
#1 (Scan.error (Scan.finite stopper (rules --| !!! (Scan.ahead (Scan.one is_eof)))) toks);  | 
|
245  | 
||
246  | 
end;  | 
|
| 42504 | 247  | 
|
248  | 
||
249  | 
(** rail expressions **)  | 
|
250  | 
||
251  | 
(* datatype *)  | 
|
252  | 
||
253  | 
datatype rails =  | 
|
254  | 
Cat of int * rail list  | 
|
255  | 
and rail =  | 
|
256  | 
Bar of rails list |  | 
|
257  | 
Plus of rails * rails |  | 
|
258  | 
Newline of int |  | 
|
259  | 
Nonterminal of string |  | 
|
| 42516 | 260  | 
Terminal of bool * string |  | 
| 55526 | 261  | 
Antiquote of bool * Antiquote.antiq;  | 
| 42504 | 262  | 
|
| 58465 | 263  | 
fun is_newline (Newline _) = true | is_newline _ = false;  | 
264  | 
||
265  | 
||
266  | 
(* prepare *)  | 
|
267  | 
||
268  | 
local  | 
|
| 42504 | 269  | 
|
270  | 
fun cat rails = Cat (0, rails);  | 
|
271  | 
||
272  | 
val empty = cat [];  | 
|
273  | 
fun is_empty (Cat (_, [])) = true | is_empty _ = false;  | 
|
274  | 
||
275  | 
fun bar [Cat (_, [rail])] = rail  | 
|
276  | 
| bar cats = Bar cats;  | 
|
277  | 
||
| 58465 | 278  | 
fun reverse_cat (Cat (y, rails)) = Cat (y, rev (map reverse rails))  | 
279  | 
and reverse (Bar cats) = Bar (map reverse_cat cats)  | 
|
280  | 
| reverse (Plus (cat1, cat2)) = Plus (reverse_cat cat1, reverse_cat cat2)  | 
|
281  | 
| reverse x = x;  | 
|
282  | 
||
283  | 
fun plus (cat1, cat2) = Plus (cat1, reverse_cat cat2);  | 
|
284  | 
||
285  | 
in  | 
|
| 42504 | 286  | 
|
| 58465 | 287  | 
fun prepare_trees (CAT (ts, _)) = Cat (0, map prepare_tree ts)  | 
288  | 
and prepare_tree (BAR (Ts, _)) = bar (map prepare_trees Ts)  | 
|
289  | 
| prepare_tree (STAR (Ts, _)) =  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58978 
diff
changeset
 | 
290  | 
let val (cat1, cat2) = apply2 prepare_trees Ts in  | 
| 58465 | 291  | 
if is_empty cat2 then plus (empty, cat1)  | 
292  | 
else bar [empty, cat [plus (cat1, cat2)]]  | 
|
293  | 
end  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58978 
diff
changeset
 | 
294  | 
| prepare_tree (PLUS (Ts, _)) = plus (apply2 prepare_trees Ts)  | 
| 58465 | 295  | 
| prepare_tree (MAYBE (t, _)) = bar [empty, cat [prepare_tree t]]  | 
296  | 
| prepare_tree (NEWLINE _) = Newline 0  | 
|
297  | 
| prepare_tree (NONTERMINAL (a, _)) = Nonterminal a  | 
|
298  | 
| prepare_tree (TERMINAL (a, _)) = Terminal a  | 
|
299  | 
| prepare_tree (ANTIQUOTE (a, _)) = Antiquote a;  | 
|
| 42504 | 300  | 
|
| 58465 | 301  | 
end;  | 
| 42504 | 302  | 
|
303  | 
||
304  | 
(* read *)  | 
|
305  | 
||
| 59064 | 306  | 
fun read ctxt source =  | 
| 55613 | 307  | 
let  | 
| 59064 | 308  | 
val _ = Context_Position.report ctxt (Input.pos_of source) Markup.language_rail;  | 
309  | 
val toks = tokenize (Input.source_explode source);  | 
|
| 55613 | 310  | 
val _ = Context_Position.reports ctxt (maps reports_of_token toks);  | 
| 
67387
 
ff07dd9c7cb4
clarified rail token language: white space and formal comments;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
311  | 
val rules = parse_rules (filter is_proper toks);  | 
| 71674 | 312  | 
val _ = Context_Position.reports ctxt (maps (reports_of_tree ctxt o #2) rules);  | 
| 58465 | 313  | 
in map (apsnd prepare_tree) rules end;  | 
| 42504 | 314  | 
|
315  | 
||
316  | 
(* latex output *)  | 
|
317  | 
||
318  | 
local  | 
|
319  | 
||
320  | 
fun vertical_range_cat (Cat (_, rails)) y =  | 
|
321  | 
let val (rails', (_, y')) =  | 
|
322  | 
fold_map (fn rail => fn (y0, y') =>  | 
|
323  | 
if is_newline rail then (Newline (y' + 1), (y' + 1, y' + 2))  | 
|
324  | 
else  | 
|
325  | 
let val (rail', y0') = vertical_range rail y0;  | 
|
326  | 
in (rail', (y0, Int.max (y0', y'))) end) rails (y, y + 1)  | 
|
327  | 
in (Cat (y, rails'), y') end  | 
|
328  | 
||
329  | 
and vertical_range (Bar cats) y =  | 
|
330  | 
let val (cats', y') = fold_map vertical_range_cat cats y  | 
|
331  | 
in (Bar cats', Int.max (y + 1, y')) end  | 
|
332  | 
| vertical_range (Plus (cat1, cat2)) y =  | 
|
333  | 
let val ([cat1', cat2'], y') = fold_map vertical_range_cat [cat1, cat2] y;  | 
|
334  | 
in (Plus (cat1', cat2'), Int.max (y + 1, y')) end  | 
|
335  | 
| vertical_range (Newline _) y = (Newline (y + 2), y + 3)  | 
|
336  | 
| vertical_range atom y = (atom, y + 1);  | 
|
337  | 
||
| 62748 | 338  | 
in  | 
339  | 
||
| 67381 | 340  | 
fun output_rules ctxt rules =  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
341  | 
let  | 
| 67463 | 342  | 
val output_antiq =  | 
| 
67571
 
f858fe5531ac
more uniform treatment of formal comments within document source;
 
wenzelm 
parents: 
67463 
diff
changeset
 | 
343  | 
Antiquote.Antiq #>  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
344  | 
Document_Antiquotation.evaluate Latex.symbols ctxt;  | 
| 42516 | 345  | 
fun output_text b s =  | 
| 74790 | 346  | 
Latex.string (Output.output s)  | 
347  | 
|> b ? Latex.macro "isakeyword"  | 
|
348  | 
|> Latex.macro "isa";  | 
|
| 42504 | 349  | 
|
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
350  | 
fun output_cat c (Cat (_, rails)) = outputs c rails  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
351  | 
and outputs c [rail] = output c rail  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
352  | 
| outputs _ rails = maps (output "") rails  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
353  | 
and output _ (Bar []) = []  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
354  | 
| output c (Bar [cat]) = output_cat c cat  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
355  | 
| output _ (Bar (cat :: cats)) =  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
356  | 
          Latex.string ("\\rail@bar\n") @ output_cat "" cat @
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
357  | 
maps (fn Cat (y, rails) =>  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
358  | 
            Latex.string ("\\rail@nextbar{" ^ string_of_int y ^ "}\n") @ outputs "" rails) cats @
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
359  | 
Latex.string "\\rail@endbar\n"  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
360  | 
| output c (Plus (cat, Cat (y, rails))) =  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
361  | 
Latex.string "\\rail@plus\n" @ output_cat c cat @  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
362  | 
          Latex.string ("\\rail@nextplus{" ^ string_of_int y ^ "}\n") @ outputs "c" rails @
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
363  | 
Latex.string "\\rail@endplus\n"  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
364  | 
      | output _ (Newline y) = Latex.string ("\\rail@cr{" ^ string_of_int y ^ "}\n")
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
365  | 
| output c (Nonterminal s) =  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
366  | 
          Latex.string ("\\rail@" ^ c ^ "nont{") @ output_text false s @ Latex.string "}[]\n"
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
367  | 
| output c (Terminal (b, s)) =  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
368  | 
          Latex.string ("\\rail@" ^ c ^ "term{") @ output_text b s @ Latex.string "}[]\n"
 | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
369  | 
| output c (Antiquote (b, a)) =  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
370  | 
          Latex.string ("\\rail@" ^ c ^ (if b then "term{" else "nont{")) @
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
371  | 
Latex.output (output_antiq a) @  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
372  | 
Latex.string "}[]\n";  | 
| 42504 | 373  | 
|
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
374  | 
fun output_rule (name, rail) =  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
375  | 
let  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
376  | 
val (rail', y') = vertical_range rail 0;  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
377  | 
val out_name =  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
378  | 
(case name of  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
379  | 
Antiquote.Text "" => []  | 
| 
42661
 
824d3f1d8de6
proper treatment of empty name -- avoid excessive vertical space;
 
wenzelm 
parents: 
42657 
diff
changeset
 | 
380  | 
| Antiquote.Text s => output_text false s  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
381  | 
| Antiquote.Antiq a => output_antiq a);  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
382  | 
in  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
383  | 
        Latex.string ("\\rail@begin{" ^ string_of_int y' ^ "}{") @ out_name @ Latex.string "}\n" @
 | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
384  | 
output "" rail' @  | 
| 
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
385  | 
Latex.string "\\rail@end\n"  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42507 
diff
changeset
 | 
386  | 
end;  | 
| 
74784
 
d2522bb4db1b
symbolic latex_output via XML, interpreted in Isabelle/Scala;
 
wenzelm 
parents: 
74781 
diff
changeset
 | 
387  | 
in Latex.environment_text "railoutput" (maps output_rule rules) end;  | 
| 42504 | 388  | 
|
| 53171 | 389  | 
val _ = Theory.setup  | 
| 73761 | 390  | 
(Document_Output.antiquotation_raw_embedded \<^binding>\<open>rail\<close> (Scan.lift Args.text_input)  | 
| 67463 | 391  | 
(fn ctxt => output_rules ctxt o read ctxt));  | 
| 42504 | 392  | 
|
393  | 
end;  | 
|
394  | 
||
395  | 
end;  |