| author | panny | 
| Tue, 03 Dec 2013 02:51:20 +0100 | |
| changeset 54629 | a692901ecdc2 | 
| parent 54548 | 41e4ba92a979 | 
| child 55033 | 8e8243975860 | 
| 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 | |
| 49245 
cb70157293c0
manage Isabelle/jEdit options as Isabelle/Scala options (with persistent preferences);
 wenzelm parents: 
48996diff
changeset | 13 | import java.util.Locale | 
| 50414 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 14 | import java.util.concurrent.{Future => JFuture, TimeUnit}
 | 
| 37018 | 15 | |
| 16 | ||
| 34136 | 17 | object Library | 
| 18 | {
 | |
| 43652 | 19 | /* user errors */ | 
| 20 | ||
| 21 | object ERROR | |
| 22 |   {
 | |
| 23 | def apply(message: String): Throwable = new RuntimeException(message) | |
| 48479 
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
 wenzelm parents: 
48425diff
changeset | 24 | def unapply(exn: Throwable): Option[String] = Exn.user_message(exn) | 
| 43652 | 25 | } | 
| 26 | ||
| 27 | def error(message: String): Nothing = throw ERROR(message) | |
| 28 | ||
| 54548 | 29 | def cat_message(msg1: String, msg2: String): String = | 
| 30 | if (msg1 == "") msg2 | |
| 31 | else msg1 + "\n" + msg2 | |
| 32 | ||
| 43652 | 33 | def cat_error(msg1: String, msg2: String): Nothing = | 
| 54548 | 34 | error(cat_message(msg1, msg2)) | 
| 43652 | 35 | |
| 36 | ||
| 48996 | 37 | /* separated chunks */ | 
| 36688 | 38 | |
| 39 | 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 | 40 |   {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 41 | val result = new mutable.ListBuffer[A] | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 42 | var first = true | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 43 |     for (x <- list) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 44 |       if (first) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 45 | first = false | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 46 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 47 | } | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 48 |       else {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 49 | result += s | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 50 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 51 | } | 
| 36688 | 52 | } | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 53 | result.toList | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 54 | } | 
| 36688 | 55 | |
| 48996 | 56 | def separated_chunks(sep: Char, source: CharSequence): Iterator[CharSequence] = | 
| 57 |     new Iterator[CharSequence] {
 | |
| 58 | private val end = source.length | |
| 59 | private def next_chunk(i: Int): Option[(CharSequence, Int)] = | |
| 60 |       {
 | |
| 61 |         if (i < end) {
 | |
| 62 | var j = i; do j += 1 while (j < end && source.charAt(j) != sep) | |
| 63 | Some((source.subSequence(i + 1, j), j)) | |
| 64 | } | |
| 65 | else None | |
| 43598 | 66 | } | 
| 48996 | 67 | private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1) | 
| 68 | ||
| 69 | def hasNext(): Boolean = state.isDefined | |
| 70 | def next(): CharSequence = | |
| 71 |         state match {
 | |
| 72 |           case Some((s, i)) => { state = next_chunk(i); s }
 | |
| 73 | case None => Iterator.empty.next() | |
| 74 | } | |
| 43598 | 75 | } | 
| 76 | ||
| 48996 | 77 | def space_explode(sep: Char, str: String): List[String] = | 
| 78 | separated_chunks(sep, str).map(_.toString).toList | |
| 79 | ||
| 80 | ||
| 81 | /* lines */ | |
| 82 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 83 | 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 | 84 |     new Iterable[CharSequence] {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 85 | def iterator: Iterator[CharSequence] = | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 86 | 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 | 87 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 88 | |
| 48996 | 89 |   def cat_lines(lines: TraversableOnce[String]): String = lines.mkString("\n")
 | 
| 90 | ||
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 91 |   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 | 92 | |
| 48996 | 93 | def first_line(source: CharSequence): String = | 
| 94 |   {
 | |
| 95 |     val lines = separated_chunks('\n', source)
 | |
| 96 | if (lines.hasNext) lines.next.toString | |
| 97 | else "" | |
| 98 | } | |
| 99 | ||
| 50847 | 100 | |
| 101 | /* strings */ | |
| 102 | ||
| 50299 | 103 | def lowercase(str: String): String = str.toLowerCase(Locale.ENGLISH) | 
| 104 | def uppercase(str: String): String = str.toUpperCase(Locale.ENGLISH) | |
| 105 | ||
| 49245 
cb70157293c0
manage Isabelle/jEdit options as Isabelle/Scala options (with persistent preferences);
 wenzelm parents: 
48996diff
changeset | 106 | def capitalize(str: String): String = | 
| 
cb70157293c0
manage Isabelle/jEdit options as Isabelle/Scala options (with persistent preferences);
 wenzelm parents: 
48996diff
changeset | 107 | if (str.length == 0) str | 
| 50299 | 108 | else uppercase(str.substring(0, 1)) + str.substring(1) | 
| 49245 
cb70157293c0
manage Isabelle/jEdit options as Isabelle/Scala options (with persistent preferences);
 wenzelm parents: 
48996diff
changeset | 109 | |
| 50847 | 110 | def try_unprefix(prfx: String, s: String): Option[String] = | 
| 111 | if (s.startsWith(prfx)) Some(s.substring(prfx.length)) else None | |
| 112 | ||
| 52444 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 113 | def trim_line(s: String): String = | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 114 |     if (s.endsWith("\r\n")) s.substring(0, s.length - 2)
 | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 115 |     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 | 116 | else s | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 117 | |
| 43598 | 118 | |
| 48996 | 119 | /* quote */ | 
| 46196 | 120 | |
| 43598 | 121 | def quote(s: String): String = "\"" + s + "\"" | 
| 122 |   def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | |
| 48362 | 123 |   def commas_quote(ss: Iterable[String]): String = ss.iterator.map(quote).mkString(", ")
 | 
| 43598 | 124 | |
| 36688 | 125 | |
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 126 | /* CharSequence */ | 
| 34141 | 127 | |
| 128 | class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence | |
| 129 |   {
 | |
| 130 | require(0 <= start && start <= end && end <= text.length) | |
| 131 | ||
| 132 | def this(text: CharSequence) = this(text, 0, text.length) | |
| 133 | ||
| 134 | def length: Int = end - start | |
| 135 | def charAt(i: Int): Char = text.charAt(end - i - 1) | |
| 136 | ||
| 137 | def subSequence(i: Int, j: Int): CharSequence = | |
| 138 | if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i) | |
| 139 | else throw new IndexOutOfBoundsException | |
| 140 | ||
| 141 | override def toString: String = | |
| 142 |     {
 | |
| 143 | val buf = new StringBuilder(length) | |
| 144 | for (i <- 0 until length) | |
| 145 | buf.append(charAt(i)) | |
| 146 | buf.toString | |
| 147 | } | |
| 148 | } | |
| 149 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 150 | 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 | 151 |   {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 152 | def length: Int = text.length + 1 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 153 | 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 | 154 | 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 | 155 | 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 | 156 | else text.subSequence(i, j) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 157 | 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 | 158 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 159 | |
| 34141 | 160 | |
| 50414 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 161 | /* Java futures */ | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 162 | |
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 163 | def future_value[A](x: A) = new JFuture[A] | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 164 |   {
 | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 165 | def cancel(may_interrupt: Boolean): Boolean = false | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 166 | def isCancelled(): Boolean = false | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 167 | def isDone(): Boolean = true | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 168 | def get(): A = x | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 169 | def get(timeout: Long, time_unit: TimeUnit): A = x | 
| 
e17a1f179bb0
explore theory_body_files via future, for improved performance;
 wenzelm parents: 
50299diff
changeset | 170 | } | 
| 34136 | 171 | } | 
| 43652 | 172 | |
| 173 | ||
| 174 | class Basic_Library | |
| 175 | {
 | |
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 176 | val ERROR = Library.ERROR | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 177 | val error = Library.error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 178 | val cat_error = Library.cat_error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 179 | |
| 43652 | 180 | val space_explode = Library.space_explode _ | 
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 181 | val split_lines = Library.split_lines _ | 
| 46196 | 182 | val cat_lines = Library.cat_lines _ | 
| 43652 | 183 | val quote = Library.quote _ | 
| 184 | val commas = Library.commas _ | |
| 185 | val commas_quote = Library.commas_quote _ | |
| 49470 | 186 | |
| 187 | ||
| 188 | /* parallel tasks */ | |
| 189 | ||
| 190 | implicit def function_as_callable[A](f: () => A) = | |
| 191 |     new java.util.concurrent.Callable[A] { def call = f() }
 | |
| 192 | ||
| 193 | val default_thread_pool = | |
| 194 | scala.collection.parallel.ThreadPoolTasks.defaultThreadPool | |
| 43652 | 195 | } |