| author | wenzelm | 
| Fri, 25 Apr 2014 14:18:13 +0200 | |
| changeset 56720 | e1317a26f8c0 | 
| parent 56472 | f4abde849190 | 
| child 56746 | d37a5d09a277 | 
| 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 | ||
| 55618 | 9 | |
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 10 | import scala.collection.mutable | 
| 31522 | 11 | import scala.util.matching.Regex | 
| 48922 | 12 | import scala.annotation.tailrec | 
| 27901 | 13 | |
| 14 | ||
| 31522 | 15 | object Symbol | 
| 16 | {
 | |
| 43696 | 17 | type Symbol = String | 
| 18 | ||
| 55884 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 19 | // counting Isabelle symbols, starting from 1 | 
| 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 20 | type Offset = Text.Offset | 
| 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 21 | type Range = Text.Range | 
| 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 22 | |
| 43696 | 23 | |
| 43418 | 24 | /* ASCII characters */ | 
| 25 | ||
| 26 | def is_ascii_letter(c: Char): Boolean = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' | |
| 55497 | 27 | |
| 43418 | 28 | def is_ascii_digit(c: Char): Boolean = '0' <= c && c <= '9' | 
| 55497 | 29 | |
| 30 | def is_ascii_hex(c: Char): Boolean = | |
| 31 | '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f' | |
| 32 | ||
| 43418 | 33 | def is_ascii_quasi(c: Char): Boolean = c == '_' || c == '\'' | 
| 34 | ||
| 55497 | 35 | def is_ascii_blank(c: Char): Boolean = " \t\n\u000b\f\r".contains(c) | 
| 36 | ||
| 43418 | 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 = | |
| 50238 
98d35a7368bd
more uniform Symbol.is_ascii_identifier in ML/Scala;
 wenzelm parents: 
50233diff
changeset | 41 | s.length > 0 && is_ascii_letter(s(0)) && s.forall(is_ascii_letdig) | 
| 43418 | 42 | |
| 43 | ||
| 48775 | 44 | /* symbol matching */ | 
| 27901 | 45 | |
| 48775 | 46 |   private val symbol_total = new Regex("""(?xs)
 | 
| 47 | [\ud800-\udbff][\udc00-\udfff] | \r\n | | |
| 48 | \\ < (?: \^raw: [\x20-\x7e\u0100-\uffff && [^.>]]* | \^? ([A-Za-z][A-Za-z0-9_']*)? ) >? | | |
| 49 | .""") | |
| 27924 | 50 | |
| 48775 | 51 | private def is_plain(c: Char): Boolean = | 
| 52 | !(c == '\r' || c == '\\' || Character.isHighSurrogate(c)) | |
| 48773 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 53 | |
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 54 | def is_malformed(s: Symbol): Boolean = | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 55 |     s.length match {
 | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 56 | case 1 => | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 57 | val c = s(0) | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 58 | Character.isHighSurrogate(c) || Character.isLowSurrogate(c) || c == '\ufffd' | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 59 | case 2 => | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 60 | val c1 = s(0) | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 61 | val c2 = s(1) | 
| 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 62 | !(c1 == '\r' && c2 == '\n' || Character.isSurrogatePair(c1, c2)) | 
| 48774 | 63 |       case _ => !s.endsWith(">") || s == "\\<>" || s == "\\<^>"
 | 
| 48773 
0e1bab274672
more liberal scanning of potentially malformed symbols;
 wenzelm parents: 
48704diff
changeset | 64 | } | 
| 34137 | 65 | |
| 54734 
b91afc3aa3e6
clarified Proof General legacy: special treatment of \<^newline> only in TTY mode;
 wenzelm parents: 
53400diff
changeset | 66 | def is_newline(s: Symbol): Boolean = | 
| 43675 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 67 | s == "\n" || s == "\r" || s == "\r\n" | 
| 38877 | 68 | |
| 34137 | 69 | class Matcher(text: CharSequence) | 
| 70 |   {
 | |
| 48775 | 71 | private val matcher = symbol_total.pattern.matcher(text) | 
| 34137 | 72 | def apply(start: Int, end: Int): Int = | 
| 73 |     {
 | |
| 74 | require(0 <= start && start < end && end <= text.length) | |
| 34316 
f879b649ac4c
clarified Symbol.is_plain/is_wellformed -- is_closed was rejecting plain backslashes;
 wenzelm parents: 
34193diff
changeset | 75 | if (is_plain(text.charAt(start))) 1 | 
| 34138 | 76 |       else {
 | 
| 34137 | 77 | matcher.region(start, end).lookingAt | 
| 78 | matcher.group.length | |
| 79 | } | |
| 80 | } | |
| 31522 | 81 | } | 
| 27937 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 82 | |
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 83 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 84 | /* iterator */ | 
| 33998 | 85 | |
| 43696 | 86 | private val char_symbols: Array[Symbol] = | 
| 43675 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 87 | (0 until 256).iterator.map(i => new String(Array(i.toChar))).toArray | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 88 | |
| 43696 | 89 | def iterator(text: CharSequence): Iterator[Symbol] = | 
| 90 | new Iterator[Symbol] | |
| 40522 | 91 |     {
 | 
| 43489 | 92 | private val matcher = new Matcher(text) | 
| 93 | private var i = 0 | |
| 94 | def hasNext = i < text.length | |
| 95 | def next = | |
| 96 |       {
 | |
| 97 | val n = matcher(i, text.length) | |
| 43675 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 98 | val s = | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 99 | if (n == 0) "" | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 100 |           else if (n == 1) {
 | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 101 | val c = text.charAt(i) | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 102 | if (c < char_symbols.length) char_symbols(c) | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 103 | else text.subSequence(i, i + n).toString | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 104 | } | 
| 
8252d51d70e2
simplified Symbol.iterator: produce strings, which are mostly preallocated;
 wenzelm parents: 
43511diff
changeset | 105 | else text.subSequence(i, i + n).toString | 
| 43489 | 106 | i += n | 
| 107 | s | |
| 108 | } | |
| 33998 | 109 | } | 
| 43489 | 110 | |
| 44949 | 111 | def explode(text: CharSequence): List[Symbol] = iterator(text).toList | 
| 112 | ||
| 48922 | 113 | def advance_line_column(pos: (Int, Int), text: CharSequence): (Int, Int) = | 
| 114 |   {
 | |
| 115 | var (line, column) = pos | |
| 116 |     for (sym <- iterator(text)) {
 | |
| 54734 
b91afc3aa3e6
clarified Proof General legacy: special treatment of \<^newline> only in TTY mode;
 wenzelm parents: 
53400diff
changeset | 117 |       if (is_newline(sym)) { line += 1; column = 1 }
 | 
| 48922 | 118 | else column += 1 | 
| 119 | } | |
| 120 | (line, column) | |
| 121 | } | |
| 122 | ||
| 33998 | 123 | |
| 124 | /* decoding offsets */ | |
| 125 | ||
| 52507 | 126 | object Index | 
| 127 |   {
 | |
| 56471 
2293a4350716
more frugal Symbol.Index -- no need to waste space on mostly empty arrays;
 wenzelm parents: 
56338diff
changeset | 128 | private sealed case class Entry(chr: Int, sym: Int) | 
| 52507 | 129 | |
| 56472 | 130 | val empty: Index = new Index(Nil) | 
| 56471 
2293a4350716
more frugal Symbol.Index -- no need to waste space on mostly empty arrays;
 wenzelm parents: 
56338diff
changeset | 131 | |
| 
2293a4350716
more frugal Symbol.Index -- no need to waste space on mostly empty arrays;
 wenzelm parents: 
56338diff
changeset | 132 | def apply(text: CharSequence): Index = | 
| 31929 | 133 |     {
 | 
| 34137 | 134 | val matcher = new Matcher(text) | 
| 56471 
2293a4350716
more frugal Symbol.Index -- no need to waste space on mostly empty arrays;
 wenzelm parents: 
56338diff
changeset | 135 | val buf = new mutable.ListBuffer[Entry] | 
| 31929 | 136 | var chr = 0 | 
| 137 | var sym = 0 | |
| 33998 | 138 |       while (chr < text.length) {
 | 
| 34137 | 139 | val n = matcher(chr, text.length) | 
| 140 | chr += n | |
| 31929 | 141 | sym += 1 | 
| 34137 | 142 | if (n > 1) buf += Entry(chr, sym) | 
| 31929 | 143 | } | 
| 56472 | 144 | if (buf.isEmpty) empty else new Index(buf.toList) | 
| 31929 | 145 | } | 
| 56471 
2293a4350716
more frugal Symbol.Index -- no need to waste space on mostly empty arrays;
 wenzelm parents: 
56338diff
changeset | 146 | } | 
| 55430 | 147 | |
| 56472 | 148 | final class Index private(entries: List[Index.Entry]) | 
| 56471 
2293a4350716
more frugal Symbol.Index -- no need to waste space on mostly empty arrays;
 wenzelm parents: 
56338diff
changeset | 149 |   {
 | 
| 56472 | 150 | private val hash: Int = entries.hashCode | 
| 151 | private val index: Array[Index.Entry] = entries.toArray | |
| 152 | ||
| 55884 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 153 | def decode(symbol_offset: Offset): Text.Offset = | 
| 31929 | 154 |     {
 | 
| 55884 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 155 | val sym = symbol_offset - 1 | 
| 31929 | 156 | val end = index.length | 
| 48922 | 157 | @tailrec def bisect(a: Int, b: Int): Int = | 
| 31929 | 158 |       {
 | 
| 159 |         if (a < b) {
 | |
| 160 | val c = (a + b) / 2 | |
| 161 | if (sym < index(c).sym) bisect(a, c) | |
| 162 | else if (c + 1 == end || sym < index(c + 1).sym) c | |
| 163 | else bisect(c + 1, b) | |
| 164 | } | |
| 165 | else -1 | |
| 166 | } | |
| 167 | val i = bisect(0, end) | |
| 168 | if (i < 0) sym | |
| 169 | else index(i).chr + sym - index(i).sym | |
| 170 | } | |
| 55884 
f2c0eaedd579
tuned signature -- emphasize symbol positions (prover) vs. decoded text offsets (editor);
 wenzelm parents: 
55618diff
changeset | 171 | def decode(symbol_range: Range): Text.Range = symbol_range.map(decode(_)) | 
| 56335 
8953d4cc060a
store blob content within document node: aux. files that were once open are made persistent;
 wenzelm parents: 
55884diff
changeset | 172 | |
| 56338 
f968f4e3d520
proper structural hashCode, which is required for Command.File equals (NB: Array has physical object identity);
 wenzelm parents: 
56335diff
changeset | 173 | override def hashCode: Int = hash | 
| 56335 
8953d4cc060a
store blob content within document node: aux. files that were once open are made persistent;
 wenzelm parents: 
55884diff
changeset | 174 | override def equals(that: Any): Boolean = | 
| 
8953d4cc060a
store blob content within document node: aux. files that were once open are made persistent;
 wenzelm parents: 
55884diff
changeset | 175 |       that match {
 | 
| 
8953d4cc060a
store blob content within document node: aux. files that were once open are made persistent;
 wenzelm parents: 
55884diff
changeset | 176 | case other: Index => index.sameElements(other.index) | 
| 
8953d4cc060a
store blob content within document node: aux. files that were once open are made persistent;
 wenzelm parents: 
55884diff
changeset | 177 | case _ => false | 
| 
8953d4cc060a
store blob content within document node: aux. files that were once open are made persistent;
 wenzelm parents: 
55884diff
changeset | 178 | } | 
| 31929 | 179 | } | 
| 180 | ||
| 181 | ||
| 33998 | 182 | /* recoding text */ | 
| 27937 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 183 | |
| 31522 | 184 | private class Recoder(list: List[(String, String)]) | 
| 185 |   {
 | |
| 186 | private val (min, max) = | |
| 187 |     {
 | |
| 27937 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 188 | var min = '\uffff' | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 189 | var max = '\u0000' | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 190 |       for ((x, _) <- list) {
 | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 191 | val c = x(0) | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 192 | if (c < min) min = c | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 193 | if (c > max) max = c | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 194 | } | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 195 | (min, max) | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 196 | } | 
| 40443 | 197 | private val table = | 
| 198 |     {
 | |
| 199 | var tab = Map[String, String]() | |
| 200 |       for ((x, y) <- list) {
 | |
| 201 |         tab.get(x) match {
 | |
| 202 | case None => tab += (x -> y) | |
| 203 | case Some(z) => | |
| 44181 | 204 |             error("Duplicate mapping of " + quote(x) + " to " + quote(y) + " vs. " + quote(z))
 | 
| 40443 | 205 | } | 
| 206 | } | |
| 207 | tab | |
| 208 | } | |
| 31522 | 209 | def recode(text: String): String = | 
| 210 |     {
 | |
| 27937 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 211 | val len = text.length | 
| 48775 | 212 | val matcher = symbol_total.pattern.matcher(text) | 
| 27937 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 213 | val result = new StringBuilder(len) | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 214 | var i = 0 | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 215 |       while (i < len) {
 | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 216 | val c = text(i) | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 217 |         if (min <= c && c <= max) {
 | 
| 31929 | 218 | matcher.region(i, len).lookingAt | 
| 27938 | 219 | val x = matcher.group | 
| 52888 | 220 | result.append(table.getOrElse(x, x)) | 
| 27937 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 221 | i = matcher.end | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 222 | } | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 223 |         else { result.append(c); i += 1 }
 | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 224 | } | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 225 | result.toString | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 226 | } | 
| 
fdf77e7be01a
more robust pattern: look at longer matches first, added catch-all case;
 wenzelm parents: 
27935diff
changeset | 227 | } | 
| 27924 | 228 | |
| 27918 | 229 | |
| 27923 
7ebe9d38743a
use scala.collection.jcl.HashMap, which seems to be more efficient;
 wenzelm parents: 
27918diff
changeset | 230 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 231 | /** symbol interpretation **/ | 
| 27927 | 232 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 233 | private lazy val symbols = | 
| 50564 
c6fde2fc4217
allow to suppress ISABELLE_SYMBOLS for experiments;
 wenzelm parents: 
50291diff
changeset | 234 |     new Interpretation(File.try_read(Path.split(Isabelle_System.getenv("ISABELLE_SYMBOLS"))))
 | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 235 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 236 | private class Interpretation(symbols_spec: String) | 
| 29569 
f3f529b5d8fb
more general init of Symbol.Interpretation, independent of IsabelleSystem instance;
 wenzelm parents: 
29174diff
changeset | 237 |   {
 | 
| 31522 | 238 | /* read symbols */ | 
| 239 | ||
| 50136 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 240 |     private val No_Decl = new Regex("""(?xs) ^\s* (?: \#.* )? $ """)
 | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 241 |     private val Key = new Regex("""(?xs) (.+): """)
 | 
| 31522 | 242 | |
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 243 | private def read_decl(decl: String): (Symbol, Properties.T) = | 
| 31522 | 244 |     {
 | 
| 245 |       def err() = error("Bad symbol declaration: " + decl)
 | |
| 246 | ||
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 247 | def read_props(props: List[String]): Properties.T = | 
| 31522 | 248 |       {
 | 
| 249 |         props match {
 | |
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 250 | case Nil => Nil | 
| 31522 | 251 | case _ :: Nil => err() | 
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 252 | case Key(x) :: y :: rest => (x -> y) :: read_props(rest) | 
| 31522 | 253 | case _ => err() | 
| 254 | } | |
| 255 | } | |
| 256 |       decl.split("\\s+").toList match {
 | |
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 257 | case sym :: props if sym.length > 1 && !is_malformed(sym) => | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 258 | (sym, read_props(props)) | 
| 34193 | 259 | case _ => err() | 
| 31522 | 260 | } | 
| 261 | } | |
| 262 | ||
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 263 | private val symbols: List[(Symbol, Properties.T)] = | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 264 | (((List.empty[(Symbol, Properties.T)], Set.empty[Symbol]) /: | 
| 50136 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 265 | split_lines(symbols_spec).reverse) | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 266 |         { case (res, No_Decl()) => res
 | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 267 | case ((list, known), decl) => | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 268 | val (sym, props) = read_decl(decl) | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 269 | if (known(sym)) (list, known) | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 270 | else ((sym, props) :: list, known + sym) | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 271 | })._1 | 
| 31522 | 272 | |
| 273 | ||
| 53400 | 274 | /* basic properties */ | 
| 275 | ||
| 276 | val properties: Map[Symbol, Properties.T] = Map(symbols: _*) | |
| 31651 | 277 | |
| 43696 | 278 | val names: Map[Symbol, String] = | 
| 34134 | 279 |     {
 | 
| 43456 
8a6de1a6e1dc
names for control symbols without "^", which is relevant for completion;
 wenzelm parents: 
43455diff
changeset | 280 |       val name = new Regex("""\\<\^?([A-Za-z][A-Za-z0-9_']*)>""")
 | 
| 31651 | 281 | Map((for ((sym @ name(a), _) <- symbols) yield (sym -> a)): _*) | 
| 282 | } | |
| 283 | ||
| 50136 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 284 | val groups: List[(String, List[Symbol])] = | 
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 285 |       symbols.map({ case (sym, props) =>
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 286 |         val gs = for (("group", g) <- props) yield g
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 287 | if (gs.isEmpty) List(sym -> "unsorted") else gs.map(sym -> _) | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 288 | }).flatten | 
| 50136 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 289 |         .groupBy(_._2).toList.map({ case (group, list) => (group, list.map(_._1)) })
 | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 290 | .sortBy(_._1) | 
| 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 291 | |
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 292 | val abbrevs: Multi_Map[Symbol, String] = | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 293 | Multi_Map(( | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 294 |         for {
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 295 | (sym, props) <- symbols | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 296 |           ("abbrev", a) <- props.reverse
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 297 | } yield (sym -> a)): _*) | 
| 43488 | 298 | |
| 299 | ||
| 43490 | 300 | /* recoding */ | 
| 31522 | 301 | |
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 302 |     private val Code = new Properties.String("code")
 | 
| 31522 | 303 | private val (decoder, encoder) = | 
| 304 |     {
 | |
| 305 | val mapping = | |
| 306 |         for {
 | |
| 307 | (sym, props) <- symbols | |
| 46997 | 308 | code = | 
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 309 |             props match {
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 310 | case Code(s) => | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 311 |                 try { Integer.decode(s).intValue }
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 312 |                 catch { case _: NumberFormatException => error("Bad code for symbol " + sym) }
 | 
| 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 313 |               case _ => error("Missing code for symbol " + sym)
 | 
| 31522 | 314 | } | 
| 46997 | 315 | ch = new String(Character.toChars(code)) | 
| 34193 | 316 |         } yield {
 | 
| 317 |           if (code < 128) error("Illegal ASCII code for symbol " + sym)
 | |
| 318 | else (sym, ch) | |
| 319 | } | |
| 31545 
5f1f0a20af4d
discontinued escaped symbols such as \\<forall> -- only one backslash should be used;
 wenzelm parents: 
31523diff
changeset | 320 | (new Recoder(mapping), | 
| 31548 | 321 |        new Recoder(mapping map { case (x, y) => (y, x) }))
 | 
| 31522 | 322 | } | 
| 27918 | 323 | |
| 34098 | 324 | def decode(text: String): String = decoder.recode(text) | 
| 325 | def encode(text: String): String = encoder.recode(text) | |
| 34134 | 326 | |
| 43490 | 327 | private def recode_set(elems: String*): Set[String] = | 
| 328 |     {
 | |
| 329 | val content = elems.toList | |
| 330 | Set((content ::: content.map(decode)): _*) | |
| 331 | } | |
| 332 | ||
| 333 | private def recode_map[A](elems: (String, A)*): Map[String, A] = | |
| 334 |     {
 | |
| 335 | val content = elems.toList | |
| 336 |       Map((content ::: content.map({ case (sym, a) => (decode(sym), a) })): _*)
 | |
| 337 | } | |
| 338 | ||
| 339 | ||
| 340 | /* user fonts */ | |
| 341 | ||
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 342 |     private val Font = new Properties.String("font")
 | 
| 43696 | 343 | val fonts: Map[Symbol, String] = | 
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 344 | recode_map((for ((sym, Font(font)) <- symbols) yield (sym -> font)): _*) | 
| 43490 | 345 | |
| 346 | val font_names: List[String] = Set(fonts.toList.map(_._2): _*).toList | |
| 347 | val font_index: Map[String, Int] = Map((font_names zip (0 until font_names.length).toList): _*) | |
| 348 | ||
| 34134 | 349 | |
| 350 | /* classification */ | |
| 351 | ||
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 352 | val letters = recode_set( | 
| 34134 | 353 | "A", "B", "C", "D", "E", "F", "G", "H", "I", "J", "K", "L", "M", | 
| 354 | "N", "O", "P", "Q", "R", "S", "T", "U", "V", "W", "X", "Y", "Z", | |
| 355 | "a", "b", "c", "d", "e", "f", "g", "h", "i", "j", "k", "l", "m", | |
| 356 | "n", "o", "p", "q", "r", "s", "t", "u", "v", "w", "x", "y", "z", | |
| 357 | ||
| 358 | "\\<A>", "\\<B>", "\\<C>", "\\<D>", "\\<E>", "\\<F>", "\\<G>", | |
| 359 | "\\<H>", "\\<I>", "\\<J>", "\\<K>", "\\<L>", "\\<M>", "\\<N>", | |
| 360 | "\\<O>", "\\<P>", "\\<Q>", "\\<R>", "\\<S>", "\\<T>", "\\<U>", | |
| 361 | "\\<V>", "\\<W>", "\\<X>", "\\<Y>", "\\<Z>", "\\<a>", "\\<b>", | |
| 362 | "\\<c>", "\\<d>", "\\<e>", "\\<f>", "\\<g>", "\\<h>", "\\<i>", | |
| 363 | "\\<j>", "\\<k>", "\\<l>", "\\<m>", "\\<n>", "\\<o>", "\\<p>", | |
| 364 | "\\<q>", "\\<r>", "\\<s>", "\\<t>", "\\<u>", "\\<v>", "\\<w>", | |
| 365 | "\\<x>", "\\<y>", "\\<z>", | |
| 366 | ||
| 367 | "\\<AA>", "\\<BB>", "\\<CC>", "\\<DD>", "\\<EE>", "\\<FF>", | |
| 368 | "\\<GG>", "\\<HH>", "\\<II>", "\\<JJ>", "\\<KK>", "\\<LL>", | |
| 369 | "\\<MM>", "\\<NN>", "\\<OO>", "\\<PP>", "\\<QQ>", "\\<RR>", | |
| 370 | "\\<SS>", "\\<TT>", "\\<UU>", "\\<VV>", "\\<WW>", "\\<XX>", | |
| 371 | "\\<YY>", "\\<ZZ>", "\\<aa>", "\\<bb>", "\\<cc>", "\\<dd>", | |
| 372 | "\\<ee>", "\\<ff>", "\\<gg>", "\\<hh>", "\\<ii>", "\\<jj>", | |
| 373 | "\\<kk>", "\\<ll>", "\\<mm>", "\\<nn>", "\\<oo>", "\\<pp>", | |
| 374 | "\\<qq>", "\\<rr>", "\\<ss>", "\\<tt>", "\\<uu>", "\\<vv>", | |
| 375 | "\\<ww>", "\\<xx>", "\\<yy>", "\\<zz>", | |
| 376 | ||
| 377 | "\\<alpha>", "\\<beta>", "\\<gamma>", "\\<delta>", "\\<epsilon>", | |
| 378 | "\\<zeta>", "\\<eta>", "\\<theta>", "\\<iota>", "\\<kappa>", | |
| 379 | "\\<mu>", "\\<nu>", "\\<xi>", "\\<pi>", "\\<rho>", "\\<sigma>", | |
| 380 | "\\<tau>", "\\<upsilon>", "\\<phi>", "\\<chi>", "\\<psi>", | |
| 381 | "\\<omega>", "\\<Gamma>", "\\<Delta>", "\\<Theta>", "\\<Lambda>", | |
| 382 | "\\<Xi>", "\\<Pi>", "\\<Sigma>", "\\<Upsilon>", "\\<Phi>", | |
| 52616 
3ac2878764f9
more robust identifier syntax: sub/superscript counts as modifier of LETDIG part instead of LETTER, both isub/isup and sub/sup are allowed;
 wenzelm parents: 
52507diff
changeset | 383 | "\\<Psi>", "\\<Omega>") | 
| 34134 | 384 | |
| 54734 
b91afc3aa3e6
clarified Proof General legacy: special treatment of \<^newline> only in TTY mode;
 wenzelm parents: 
53400diff
changeset | 385 |     val blanks = recode_set(" ", "\t", "\n", "\u000B", "\f", "\r", "\r\n")
 | 
| 34138 | 386 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 387 | val sym_chars = | 
| 34138 | 388 |       Set("!", "#", "$", "%", "&", "*", "+", "-", "/", "<", "=", ">", "?", "@", "^", "_", "|", "~")
 | 
| 34134 | 389 | |
| 44992 
aa34d2d049ce
refined Symbol.is_symbolic -- cover recoded versions as well;
 wenzelm parents: 
44949diff
changeset | 390 |     val symbolic = recode_set((for { (sym, _) <- symbols; if raw_symbolic(sym) } yield sym): _*)
 | 
| 
aa34d2d049ce
refined Symbol.is_symbolic -- cover recoded versions as well;
 wenzelm parents: 
44949diff
changeset | 391 | |
| 43455 | 392 | |
| 55033 | 393 | /* cartouches */ | 
| 394 | ||
| 395 | val open_decoded = decode(open) | |
| 396 | val close_decoded = decode(close) | |
| 397 | ||
| 398 | ||
| 43488 | 399 | /* control symbols */ | 
| 400 | ||
| 43696 | 401 | val ctrl_decoded: Set[Symbol] = | 
| 43488 | 402 |       Set((for ((sym, _) <- symbols if sym.startsWith("\\<^")) yield decode(sym)): _*)
 | 
| 403 | ||
| 44238 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 404 |     val sub_decoded = decode("\\<^sub>")
 | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 405 |     val sup_decoded = decode("\\<^sup>")
 | 
| 43511 | 406 |     val bsub_decoded = decode("\\<^bsub>")
 | 
| 407 |     val esub_decoded = decode("\\<^esub>")
 | |
| 408 |     val bsup_decoded = decode("\\<^bsup>")
 | |
| 409 |     val esup_decoded = decode("\\<^esup>")
 | |
| 44238 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 410 |     val bold_decoded = decode("\\<^bold>")
 | 
| 27918 | 411 | } | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 412 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 413 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 414 | /* tables */ | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 415 | |
| 53400 | 416 | def properties: Map[Symbol, Properties.T] = symbols.properties | 
| 43696 | 417 | def names: Map[Symbol, String] = symbols.names | 
| 50136 
a96bd08258a2
support for symbol groups, retaining original order of declarations;
 wenzelm parents: 
48922diff
changeset | 418 | def groups: List[(String, List[Symbol])] = symbols.groups | 
| 53316 
c3e549e0d3c7
allow multiple symbol properties, notably groups and abbrevs;
 wenzelm parents: 
53021diff
changeset | 419 | def abbrevs: Multi_Map[Symbol, String] = symbols.abbrevs | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 420 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 421 | def decode(text: String): String = symbols.decode(text) | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 422 | def encode(text: String): String = symbols.encode(text) | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 423 | |
| 53337 
b3817a0e3211
sort items according to persistent history of frequency of use;
 wenzelm parents: 
53316diff
changeset | 424 | def decode_string: XML.Decode.T[String] = (x => decode(XML.Decode.string(x))) | 
| 
b3817a0e3211
sort items according to persistent history of frequency of use;
 wenzelm parents: 
53316diff
changeset | 425 | def encode_string: XML.Encode.T[String] = (x => XML.Encode.string(encode(x))) | 
| 
b3817a0e3211
sort items according to persistent history of frequency of use;
 wenzelm parents: 
53316diff
changeset | 426 | |
| 50291 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 427 | def decode_strict(text: String): String = | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 428 |   {
 | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 429 | val decoded = decode(text) | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 430 | if (encode(decoded) == text) decoded | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 431 |     else {
 | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 432 | val bad = new mutable.ListBuffer[Symbol] | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 433 | for (s <- iterator(text) if encode(decode(s)) != s && !bad.contains(s)) | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 434 | bad += s | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 435 |       error("Bad Unicode symbols in text: " + commas_quote(bad))
 | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 436 | } | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 437 | } | 
| 
674893679352
prefer Symbol.decode_strict in batch mode, to avoid files with spurious Unicode symbols that clash with Isabelle symbol interpretation;
 wenzelm parents: 
50238diff
changeset | 438 | |
| 43696 | 439 | def fonts: Map[Symbol, String] = symbols.fonts | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 440 | def font_names: List[String] = symbols.font_names | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 441 | def font_index: Map[String, Int] = symbols.font_index | 
| 43696 | 442 | def lookup_font(sym: Symbol): Option[Int] = symbols.fonts.get(sym).map(font_index(_)) | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 443 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 444 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 445 | /* classification */ | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 446 | |
| 43696 | 447 | def is_letter(sym: Symbol): Boolean = symbols.letters.contains(sym) | 
| 448 | def is_digit(sym: Symbol): Boolean = sym.length == 1 && '0' <= sym(0) && sym(0) <= '9' | |
| 449 | def is_quasi(sym: Symbol): Boolean = sym == "_" || sym == "'" | |
| 450 | def is_letdig(sym: Symbol): Boolean = is_letter(sym) || is_digit(sym) || is_quasi(sym) | |
| 451 | def is_blank(sym: Symbol): Boolean = symbols.blanks.contains(sym) | |
| 44992 
aa34d2d049ce
refined Symbol.is_symbolic -- cover recoded versions as well;
 wenzelm parents: 
44949diff
changeset | 452 | |
| 55033 | 453 | |
| 454 | /* cartouches */ | |
| 455 | ||
| 456 | val open = "\\<open>" | |
| 457 | val close = "\\<close>" | |
| 458 | ||
| 459 | def open_decoded: Symbol = symbols.open_decoded | |
| 460 | def close_decoded: Symbol = symbols.close_decoded | |
| 461 | ||
| 462 | def is_open(sym: Symbol): Boolean = sym == open_decoded || sym == open | |
| 463 | def is_close(sym: Symbol): Boolean = sym == close_decoded || sym == close | |
| 464 | ||
| 465 | ||
| 466 | /* symbols for symbolic identifiers */ | |
| 44992 
aa34d2d049ce
refined Symbol.is_symbolic -- cover recoded versions as well;
 wenzelm parents: 
44949diff
changeset | 467 | |
| 
aa34d2d049ce
refined Symbol.is_symbolic -- cover recoded versions as well;
 wenzelm parents: 
44949diff
changeset | 468 | private def raw_symbolic(sym: Symbol): Boolean = | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 469 |     sym.startsWith("\\<") && sym.endsWith(">") && !sym.startsWith("\\<^")
 | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 470 | |
| 55033 | 471 | def is_symbolic(sym: Symbol): Boolean = | 
| 472 | !is_open(sym) && !is_close(sym) && (raw_symbolic(sym) || symbols.symbolic.contains(sym)) | |
| 473 | ||
| 474 | def is_symbolic_char(sym: Symbol): Boolean = symbols.sym_chars.contains(sym) | |
| 475 | ||
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 476 | |
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 477 | /* control symbols */ | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 478 | |
| 43696 | 479 | def is_ctrl(sym: Symbol): Boolean = | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 480 |     sym.startsWith("\\<^") || symbols.ctrl_decoded.contains(sym)
 | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 481 | |
| 43696 | 482 | def is_controllable(sym: Symbol): Boolean = | 
| 55033 | 483 | !is_blank(sym) && !is_ctrl(sym) && !is_open(sym) && !is_close(sym) && !is_malformed(sym) | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43675diff
changeset | 484 | |
| 44238 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 485 | def sub_decoded: Symbol = symbols.sub_decoded | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 486 | def sup_decoded: Symbol = symbols.sup_decoded | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 487 | def bsub_decoded: Symbol = symbols.bsub_decoded | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 488 | def esub_decoded: Symbol = symbols.esub_decoded | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 489 | def bsup_decoded: Symbol = symbols.bsup_decoded | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 490 | def esup_decoded: Symbol = symbols.esup_decoded | 
| 
36120feb70ed
some convenience actions/shortcuts for control symbols;
 wenzelm parents: 
44181diff
changeset | 491 | def bold_decoded: Symbol = symbols.bold_decoded | 
| 27901 | 492 | } |