| author | haftmann |
| Sat, 02 Jul 2011 10:37:35 +0200 | |
| changeset 43639 | 9cba66fb109a |
| parent 43511 | d138e7482a1b |
| child 43675 | 8252d51d70e2 |
| permissions | -rw-r--r-- |
| 27901 | 1 |
/* Title: Pure/General/symbol.scala |
2 |
Author: Makarius |
|
3 |
||
| 27924 | 4 |
Detecting and recoding Isabelle symbols. |
| 27901 | 5 |
*/ |
6 |
||
7 |
package isabelle |
|
8 |
||
| 27918 | 9 |
import scala.io.Source |
|
36011
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
wenzelm
parents:
34316
diff
changeset
|
10 |
import scala.collection.mutable |
| 31522 | 11 |
import scala.util.matching.Regex |
| 27901 | 12 |
|
13 |
||
| 31522 | 14 |
object Symbol |
15 |
{
|
|
| 36763 | 16 |
/* spaces */ |
17 |
||
| 36816 | 18 |
val spc = ' ' |
19 |
val space = " " |
|
20 |
||
21 |
private val static_spaces = space * 4000 |
|
| 36763 | 22 |
|
23 |
def spaces(k: Int): String = |
|
24 |
{
|
|
25 |
require(k >= 0) |
|
26 |
if (k < static_spaces.length) static_spaces.substring(0, k) |
|
| 36816 | 27 |
else space * k |
| 36763 | 28 |
} |
29 |
||
30 |
||
| 43418 | 31 |
/* ASCII characters */ |
32 |
||
33 |
def is_ascii_letter(c: Char): Boolean = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' |
|
34 |
def is_ascii_digit(c: Char): Boolean = '0' <= c && c <= '9' |
|
35 |
def is_ascii_quasi(c: Char): Boolean = c == '_' || c == '\'' |
|
36 |
||
37 |
def is_ascii_letdig(c: Char): Boolean = |
|
38 |
is_ascii_letter(c) || is_ascii_digit(c) || is_ascii_quasi(c) |
|
39 |
||
40 |
def is_ascii_identifier(s: String): Boolean = |
|
41 |
s.length > 0 && is_ascii_letter(s(0)) && s.substring(1).forall(is_ascii_letdig) |
|
42 |
||
43 |
||
| 33998 | 44 |
/* Symbol regexps */ |
| 27901 | 45 |
|
| 31522 | 46 |
private val plain = new Regex("""(?xs)
|
|
40524
6131d7a78ad3
treat Unicode "replacement character" (i.e. decoding error) is malformed;
wenzelm
parents:
40523
diff
changeset
|
47 |
[^\r\\\ud800-\udfff\ufffd] | [\ud800-\udbff][\udc00-\udfff] """) |
|
37556
2bf29095d26f
treat alternative newline symbols as in Isabelle/ML;
wenzelm
parents:
36816
diff
changeset
|
48 |
|
| 40522 | 49 |
private val physical_newline = new Regex("""(?xs) \n | \r\n | \r """)
|
| 27901 | 50 |
|
| 31522 | 51 |
private val symbol = new Regex("""(?xs)
|
|
31545
5f1f0a20af4d
discontinued escaped symbols such as \\<forall> -- only one backslash should be used;
wenzelm
parents:
31523
diff
changeset
|
52 |
\\ < (?: |
| 27924 | 53 |
\^? [A-Za-z][A-Za-z0-9_']* | |
54 |
\^raw: [\x20-\x7e\u0100-\uffff && [^.>]]* ) >""") |
|
55 |
||
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
56 |
private val malformed_symbol = new Regex("(?xs) (?!" + symbol + ")" +
|
| 40529 | 57 |
""" [\ud800-\udbff\ufffd] | \\<\^? """) |
| 27924 | 58 |
|
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
59 |
val regex_total = |
|
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
60 |
new Regex(plain + "|" + physical_newline + "|" + symbol + "|" + malformed_symbol + "| .") |
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
61 |
|
| 34137 | 62 |
|
63 |
/* basic matching */ |
|
64 |
||
|
37556
2bf29095d26f
treat alternative newline symbols as in Isabelle/ML;
wenzelm
parents:
36816
diff
changeset
|
65 |
def is_plain(c: Char): Boolean = !(c == '\r' || c == '\\' || '\ud800' <= c && c <= '\udfff') |
| 34137 | 66 |
|
| 38877 | 67 |
def is_physical_newline(s: CharSequence): Boolean = |
68 |
"\n".contentEquals(s) || "\r".contentEquals(s) || "\r\n".contentEquals(s) |
|
69 |
||
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
70 |
def is_malformed(s: CharSequence): Boolean = |
|
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
71 |
!(s.length == 1 && is_plain(s.charAt(0))) && malformed_symbol.pattern.matcher(s).matches |
| 34137 | 72 |
|
73 |
class Matcher(text: CharSequence) |
|
74 |
{
|
|
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
75 |
private val matcher = regex_total.pattern.matcher(text) |
| 34137 | 76 |
def apply(start: Int, end: Int): Int = |
77 |
{
|
|
78 |
require(0 <= start && start < end && end <= text.length) |
|
|
34316
f879b649ac4c
clarified Symbol.is_plain/is_wellformed -- is_closed was rejecting plain backslashes;
wenzelm
parents:
34193
diff
changeset
|
79 |
if (is_plain(text.charAt(start))) 1 |
| 34138 | 80 |
else {
|
| 34137 | 81 |
matcher.region(start, end).lookingAt |
82 |
matcher.group.length |
|
83 |
} |
|
84 |
} |
|
| 31522 | 85 |
} |
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
86 |
|
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
87 |
|
| 43489 | 88 |
/* efficient iterators */ |
| 33998 | 89 |
|
| 43489 | 90 |
def iterator(text: CharSequence): Iterator[CharSequence] = |
91 |
new Iterator[CharSequence] |
|
| 40522 | 92 |
{
|
| 43489 | 93 |
private val matcher = new Matcher(text) |
94 |
private var i = 0 |
|
95 |
def hasNext = i < text.length |
|
96 |
def next = |
|
97 |
{
|
|
98 |
val n = matcher(i, text.length) |
|
99 |
val s = text.subSequence(i, i + n) |
|
100 |
i += n |
|
101 |
s |
|
102 |
} |
|
| 33998 | 103 |
} |
| 43489 | 104 |
|
105 |
private val char_symbols: Array[String] = |
|
| 43490 | 106 |
(0 until 128).iterator.map(i => new String(Array(i.toChar))).toArray |
| 43489 | 107 |
|
108 |
private def make_string(sym: CharSequence): String = |
|
109 |
sym.length match {
|
|
110 |
case 0 => "" |
|
111 |
case 1 => |
|
112 |
val c = sym.charAt(0) |
|
113 |
if (c < char_symbols.length) char_symbols(c) |
|
114 |
else sym.toString |
|
115 |
case _ => sym.toString |
|
116 |
} |
|
117 |
||
118 |
def iterator_string(text: CharSequence): Iterator[String] = |
|
119 |
iterator(text).map(make_string) |
|
| 33998 | 120 |
|
121 |
||
122 |
/* decoding offsets */ |
|
123 |
||
124 |
class Index(text: CharSequence) |
|
| 31929 | 125 |
{
|
126 |
case class Entry(chr: Int, sym: Int) |
|
127 |
val index: Array[Entry] = |
|
128 |
{
|
|
| 34137 | 129 |
val matcher = new Matcher(text) |
| 31929 | 130 |
val buf = new mutable.ArrayBuffer[Entry] |
131 |
var chr = 0 |
|
132 |
var sym = 0 |
|
| 33998 | 133 |
while (chr < text.length) {
|
| 34137 | 134 |
val n = matcher(chr, text.length) |
135 |
chr += n |
|
| 31929 | 136 |
sym += 1 |
| 34137 | 137 |
if (n > 1) buf += Entry(chr, sym) |
| 31929 | 138 |
} |
139 |
buf.toArray |
|
140 |
} |
|
|
38479
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
wenzelm
parents:
37556
diff
changeset
|
141 |
def decode(sym1: Int): Int = |
| 31929 | 142 |
{
|
|
38479
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
wenzelm
parents:
37556
diff
changeset
|
143 |
val sym = sym1 - 1 |
| 31929 | 144 |
val end = index.length |
145 |
def bisect(a: Int, b: Int): Int = |
|
146 |
{
|
|
147 |
if (a < b) {
|
|
148 |
val c = (a + b) / 2 |
|
149 |
if (sym < index(c).sym) bisect(a, c) |
|
150 |
else if (c + 1 == end || sym < index(c + 1).sym) c |
|
151 |
else bisect(c + 1, b) |
|
152 |
} |
|
153 |
else -1 |
|
154 |
} |
|
155 |
val i = bisect(0, end) |
|
156 |
if (i < 0) sym |
|
157 |
else index(i).chr + sym - index(i).sym |
|
158 |
} |
|
|
38479
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
wenzelm
parents:
37556
diff
changeset
|
159 |
def decode(range: Text.Range): Text.Range = range.map(decode(_)) |
| 31929 | 160 |
} |
161 |
||
162 |
||
| 33998 | 163 |
/* recoding text */ |
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
164 |
|
| 31522 | 165 |
private class Recoder(list: List[(String, String)]) |
166 |
{
|
|
167 |
private val (min, max) = |
|
168 |
{
|
|
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
169 |
var min = '\uffff' |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
170 |
var max = '\u0000' |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
171 |
for ((x, _) <- list) {
|
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
172 |
val c = x(0) |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
173 |
if (c < min) min = c |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
174 |
if (c > max) max = c |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
175 |
} |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
176 |
(min, max) |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
177 |
} |
| 40443 | 178 |
private val table = |
179 |
{
|
|
180 |
var tab = Map[String, String]() |
|
181 |
for ((x, y) <- list) {
|
|
182 |
tab.get(x) match {
|
|
183 |
case None => tab += (x -> y) |
|
184 |
case Some(z) => |
|
185 |
error("Duplicate mapping of \"" + x + "\" to \"" + y + "\" vs. \"" + z + "\"")
|
|
186 |
} |
|
187 |
} |
|
188 |
tab |
|
189 |
} |
|
| 31522 | 190 |
def recode(text: String): String = |
191 |
{
|
|
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
192 |
val len = text.length |
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
193 |
val matcher = regex_total.pattern.matcher(text) |
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
194 |
val result = new StringBuilder(len) |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
195 |
var i = 0 |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
196 |
while (i < len) {
|
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
197 |
val c = text(i) |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
198 |
if (min <= c && c <= max) {
|
| 31929 | 199 |
matcher.region(i, len).lookingAt |
| 27938 | 200 |
val x = matcher.group |
| 31522 | 201 |
result.append(table.get(x) getOrElse x) |
|
27937
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
202 |
i = matcher.end |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
203 |
} |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
204 |
else { result.append(c); i += 1 }
|
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
205 |
} |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
206 |
result.toString |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
207 |
} |
|
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
wenzelm
parents:
27935
diff
changeset
|
208 |
} |
| 27924 | 209 |
|
| 27918 | 210 |
|
|
27923
7ebe9d38743a
use scala.collection.jcl.HashMap, which seems to be more efficient;
wenzelm
parents:
27918
diff
changeset
|
211 |
|
| 27927 | 212 |
/** Symbol interpretation **/ |
213 |
||
| 34137 | 214 |
class Interpretation(symbol_decls: List[String]) |
|
29569
f3f529b5d8fb
more general init of Symbol.Interpretation, independent of IsabelleSystem instance;
wenzelm
parents:
29174
diff
changeset
|
215 |
{
|
| 31522 | 216 |
/* read symbols */ |
217 |
||
218 |
private val empty = new Regex("""(?xs) ^\s* (?: \#.* )? $ """)
|
|
219 |
private val key = new Regex("""(?xs) (.+): """)
|
|
220 |
||
221 |
private def read_decl(decl: String): (String, Map[String, String]) = |
|
222 |
{
|
|
223 |
def err() = error("Bad symbol declaration: " + decl)
|
|
224 |
||
225 |
def read_props(props: List[String]): Map[String, String] = |
|
226 |
{
|
|
227 |
props match {
|
|
228 |
case Nil => Map() |
|
229 |
case _ :: Nil => err() |
|
230 |
case key(x) :: y :: rest => read_props(rest) + (x -> y) |
|
231 |
case _ => err() |
|
232 |
} |
|
233 |
} |
|
234 |
decl.split("\\s+").toList match {
|
|
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
235 |
case sym :: props if sym.length > 1 && !is_malformed(sym) => (sym, read_props(props)) |
| 34193 | 236 |
case _ => err() |
| 31522 | 237 |
} |
238 |
} |
|
239 |
||
240 |
private val symbols: List[(String, Map[String, String])] = |
|
| 40443 | 241 |
Map(( |
242 |
for (decl <- symbol_decls if !empty.pattern.matcher(decl).matches) |
|
243 |
yield read_decl(decl)): _*) toList |
|
| 31522 | 244 |
|
245 |
||
| 31651 | 246 |
/* misc properties */ |
247 |
||
| 34134 | 248 |
val names: Map[String, String] = |
249 |
{
|
|
|
43456
8a6de1a6e1dc
names for control symbols without "^", which is relevant for completion;
wenzelm
parents:
43455
diff
changeset
|
250 |
val name = new Regex("""\\<\^?([A-Za-z][A-Za-z0-9_']*)>""")
|
| 31651 | 251 |
Map((for ((sym @ name(a), _) <- symbols) yield (sym -> a)): _*) |
252 |
} |
|
253 |
||
| 43488 | 254 |
val abbrevs: Map[String, String] = |
255 |
Map(( |
|
256 |
for ((sym, props) <- symbols if props.isDefinedAt("abbrev"))
|
|
257 |
yield (sym -> props("abbrev"))): _*)
|
|
258 |
||
259 |
||
| 43490 | 260 |
/* recoding */ |
| 31522 | 261 |
|
262 |
private val (decoder, encoder) = |
|
263 |
{
|
|
264 |
val mapping = |
|
265 |
for {
|
|
266 |
(sym, props) <- symbols |
|
267 |
val code = |
|
268 |
try { Integer.decode(props("code")).intValue }
|
|
269 |
catch {
|
|
270 |
case _: NoSuchElementException => error("Missing code for symbol " + sym)
|
|
271 |
case _: NumberFormatException => error("Bad code for symbol " + sym)
|
|
272 |
} |
|
273 |
val ch = new String(Character.toChars(code)) |
|
| 34193 | 274 |
} yield {
|
275 |
if (code < 128) error("Illegal ASCII code for symbol " + sym)
|
|
276 |
else (sym, ch) |
|
277 |
} |
|
|
31545
5f1f0a20af4d
discontinued escaped symbols such as \\<forall> -- only one backslash should be used;
wenzelm
parents:
31523
diff
changeset
|
278 |
(new Recoder(mapping), |
| 31548 | 279 |
new Recoder(mapping map { case (x, y) => (y, x) }))
|
| 31522 | 280 |
} |
| 27918 | 281 |
|
| 34098 | 282 |
def decode(text: String): String = decoder.recode(text) |
283 |
def encode(text: String): String = encoder.recode(text) |
|
| 34134 | 284 |
|
| 43490 | 285 |
private def recode_set(elems: String*): Set[String] = |
286 |
{
|
|
287 |
val content = elems.toList |
|
288 |
Set((content ::: content.map(decode)): _*) |
|
289 |
} |
|
290 |
||
291 |
private def recode_map[A](elems: (String, A)*): Map[String, A] = |
|
292 |
{
|
|
293 |
val content = elems.toList |
|
294 |
Map((content ::: content.map({ case (sym, a) => (decode(sym), a) })): _*)
|
|
295 |
} |
|
296 |
||
297 |
||
298 |
/* user fonts */ |
|
299 |
||
300 |
val fonts: Map[String, String] = |
|
301 |
recode_map(( |
|
302 |
for ((sym, props) <- symbols if props.isDefinedAt("font"))
|
|
303 |
yield (sym -> props("font"))): _*)
|
|
304 |
||
305 |
val font_names: List[String] = Set(fonts.toList.map(_._2): _*).toList |
|
306 |
val font_index: Map[String, Int] = Map((font_names zip (0 until font_names.length).toList): _*) |
|
307 |
||
308 |
def lookup_font(sym: String): Option[Int] = fonts.get(sym).map(font_index(_)) |
|
309 |
||
| 34134 | 310 |
|
311 |
/* classification */ |
|
312 |
||
| 43490 | 313 |
private val letters = recode_set( |
| 34134 | 314 |
"A", "B", "C", "D", "E", "F", "G", "H", "I", "J", "K", "L", "M", |
315 |
"N", "O", "P", "Q", "R", "S", "T", "U", "V", "W", "X", "Y", "Z", |
|
316 |
"a", "b", "c", "d", "e", "f", "g", "h", "i", "j", "k", "l", "m", |
|
317 |
"n", "o", "p", "q", "r", "s", "t", "u", "v", "w", "x", "y", "z", |
|
318 |
||
319 |
"\\<A>", "\\<B>", "\\<C>", "\\<D>", "\\<E>", "\\<F>", "\\<G>", |
|
320 |
"\\<H>", "\\<I>", "\\<J>", "\\<K>", "\\<L>", "\\<M>", "\\<N>", |
|
321 |
"\\<O>", "\\<P>", "\\<Q>", "\\<R>", "\\<S>", "\\<T>", "\\<U>", |
|
322 |
"\\<V>", "\\<W>", "\\<X>", "\\<Y>", "\\<Z>", "\\<a>", "\\<b>", |
|
323 |
"\\<c>", "\\<d>", "\\<e>", "\\<f>", "\\<g>", "\\<h>", "\\<i>", |
|
324 |
"\\<j>", "\\<k>", "\\<l>", "\\<m>", "\\<n>", "\\<o>", "\\<p>", |
|
325 |
"\\<q>", "\\<r>", "\\<s>", "\\<t>", "\\<u>", "\\<v>", "\\<w>", |
|
326 |
"\\<x>", "\\<y>", "\\<z>", |
|
327 |
||
328 |
"\\<AA>", "\\<BB>", "\\<CC>", "\\<DD>", "\\<EE>", "\\<FF>", |
|
329 |
"\\<GG>", "\\<HH>", "\\<II>", "\\<JJ>", "\\<KK>", "\\<LL>", |
|
330 |
"\\<MM>", "\\<NN>", "\\<OO>", "\\<PP>", "\\<QQ>", "\\<RR>", |
|
331 |
"\\<SS>", "\\<TT>", "\\<UU>", "\\<VV>", "\\<WW>", "\\<XX>", |
|
332 |
"\\<YY>", "\\<ZZ>", "\\<aa>", "\\<bb>", "\\<cc>", "\\<dd>", |
|
333 |
"\\<ee>", "\\<ff>", "\\<gg>", "\\<hh>", "\\<ii>", "\\<jj>", |
|
334 |
"\\<kk>", "\\<ll>", "\\<mm>", "\\<nn>", "\\<oo>", "\\<pp>", |
|
335 |
"\\<qq>", "\\<rr>", "\\<ss>", "\\<tt>", "\\<uu>", "\\<vv>", |
|
336 |
"\\<ww>", "\\<xx>", "\\<yy>", "\\<zz>", |
|
337 |
||
338 |
"\\<alpha>", "\\<beta>", "\\<gamma>", "\\<delta>", "\\<epsilon>", |
|
339 |
"\\<zeta>", "\\<eta>", "\\<theta>", "\\<iota>", "\\<kappa>", |
|
340 |
"\\<mu>", "\\<nu>", "\\<xi>", "\\<pi>", "\\<rho>", "\\<sigma>", |
|
341 |
"\\<tau>", "\\<upsilon>", "\\<phi>", "\\<chi>", "\\<psi>", |
|
342 |
"\\<omega>", "\\<Gamma>", "\\<Delta>", "\\<Theta>", "\\<Lambda>", |
|
343 |
"\\<Xi>", "\\<Pi>", "\\<Sigma>", "\\<Upsilon>", "\\<Phi>", |
|
344 |
"\\<Psi>", "\\<Omega>", |
|
345 |
||
346 |
"\\<^isub>", "\\<^isup>") |
|
347 |
||
| 34138 | 348 |
private val blanks = |
| 43490 | 349 |
recode_set(space, "\t", "\n", "\u000B", "\f", "\r", "\\<spacespace>", "\\<^newline>") |
| 34138 | 350 |
|
351 |
private val sym_chars = |
|
352 |
Set("!", "#", "$", "%", "&", "*", "+", "-", "/", "<", "=", ">", "?", "@", "^", "_", "|", "~")
|
|
| 34134 | 353 |
|
354 |
def is_letter(sym: String): Boolean = letters.contains(sym) |
|
| 34138 | 355 |
def is_digit(sym: String): Boolean = sym.length == 1 && '0' <= sym(0) && sym(0) <= '9' |
| 34134 | 356 |
def is_quasi(sym: String): Boolean = sym == "_" || sym == "'" |
| 34138 | 357 |
def is_letdig(sym: String): Boolean = is_letter(sym) || is_digit(sym) || is_quasi(sym) |
| 34134 | 358 |
def is_blank(sym: String): Boolean = blanks.contains(sym) |
| 34138 | 359 |
def is_symbolic_char(sym: String): Boolean = sym_chars.contains(sym) |
|
40523
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
360 |
def is_symbolic(sym: String): Boolean = |
|
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents:
40522
diff
changeset
|
361 |
sym.startsWith("\\<") && sym.endsWith(">") && !sym.startsWith("\\<^")
|
| 43455 | 362 |
|
363 |
||
| 43488 | 364 |
/* control symbols */ |
365 |
||
366 |
private val ctrl_decoded: Set[String] = |
|
367 |
Set((for ((sym, _) <- symbols if sym.startsWith("\\<^")) yield decode(sym)): _*)
|
|
368 |
||
369 |
def is_ctrl(sym: String): Boolean = |
|
370 |
sym.startsWith("\\<^") || ctrl_decoded.contains(sym)
|
|
| 43455 | 371 |
|
|
43443
5d9693c2337e
basic support for extended syntax styles: sub/superscript;
wenzelm
parents:
43418
diff
changeset
|
372 |
def is_controllable(sym: String): Boolean = |
| 43488 | 373 |
!is_blank(sym) && !is_ctrl(sym) && !is_malformed(sym) |
| 43455 | 374 |
|
375 |
private val subscript_decoded = Set(decode("\\<^sub>"), decode("\\<^isub>"))
|
|
376 |
private val superscript_decoded = Set(decode("\\<^sup>"), decode("\\<^isup>"))
|
|
377 |
def is_subscript_decoded(sym: String): Boolean = subscript_decoded.contains(sym) |
|
378 |
def is_superscript_decoded(sym: String): Boolean = superscript_decoded.contains(sym) |
|
| 43511 | 379 |
|
380 |
val bold_decoded = decode("\\<^bold>")
|
|
381 |
val bsub_decoded = decode("\\<^bsub>")
|
|
382 |
val esub_decoded = decode("\\<^esub>")
|
|
383 |
val bsup_decoded = decode("\\<^bsup>")
|
|
384 |
val esup_decoded = decode("\\<^esup>")
|
|
| 27918 | 385 |
} |
| 27901 | 386 |
} |