| author | skalberg | 
| Thu, 28 Aug 2003 01:56:40 +0200 | |
| changeset 14171 | 0cab06e3bbd0 | 
| parent 13730 | 09aeb7346d3f | 
| child 14173 | a3690aeb79d4 | 
| permissions | -rw-r--r-- | 
| 6118 | 1  | 
(* Title: Pure/General/symbol.ML  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Markus Wenzel, TU Muenchen  | 
| 8806 | 4  | 
License: GPL (GNU GENERAL PUBLIC LICENSE)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
|
| 12116 | 6  | 
Generalized characters with infinitely many named symbols.  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
signature SYMBOL =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
type symbol  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
val space: symbol  | 
| 10953 | 13  | 
val spaces: int -> symbol  | 
| 6857 | 14  | 
val sync: symbol  | 
15  | 
val is_sync: symbol -> bool  | 
|
16  | 
val not_sync: symbol -> bool  | 
|
| 10747 | 17  | 
val malformed: symbol  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
val eof: symbol  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
val is_eof: symbol -> bool  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
val not_eof: symbol -> bool  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
val stopper: symbol * (symbol -> bool)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
val is_ascii: symbol -> bool  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
val is_letter: symbol -> bool  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
val is_digit: symbol -> bool  | 
| 12904 | 25  | 
val is_quasi: symbol -> bool  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
val is_quasi_letter: symbol -> bool  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
val is_letdig: symbol -> bool  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
val is_blank: symbol -> bool  | 
| 
13559
 
51c3ac47d127
added checking so that (rename_tac "x y") is rejected, since
 
paulson 
parents: 
12904 
diff
changeset
 | 
29  | 
val is_identifier: symbol -> bool  | 
| 8230 | 30  | 
val is_symbolic: symbol -> bool  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
val is_printable: symbol -> bool  | 
| 6272 | 32  | 
val length: symbol list -> int  | 
| 11010 | 33  | 
val strip_blanks: string -> string  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
val beginning: symbol list -> string  | 
| 13730 | 35  | 
val scan_id: string list -> string * string list  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
val scan: string list -> symbol * string list  | 
| 6640 | 37  | 
val scanner: string -> (symbol list -> 'a * symbol list) -> symbol list -> 'a  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
val source: bool -> (string, 'a) Source.source ->  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
(symbol, (string, 'a) Source.source) Source.source  | 
| 6272 | 40  | 
val explode: string -> symbol list  | 
| 12904 | 41  | 
val bump_string: string -> string  | 
| 10953 | 42  | 
val default_indent: string * int -> string  | 
43  | 
val add_mode: string -> (string -> string * real) * (string * int -> string) -> unit  | 
|
| 6692 | 44  | 
val symbolsN: string  | 
45  | 
val xsymbolsN: string  | 
|
| 10923 | 46  | 
val plain_output: string -> string  | 
| 6272 | 47  | 
val output: string -> string  | 
48  | 
val output_width: string -> string * real  | 
|
| 10953 | 49  | 
val indent: string * int -> string  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
structure Symbol: SYMBOL =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
struct  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
|
| 6272 | 56  | 
(** generalized characters **)  | 
57  | 
||
58  | 
(*symbols, which are considered the smallest entities of any Isabelle  | 
|
59  | 
string, may be of the following form:  | 
|
60  | 
(a) ASCII symbols: a  | 
|
61  | 
(b) printable symbols: \<ident>  | 
|
62  | 
(c) control symbols: \<^ident>  | 
|
63  | 
||
| 12116 | 64  | 
output is subject to the print_mode variable (default: verbatim),  | 
65  | 
actual interpretation in display is up to front-end tools;  | 
|
| 6272 | 66  | 
*)  | 
67  | 
||
68  | 
type symbol = string;  | 
|
69  | 
||
70  | 
val space = " ";  | 
|
| 10953 | 71  | 
fun spaces k = Library.replicate_string k space;  | 
| 6857 | 72  | 
val sync = "\\<^sync>";  | 
| 10747 | 73  | 
val malformed = "\\<^malformed>";  | 
| 6272 | 74  | 
val eof = "";  | 
75  | 
||
76  | 
||
77  | 
(* kinds *)  | 
|
78  | 
||
| 6857 | 79  | 
fun is_sync s = s = sync;  | 
80  | 
fun not_sync s = s <> sync;  | 
|
81  | 
||
| 6272 | 82  | 
fun is_eof s = s = eof;  | 
83  | 
fun not_eof s = s <> eof;  | 
|
84  | 
val stopper = (eof, is_eof);  | 
|
85  | 
||
86  | 
fun is_ascii s = size s = 1 andalso ord s < 128;  | 
|
87  | 
||
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
88  | 
local  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
89  | 
fun wrap s = "\\<" ^ s ^ ">"  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
90  | 
val pre_digits =  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
91  | 
["zero","one","two","three","four",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
92  | 
"five","six","seven","eight","nine"]  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
93  | 
|
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
94  | 
val cal_letters =  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
95  | 
["A","B","C","D","E","F","G","H","I","J","K","L","M",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
96  | 
"N","O","P","Q","R","S","T","U","V","W","X","Y","Z"]  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
97  | 
|
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
98  | 
val small_letters =  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
99  | 
["a","b","c","d","e","f","g","h","i","j","k","l","m",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
100  | 
"n","o","p","q","r","s","t","u","v","w","x","y","z"]  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
101  | 
|
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
102  | 
val goth_letters =  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
103  | 
["AA","BB","CC","DD","EE","FF","GG","HH","II","JJ","KK","LL","MM",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
104  | 
"NN","OO","PP","QQ","RR","SS","TT","UU","VV","WW","XX","YY","ZZ",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
105  | 
"aa","bb","cc","dd","ee","ff","gg","hh","ii","jj","kk","ll","mm",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
106  | 
"nn","oo","pp","qq","rr","ss","tt","uu","vv","ww","xx","yy","zz"]  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
107  | 
|
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
108  | 
val greek_letters =  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
109  | 
["alpha","beta","gamma","delta","epsilon","zeta","eta","theta",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
110  | 
"iota","kappa",(*"lambda",*)"mu","nu","xi","pi","rho","sigma","tau",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
111  | 
"upsilon","phi","psi","omega","Gamma","Delta","Theta","Lambda",  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
112  | 
"Xi","Pi","Sigma","Upsilon","Phi","Psi","Omega"]  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
113  | 
|
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
114  | 
val bbb_letters = ["bool","complex","nat","rat","real","int"]  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
115  | 
|
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
116  | 
val pre_letters =  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
117  | 
cal_letters @  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
118  | 
small_letters @  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
119  | 
goth_letters @  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
120  | 
greek_letters  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
121  | 
in  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
122  | 
val ext_digits = map wrap pre_digits  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
123  | 
val ext_letters = map wrap pre_letters  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
124  | 
val ext_letdigs = ext_digits @ ext_letters  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
125  | 
fun is_ext_digit s = String.isPrefix "\\<" s andalso s mem ext_digits  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
126  | 
fun is_ext_letter s = String.isPrefix "\\<" s andalso s mem ext_letters  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
127  | 
fun is_ext_letdig s = String.isPrefix "\\<" s andalso s mem ext_letdigs  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
128  | 
end  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
129  | 
|
| 6272 | 130  | 
fun is_letter s =  | 
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
131  | 
(size s = 1 andalso  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
132  | 
(ord "A" <= ord s andalso ord s <= ord "Z" orelse  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
133  | 
ord "a" <= ord s andalso ord s <= ord "z"))  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
134  | 
orelse is_ext_letter s  | 
| 6272 | 135  | 
|
136  | 
fun is_digit s =  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
137  | 
(size s = 1 andalso ord "0" <= ord s andalso ord s <= ord "9")  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
138  | 
orelse is_ext_digit s  | 
| 6272 | 139  | 
|
| 12904 | 140  | 
fun is_quasi "_" = true  | 
141  | 
| is_quasi "'" = true  | 
|
142  | 
| is_quasi _ = false;  | 
|
143  | 
||
144  | 
fun is_quasi_letter s = is_quasi s orelse is_letter s;  | 
|
| 6272 | 145  | 
|
146  | 
val is_blank =  | 
|
147  | 
fn " " => true | "\t" => true | "\n" => true | "\^L" => true  | 
|
148  | 
| "\160" => true | "\\<spacespace>" => true  | 
|
149  | 
| _ => false;  | 
|
150  | 
||
| 12904 | 151  | 
fun is_letdig s = is_quasi_letter s orelse is_digit s;  | 
| 6272 | 152  | 
|
| 8230 | 153  | 
fun is_symbolic s =  | 
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
154  | 
size s > 2 andalso nth_elem_string (2, s) <> "^"  | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
155  | 
andalso not (is_ext_letdig s);  | 
| 8230 | 156  | 
|
| 6272 | 157  | 
fun is_printable s =  | 
158  | 
size s = 1 andalso ord space <= ord s andalso ord s <= ord "~" orelse  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13730 
diff
changeset
 | 
159  | 
is_ext_letdig s orelse  | 
| 8230 | 160  | 
is_symbolic s;  | 
161  | 
||
| 
13559
 
51c3ac47d127
added checking so that (rename_tac "x y") is rejected, since
 
paulson 
parents: 
12904 
diff
changeset
 | 
162  | 
fun is_identifier s =  | 
| 
 
51c3ac47d127
added checking so that (rename_tac "x y") is rejected, since
 
paulson 
parents: 
12904 
diff
changeset
 | 
163  | 
case (explode s) of  | 
| 
 
51c3ac47d127
added checking so that (rename_tac "x y") is rejected, since
 
paulson 
parents: 
12904 
diff
changeset
 | 
164  | 
[] => false  | 
| 
 
51c3ac47d127
added checking so that (rename_tac "x y") is rejected, since
 
paulson 
parents: 
12904 
diff
changeset
 | 
165  | 
| c::cs => is_letter c andalso forall is_letdig cs;  | 
| 6272 | 166  | 
|
| 10738 | 167  | 
fun sym_length ss = foldl (fn (n, s) =>  | 
168  | 
(if not (is_printable s) then 0 else  | 
|
169  | 
(case Library.try String.substring (s, 2, 4) of  | 
|
170  | 
Some s' => if s' = "long" orelse s' = "Long" then 2 else 1  | 
|
171  | 
| None => 1)) + n) (0, ss);  | 
|
| 6272 | 172  | 
|
| 11010 | 173  | 
fun strip_blanks s =  | 
174  | 
implode (#1 (Library.take_suffix is_blank (#2 (Library.take_prefix is_blank (explode s)))));  | 
|
175  | 
||
| 6272 | 176  | 
|
177  | 
(* beginning *)  | 
|
178  | 
||
179  | 
val smash_blanks = map (fn s => if is_blank s then space else s);  | 
|
180  | 
||
181  | 
fun beginning raw_ss =  | 
|
182  | 
let  | 
|
183  | 
val (all_ss, _) = take_suffix is_blank raw_ss;  | 
|
184  | 
val dots = if length all_ss > 10 then " ..." else "";  | 
|
185  | 
val (ss, _) = take_suffix is_blank (take (10, all_ss));  | 
|
186  | 
in implode (smash_blanks ss) ^ dots end;  | 
|
187  | 
||
188  | 
||
189  | 
||
| 8998 | 190  | 
(** scanning through symbols **)  | 
| 6640 | 191  | 
|
192  | 
fun scanner msg scan chs =  | 
|
193  | 
let  | 
|
194  | 
fun err_msg cs = msg ^ ": " ^ beginning cs;  | 
|
195  | 
val fin_scan = Scan.error (Scan.finite stopper (!! (fn (cs, _) => err_msg cs) scan));  | 
|
196  | 
in  | 
|
197  | 
(case fin_scan chs of  | 
|
198  | 
(result, []) => result  | 
|
199  | 
| (_, rest) => error (err_msg rest))  | 
|
200  | 
end;  | 
|
201  | 
||
202  | 
||
203  | 
||
| 6272 | 204  | 
(** symbol input **)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
(* scan *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
val scan_id = Scan.one is_letter ^^ (Scan.any is_letdig >> implode);  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
val scan =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
211  | 
($$ "\\" --| Scan.optional ($$ "\\") "") ^^ $$ "<" ^^  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
!! (fn (cs, _) => "Malformed symbolic character specification: \\" ^ "<" ^ beginning cs)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
(Scan.optional ($$ "^") "" ^^ scan_id ^^ $$ ">") ||  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
Scan.one not_eof;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
217  | 
(* source *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
218  | 
|
| 10747 | 219  | 
val recover = Scan.any ((not o is_blank) andf not_eof) >> K [malformed];  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
fun source do_recover src =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
Source.source stopper (Scan.bulk scan) (if do_recover then Some recover else None) src;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
(* explode *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
226  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
227  | 
fun no_syms [] = true  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
  | no_syms ("\\" :: "<" :: _) = false
 | 
| 12116 | 229  | 
| no_syms (_ :: cs) = no_syms cs;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
230  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
231  | 
fun sym_explode str =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
let val chs = explode str in  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
233  | 
if no_syms chs then chs (*tune trivial case*)  | 
| 12116 | 234  | 
else the (Scan.read stopper (Scan.repeat scan) chs)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
235  | 
end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
236  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
237  | 
|
| 12904 | 238  | 
(* bump_string -- increment suffix of lowercase letters like a base 26 number *)  | 
239  | 
||
240  | 
fun bump_string str =  | 
|
241  | 
let  | 
|
242  | 
fun bump [] = ["a"]  | 
|
243  | 
      | bump ("z" :: ss) = "a" :: bump ss
 | 
|
244  | 
| bump (s :: ss) =  | 
|
245  | 
if size s = 1 andalso ord "a" <= ord s andalso ord s < ord "z"  | 
|
246  | 
then chr (ord s + 1) :: ss  | 
|
247  | 
else "a" :: s :: ss;  | 
|
248  | 
val (cs, qs) = Library.take_suffix is_quasi (sym_explode str);  | 
|
249  | 
in implode (rev (bump (rev cs)) @ qs) end;  | 
|
250  | 
||
| 6272 | 251  | 
|
252  | 
(** symbol output **)  | 
|
253  | 
||
| 10953 | 254  | 
(* default *)  | 
| 6272 | 255  | 
|
256  | 
fun string_size s = (s, real (size s));  | 
|
257  | 
||
| 13730 | 258  | 
val escape = Scan.repeat  | 
259  | 
((($$ "\\" >> K "\\\\") ^^ Scan.optional ($$ "\\" >> K "\\\\") "" ^^ $$ "<" ^^  | 
|
260  | 
Scan.optional ($$ "^") "" ^^ scan_id ^^ $$ ">") ||  | 
|
261  | 
Scan.one not_eof) >> implode;  | 
|
262  | 
||
| 6272 | 263  | 
fun default_output s =  | 
| 6320 | 264  | 
if not (exists_string (equal "\\") s) then string_size s  | 
| 13730 | 265  | 
else string_size (fst (Scan.finite stopper escape (explode s)));  | 
| 6272 | 266  | 
|
| 10953 | 267  | 
fun default_indent (_: string, k) = spaces k;  | 
| 6272 | 268  | 
|
| 10953 | 269  | 
|
| 6272 | 270  | 
(* maintain modes *)  | 
271  | 
||
| 6692 | 272  | 
val symbolsN = "symbols";  | 
273  | 
val xsymbolsN = "xsymbols";  | 
|
274  | 
||
| 12116 | 275  | 
val modes =  | 
276  | 
ref (Symtab.empty: ((string -> string * real) * (string * int -> string)) Symtab.table);  | 
|
| 6272 | 277  | 
|
278  | 
fun lookup_mode name = Symtab.lookup (! modes, name);  | 
|
279  | 
||
| 10953 | 280  | 
fun add_mode name m =  | 
| 6272 | 281  | 
(if is_none (lookup_mode name) then ()  | 
| 6320 | 282  | 
  else warning ("Redeclaration of symbol print mode " ^ quote name);
 | 
| 10953 | 283  | 
modes := Symtab.update ((name, m), ! modes));  | 
284  | 
||
285  | 
fun get_mode () =  | 
|
286  | 
if_none (get_first lookup_mode (! print_mode)) (default_output, default_indent);  | 
|
| 6272 | 287  | 
|
288  | 
||
289  | 
(* mode output *)  | 
|
290  | 
||
| 10953 | 291  | 
fun output_width x = #1 (get_mode ()) x;  | 
| 6272 | 292  | 
val output = #1 o output_width;  | 
| 10923 | 293  | 
val plain_output = #1 o default_output;  | 
| 6272 | 294  | 
|
| 10953 | 295  | 
fun indent x = #2 (get_mode ()) x;  | 
296  | 
||
| 6272 | 297  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
298  | 
(*final declarations of this structure!*)  | 
| 6272 | 299  | 
val length = sym_length;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
300  | 
val explode = sym_explode;  | 
| 6272 | 301  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
302  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
303  | 
end;  |