| author | wenzelm | 
| Tue, 22 Sep 2015 18:06:49 +0200 | |
| changeset 61251 | 2da25a27a616 | 
| parent 60215 | 5fb4990dfc73 | 
| child 61883 | c0f34fe6aa61 | 
| permissions | -rw-r--r-- | 
| 34136 | 1 | /* Title: Pure/library.scala | 
| 45673 
cd41e3903fbf
separate compilation of PIDE vs. Pure sources, which enables independent Scala library;
 wenzelm parents: 
45667diff
changeset | 2 | Module: PIDE | 
| 34136 | 3 | Author: Makarius | 
| 4 | ||
| 5 | Basic library. | |
| 6 | */ | |
| 7 | ||
| 8 | package isabelle | |
| 9 | ||
| 38258 | 10 | |
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 11 | import scala.collection.mutable | 
| 59224 | 12 | import scala.util.matching.Regex | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 13 | |
| 37018 | 14 | |
| 34136 | 15 | object Library | 
| 16 | {
 | |
| 43652 | 17 | /* user errors */ | 
| 18 | ||
| 59697 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 19 | class User_Error(message: String) extends RuntimeException(message) | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 20 |   {
 | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 21 | override def equals(that: Any): Boolean = | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 22 |       that match {
 | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 23 | case other: User_Error => message == other.getMessage | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 24 | case _ => false | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 25 | } | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 26 | override def hashCode: Int = message.hashCode | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 27 | |
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 28 |     override def toString: String = "ERROR(" + message + ")"
 | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 29 | } | 
| 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 30 | |
| 43652 | 31 | object ERROR | 
| 32 |   {
 | |
| 59697 
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
 wenzelm parents: 
59224diff
changeset | 33 | def apply(message: String): User_Error = new User_Error(message) | 
| 48479 
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
 wenzelm parents: 
48425diff
changeset | 34 | def unapply(exn: Throwable): Option[String] = Exn.user_message(exn) | 
| 43652 | 35 | } | 
| 36 | ||
| 37 | def error(message: String): Nothing = throw ERROR(message) | |
| 38 | ||
| 54548 | 39 | def cat_message(msg1: String, msg2: String): String = | 
| 40 | if (msg1 == "") msg2 | |
| 57831 
885888a880fb
more general notion of "user error" including empty message -- NB: Output.error_message needs non-empty string to emit anything;
 wenzelm parents: 
56843diff
changeset | 41 | else if (msg2 == "") msg1 | 
| 54548 | 42 | else msg1 + "\n" + msg2 | 
| 43 | ||
| 43652 | 44 | def cat_error(msg1: String, msg2: String): Nothing = | 
| 54548 | 45 | error(cat_message(msg1, msg2)) | 
| 43652 | 46 | |
| 47 | ||
| 57909 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 48 | /* integers */ | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 49 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 50 | private val small_int = 10000 | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 51 | private lazy val small_int_table = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 52 |   {
 | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 53 | val array = new Array[String](small_int) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 54 | for (i <- 0 until small_int) array(i) = i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 55 | array | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 56 | } | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 57 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 58 | def is_small_int(s: String): Boolean = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 59 |   {
 | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 60 | val len = s.length | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 61 | 1 <= len && len <= 4 && | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 62 | s.forall(c => '0' <= c && c <= '9') && | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 63 | (len == 1 || s(0) != '0') | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 64 | } | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 65 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 66 | def signed_string_of_long(i: Long): String = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 67 | if (0 <= i && i < small_int) small_int_table(i.toInt) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 68 | else i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 69 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 70 | def signed_string_of_int(i: Int): String = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 71 | if (0 <= i && i < small_int) small_int_table(i) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 72 | else i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 73 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 74 | |
| 48996 | 75 | /* separated chunks */ | 
| 36688 | 76 | |
| 77 | def separate[A](s: A, list: List[A]): List[A] = | |
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 78 |   {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 79 | val result = new mutable.ListBuffer[A] | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 80 | var first = true | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 81 |     for (x <- list) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 82 |       if (first) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 83 | first = false | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 84 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 85 | } | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 86 |       else {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 87 | result += s | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 88 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 89 | } | 
| 36688 | 90 | } | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 91 | result.toList | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 92 | } | 
| 36688 | 93 | |
| 56600 | 94 | def separated_chunks(sep: Char => Boolean, source: CharSequence): Iterator[CharSequence] = | 
| 48996 | 95 |     new Iterator[CharSequence] {
 | 
| 96 | private val end = source.length | |
| 97 | private def next_chunk(i: Int): Option[(CharSequence, Int)] = | |
| 98 |       {
 | |
| 99 |         if (i < end) {
 | |
| 56600 | 100 | var j = i; do j += 1 while (j < end && !sep(source.charAt(j))) | 
| 48996 | 101 | Some((source.subSequence(i + 1, j), j)) | 
| 102 | } | |
| 103 | else None | |
| 43598 | 104 | } | 
| 48996 | 105 | private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1) | 
| 106 | ||
| 107 | def hasNext(): Boolean = state.isDefined | |
| 108 | def next(): CharSequence = | |
| 109 |         state match {
 | |
| 60215 | 110 | case Some((s, i)) => state = next_chunk(i); s | 
| 48996 | 111 | case None => Iterator.empty.next() | 
| 112 | } | |
| 43598 | 113 | } | 
| 114 | ||
| 48996 | 115 | def space_explode(sep: Char, str: String): List[String] = | 
| 56600 | 116 | separated_chunks(_ == sep, str).map(_.toString).toList | 
| 48996 | 117 | |
| 118 | ||
| 119 | /* lines */ | |
| 120 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 121 | def terminate_lines(lines: Iterable[CharSequence]): Iterable[CharSequence] = | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 122 |     new Iterable[CharSequence] {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 123 | def iterator: Iterator[CharSequence] = | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 124 | lines.iterator.map(line => new Line_Termination(line)) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 125 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 126 | |
| 48996 | 127 |   def cat_lines(lines: TraversableOnce[String]): String = lines.mkString("\n")
 | 
| 128 | ||
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 129 |   def split_lines(str: String): List[String] = space_explode('\n', str)
 | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 130 | |
| 48996 | 131 | def first_line(source: CharSequence): String = | 
| 132 |   {
 | |
| 56600 | 133 | val lines = separated_chunks(_ == '\n', source) | 
| 48996 | 134 | if (lines.hasNext) lines.next.toString | 
| 135 | else "" | |
| 136 | } | |
| 137 | ||
| 50847 | 138 | |
| 139 | /* strings */ | |
| 140 | ||
| 141 | def try_unprefix(prfx: String, s: String): Option[String] = | |
| 142 | if (s.startsWith(prfx)) Some(s.substring(prfx.length)) else None | |
| 143 | ||
| 55033 | 144 | def try_unsuffix(sffx: String, s: String): Option[String] = | 
| 145 | if (s.endsWith(sffx)) Some(s.substring(0, s.length - sffx.length)) else None | |
| 146 | ||
| 52444 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 147 | def trim_line(s: String): String = | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 148 |     if (s.endsWith("\r\n")) s.substring(0, s.length - 2)
 | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 149 |     else if (s.endsWith("\r") || s.endsWith("\n")) s.substring(0, s.length - 1)
 | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 150 | else s | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 151 | |
| 43598 | 152 | |
| 48996 | 153 | /* quote */ | 
| 46196 | 154 | |
| 43598 | 155 | def quote(s: String): String = "\"" + s + "\"" | 
| 56843 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 156 | |
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 157 | def try_unquote(s: String): Option[String] = | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 158 |     if (s.startsWith("\"") && s.endsWith("\"")) Some(s.substring(1, s.length - 1))
 | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 159 | else None | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 160 | |
| 58592 | 161 | def perhaps_unquote(s: String): String = try_unquote(s) getOrElse s | 
| 162 | ||
| 43598 | 163 |   def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | 
| 48362 | 164 |   def commas_quote(ss: Iterable[String]): String = ss.iterator.map(quote).mkString(", ")
 | 
| 43598 | 165 | |
| 36688 | 166 | |
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 167 | /* CharSequence */ | 
| 34141 | 168 | |
| 169 | class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence | |
| 170 |   {
 | |
| 171 | require(0 <= start && start <= end && end <= text.length) | |
| 172 | ||
| 173 | def this(text: CharSequence) = this(text, 0, text.length) | |
| 174 | ||
| 175 | def length: Int = end - start | |
| 176 | def charAt(i: Int): Char = text.charAt(end - i - 1) | |
| 177 | ||
| 178 | def subSequence(i: Int, j: Int): CharSequence = | |
| 179 | if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i) | |
| 180 | else throw new IndexOutOfBoundsException | |
| 181 | ||
| 182 | override def toString: String = | |
| 183 |     {
 | |
| 184 | val buf = new StringBuilder(length) | |
| 185 | for (i <- 0 until length) | |
| 186 | buf.append(charAt(i)) | |
| 187 | buf.toString | |
| 188 | } | |
| 189 | } | |
| 190 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 191 | class Line_Termination(text: CharSequence) extends CharSequence | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 192 |   {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 193 | def length: Int = text.length + 1 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 194 | def charAt(i: Int): Char = if (i == text.length) '\n' else text.charAt(i) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 195 | def subSequence(i: Int, j: Int): CharSequence = | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 196 | if (j == text.length + 1) new Line_Termination(text.subSequence(i, j - 1)) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 197 | else text.subSequence(i, j) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 198 | override def toString: String = text.toString + "\n" | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 199 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 200 | |
| 34141 | 201 | |
| 59224 | 202 | /* regular expressions */ | 
| 203 | ||
| 204 | def make_regex(s: String): Option[Regex] = | |
| 205 |     try { Some(new Regex(s)) } catch { case ERROR(_) => None }
 | |
| 206 | ||
| 207 | ||
| 56686 | 208 | /* canonical list operations */ | 
| 209 | ||
| 60215 | 210 | def member[A, B](xs: List[A])(x: B): Boolean = xs.contains(x) | 
| 56688 | 211 | def insert[A](x: A)(xs: List[A]): List[A] = if (xs.contains(x)) xs else x :: xs | 
| 212 | def remove[A, B](x: B)(xs: List[A]): List[A] = if (member(xs)(x)) xs.filterNot(_ == x) else xs | |
| 213 | def update[A](x: A)(xs: List[A]): List[A] = x :: remove(x)(xs) | |
| 34136 | 214 | } | 
| 43652 | 215 | |
| 216 | ||
| 217 | class Basic_Library | |
| 218 | {
 | |
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 219 | val ERROR = Library.ERROR | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 220 | val error = Library.error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 221 | val cat_error = Library.cat_error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 222 | |
| 43652 | 223 | val space_explode = Library.space_explode _ | 
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 224 | val split_lines = Library.split_lines _ | 
| 46196 | 225 | val cat_lines = Library.cat_lines _ | 
| 43652 | 226 | val quote = Library.quote _ | 
| 227 | val commas = Library.commas _ | |
| 228 | val commas_quote = Library.commas_quote _ | |
| 229 | } |