| author | blanchet | 
| Thu, 26 Jun 2014 16:41:43 +0200 | |
| changeset 57383 | ba0fe0639bc8 | 
| parent 56843 | b2bfcd8cda80 | 
| child 57831 | 885888a880fb | 
| 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 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 12 | |
| 37018 | 13 | |
| 34136 | 14 | object Library | 
| 15 | {
 | |
| 43652 | 16 | /* user errors */ | 
| 17 | ||
| 18 | object ERROR | |
| 19 |   {
 | |
| 20 | def apply(message: String): Throwable = new RuntimeException(message) | |
| 48479 
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
 wenzelm parents: 
48425diff
changeset | 21 | def unapply(exn: Throwable): Option[String] = Exn.user_message(exn) | 
| 43652 | 22 | } | 
| 23 | ||
| 24 | def error(message: String): Nothing = throw ERROR(message) | |
| 25 | ||
| 54548 | 26 | def cat_message(msg1: String, msg2: String): String = | 
| 27 | if (msg1 == "") msg2 | |
| 28 | else msg1 + "\n" + msg2 | |
| 29 | ||
| 43652 | 30 | def cat_error(msg1: String, msg2: String): Nothing = | 
| 54548 | 31 | error(cat_message(msg1, msg2)) | 
| 43652 | 32 | |
| 33 | ||
| 48996 | 34 | /* separated chunks */ | 
| 36688 | 35 | |
| 36 | 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 | 37 |   {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 38 | val result = new mutable.ListBuffer[A] | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 39 | var first = true | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 40 |     for (x <- list) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 41 |       if (first) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 42 | first = false | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 43 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 44 | } | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 45 |       else {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 46 | result += s | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 47 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 48 | } | 
| 36688 | 49 | } | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 50 | result.toList | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 51 | } | 
| 36688 | 52 | |
| 56600 | 53 | def separated_chunks(sep: Char => Boolean, source: CharSequence): Iterator[CharSequence] = | 
| 48996 | 54 |     new Iterator[CharSequence] {
 | 
| 55 | private val end = source.length | |
| 56 | private def next_chunk(i: Int): Option[(CharSequence, Int)] = | |
| 57 |       {
 | |
| 58 |         if (i < end) {
 | |
| 56600 | 59 | var j = i; do j += 1 while (j < end && !sep(source.charAt(j))) | 
| 48996 | 60 | Some((source.subSequence(i + 1, j), j)) | 
| 61 | } | |
| 62 | else None | |
| 43598 | 63 | } | 
| 48996 | 64 | private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1) | 
| 65 | ||
| 66 | def hasNext(): Boolean = state.isDefined | |
| 67 | def next(): CharSequence = | |
| 68 |         state match {
 | |
| 69 |           case Some((s, i)) => { state = next_chunk(i); s }
 | |
| 70 | case None => Iterator.empty.next() | |
| 71 | } | |
| 43598 | 72 | } | 
| 73 | ||
| 48996 | 74 | def space_explode(sep: Char, str: String): List[String] = | 
| 56600 | 75 | separated_chunks(_ == sep, str).map(_.toString).toList | 
| 48996 | 76 | |
| 77 | ||
| 78 | /* lines */ | |
| 79 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 80 | 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 | 81 |     new Iterable[CharSequence] {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 82 | def iterator: Iterator[CharSequence] = | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 83 | 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 | 84 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 85 | |
| 48996 | 86 |   def cat_lines(lines: TraversableOnce[String]): String = lines.mkString("\n")
 | 
| 87 | ||
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 88 |   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 | 89 | |
| 48996 | 90 | def first_line(source: CharSequence): String = | 
| 91 |   {
 | |
| 56600 | 92 | val lines = separated_chunks(_ == '\n', source) | 
| 48996 | 93 | if (lines.hasNext) lines.next.toString | 
| 94 | else "" | |
| 95 | } | |
| 96 | ||
| 50847 | 97 | |
| 98 | /* strings */ | |
| 99 | ||
| 100 | def try_unprefix(prfx: String, s: String): Option[String] = | |
| 101 | if (s.startsWith(prfx)) Some(s.substring(prfx.length)) else None | |
| 102 | ||
| 55033 | 103 | def try_unsuffix(sffx: String, s: String): Option[String] = | 
| 104 | if (s.endsWith(sffx)) Some(s.substring(0, s.length - sffx.length)) else None | |
| 105 | ||
| 52444 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 106 | def trim_line(s: String): String = | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 107 |     if (s.endsWith("\r\n")) s.substring(0, s.length - 2)
 | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 108 |     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 | 109 | else s | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 110 | |
| 43598 | 111 | |
| 48996 | 112 | /* quote */ | 
| 46196 | 113 | |
| 43598 | 114 | def quote(s: String): String = "\"" + s + "\"" | 
| 56843 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 115 | |
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 116 | def try_unquote(s: String): Option[String] = | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 117 |     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 | 118 | else None | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 119 | |
| 43598 | 120 |   def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | 
| 48362 | 121 |   def commas_quote(ss: Iterable[String]): String = ss.iterator.map(quote).mkString(", ")
 | 
| 43598 | 122 | |
| 36688 | 123 | |
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 124 | /* CharSequence */ | 
| 34141 | 125 | |
| 126 | class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence | |
| 127 |   {
 | |
| 128 | require(0 <= start && start <= end && end <= text.length) | |
| 129 | ||
| 130 | def this(text: CharSequence) = this(text, 0, text.length) | |
| 131 | ||
| 132 | def length: Int = end - start | |
| 133 | def charAt(i: Int): Char = text.charAt(end - i - 1) | |
| 134 | ||
| 135 | def subSequence(i: Int, j: Int): CharSequence = | |
| 136 | if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i) | |
| 137 | else throw new IndexOutOfBoundsException | |
| 138 | ||
| 139 | override def toString: String = | |
| 140 |     {
 | |
| 141 | val buf = new StringBuilder(length) | |
| 142 | for (i <- 0 until length) | |
| 143 | buf.append(charAt(i)) | |
| 144 | buf.toString | |
| 145 | } | |
| 146 | } | |
| 147 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 148 | 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 | 149 |   {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 150 | def length: Int = text.length + 1 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 151 | 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 | 152 | 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 | 153 | 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 | 154 | else text.subSequence(i, j) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 155 | 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 | 156 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 157 | |
| 34141 | 158 | |
| 56686 | 159 | /* canonical list operations */ | 
| 160 | ||
| 56688 | 161 | def member[A, B](xs: List[A])(x: B): Boolean = xs.exists(_ == x) | 
| 162 | def insert[A](x: A)(xs: List[A]): List[A] = if (xs.contains(x)) xs else x :: xs | |
| 163 | def remove[A, B](x: B)(xs: List[A]): List[A] = if (member(xs)(x)) xs.filterNot(_ == x) else xs | |
| 164 | def update[A](x: A)(xs: List[A]): List[A] = x :: remove(x)(xs) | |
| 34136 | 165 | } | 
| 43652 | 166 | |
| 167 | ||
| 168 | class Basic_Library | |
| 169 | {
 | |
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 170 | val ERROR = Library.ERROR | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 171 | val error = Library.error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 172 | val cat_error = Library.cat_error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 173 | |
| 43652 | 174 | val space_explode = Library.space_explode _ | 
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 175 | val split_lines = Library.split_lines _ | 
| 46196 | 176 | val cat_lines = Library.cat_lines _ | 
| 43652 | 177 | val quote = Library.quote _ | 
| 178 | val commas = Library.commas _ | |
| 179 | val commas_quote = Library.commas_quote _ | |
| 180 | } |