| author | wenzelm | 
| Tue, 12 Aug 2014 13:18:17 +0200 | |
| changeset 57909 | 0fb331032f02 | 
| parent 57831 | 885888a880fb | 
| child 58592 | b0fff34d3247 | 
| 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: 
45667 
diff
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: 
51616 
diff
changeset
 | 
11  | 
import scala.collection.mutable  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
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: 
48425 
diff
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  | 
|
| 
57831
 
885888a880fb
more general notion of "user error" including empty message -- NB: Output.error_message needs non-empty string to emit anything;
 
wenzelm 
parents: 
56843 
diff
changeset
 | 
28  | 
else if (msg2 == "") msg1  | 
| 54548 | 29  | 
else msg1 + "\n" + msg2  | 
30  | 
||
| 43652 | 31  | 
def cat_error(msg1: String, msg2: String): Nothing =  | 
| 54548 | 32  | 
error(cat_message(msg1, msg2))  | 
| 43652 | 33  | 
|
34  | 
||
| 
57909
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
35  | 
/* integers */  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
36  | 
|
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
37  | 
private val small_int = 10000  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
38  | 
private lazy val small_int_table =  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
39  | 
  {
 | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
40  | 
val array = new Array[String](small_int)  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
41  | 
for (i <- 0 until small_int) array(i) = i.toString  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
42  | 
array  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
43  | 
}  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
44  | 
|
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
45  | 
def is_small_int(s: String): Boolean =  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
46  | 
  {
 | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
47  | 
val len = s.length  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
48  | 
1 <= len && len <= 4 &&  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
49  | 
s.forall(c => '0' <= c && c <= '9') &&  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
50  | 
(len == 1 || s(0) != '0')  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
51  | 
}  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
52  | 
|
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
53  | 
def signed_string_of_long(i: Long): String =  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
54  | 
if (0 <= i && i < small_int) small_int_table(i.toInt)  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
55  | 
else i.toString  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
56  | 
|
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
57  | 
def signed_string_of_int(i: Int): String =  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
58  | 
if (0 <= i && i < small_int) small_int_table(i)  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
59  | 
else i.toString  | 
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
60  | 
|
| 
 
0fb331032f02
more compact representation of special string values;
 
wenzelm 
parents: 
57831 
diff
changeset
 | 
61  | 
|
| 48996 | 62  | 
/* separated chunks */  | 
| 36688 | 63  | 
|
64  | 
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: 
51616 
diff
changeset
 | 
65  | 
  {
 | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
66  | 
val result = new mutable.ListBuffer[A]  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
67  | 
var first = true  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
68  | 
    for (x <- list) {
 | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
69  | 
      if (first) {
 | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
70  | 
first = false  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
71  | 
result += x  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
72  | 
}  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
73  | 
      else {
 | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
74  | 
result += s  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
75  | 
result += x  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
76  | 
}  | 
| 36688 | 77  | 
}  | 
| 
51981
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
78  | 
result.toList  | 
| 
 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 
wenzelm 
parents: 
51616 
diff
changeset
 | 
79  | 
}  | 
| 36688 | 80  | 
|
| 56600 | 81  | 
def separated_chunks(sep: Char => Boolean, source: CharSequence): Iterator[CharSequence] =  | 
| 48996 | 82  | 
    new Iterator[CharSequence] {
 | 
83  | 
private val end = source.length  | 
|
84  | 
private def next_chunk(i: Int): Option[(CharSequence, Int)] =  | 
|
85  | 
      {
 | 
|
86  | 
        if (i < end) {
 | 
|
| 56600 | 87  | 
var j = i; do j += 1 while (j < end && !sep(source.charAt(j)))  | 
| 48996 | 88  | 
Some((source.subSequence(i + 1, j), j))  | 
89  | 
}  | 
|
90  | 
else None  | 
|
| 43598 | 91  | 
}  | 
| 48996 | 92  | 
private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1)  | 
93  | 
||
94  | 
def hasNext(): Boolean = state.isDefined  | 
|
95  | 
def next(): CharSequence =  | 
|
96  | 
        state match {
 | 
|
97  | 
          case Some((s, i)) => { state = next_chunk(i); s }
 | 
|
98  | 
case None => Iterator.empty.next()  | 
|
99  | 
}  | 
|
| 43598 | 100  | 
}  | 
101  | 
||
| 48996 | 102  | 
def space_explode(sep: Char, str: String): List[String] =  | 
| 56600 | 103  | 
separated_chunks(_ == sep, str).map(_.toString).toList  | 
| 48996 | 104  | 
|
105  | 
||
106  | 
/* lines */  | 
|
107  | 
||
| 
51983
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
108  | 
def terminate_lines(lines: Iterable[CharSequence]): Iterable[CharSequence] =  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
109  | 
    new Iterable[CharSequence] {
 | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
110  | 
def iterator: Iterator[CharSequence] =  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
111  | 
lines.iterator.map(line => new Line_Termination(line))  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
112  | 
}  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
113  | 
|
| 48996 | 114  | 
  def cat_lines(lines: TraversableOnce[String]): String = lines.mkString("\n")
 | 
115  | 
||
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
116  | 
  def split_lines(str: String): List[String] = space_explode('\n', str)
 | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
117  | 
|
| 48996 | 118  | 
def first_line(source: CharSequence): String =  | 
119  | 
  {
 | 
|
| 56600 | 120  | 
val lines = separated_chunks(_ == '\n', source)  | 
| 48996 | 121  | 
if (lines.hasNext) lines.next.toString  | 
122  | 
else ""  | 
|
123  | 
}  | 
|
124  | 
||
| 50847 | 125  | 
|
126  | 
/* strings */  | 
|
127  | 
||
128  | 
def try_unprefix(prfx: String, s: String): Option[String] =  | 
|
129  | 
if (s.startsWith(prfx)) Some(s.substring(prfx.length)) else None  | 
|
130  | 
||
| 55033 | 131  | 
def try_unsuffix(sffx: String, s: String): Option[String] =  | 
132  | 
if (s.endsWith(sffx)) Some(s.substring(0, s.length - sffx.length)) else None  | 
|
133  | 
||
| 
52444
 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 
wenzelm 
parents: 
51983 
diff
changeset
 | 
134  | 
def trim_line(s: String): String =  | 
| 
 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 
wenzelm 
parents: 
51983 
diff
changeset
 | 
135  | 
    if (s.endsWith("\r\n")) s.substring(0, s.length - 2)
 | 
| 
 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 
wenzelm 
parents: 
51983 
diff
changeset
 | 
136  | 
    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: 
51983 
diff
changeset
 | 
137  | 
else s  | 
| 
 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 
wenzelm 
parents: 
51983 
diff
changeset
 | 
138  | 
|
| 43598 | 139  | 
|
| 48996 | 140  | 
/* quote */  | 
| 46196 | 141  | 
|
| 43598 | 142  | 
def quote(s: String): String = "\"" + s + "\""  | 
| 
56843
 
b2bfcd8cda80
support for path completion based on file-system content;
 
wenzelm 
parents: 
56730 
diff
changeset
 | 
143  | 
|
| 
 
b2bfcd8cda80
support for path completion based on file-system content;
 
wenzelm 
parents: 
56730 
diff
changeset
 | 
144  | 
def try_unquote(s: String): Option[String] =  | 
| 
 
b2bfcd8cda80
support for path completion based on file-system content;
 
wenzelm 
parents: 
56730 
diff
changeset
 | 
145  | 
    if (s.startsWith("\"") && s.endsWith("\"")) Some(s.substring(1, s.length - 1))
 | 
| 
 
b2bfcd8cda80
support for path completion based on file-system content;
 
wenzelm 
parents: 
56730 
diff
changeset
 | 
146  | 
else None  | 
| 
 
b2bfcd8cda80
support for path completion based on file-system content;
 
wenzelm 
parents: 
56730 
diff
changeset
 | 
147  | 
|
| 43598 | 148  | 
  def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | 
| 48362 | 149  | 
  def commas_quote(ss: Iterable[String]): String = ss.iterator.map(quote).mkString(", ")
 | 
| 43598 | 150  | 
|
| 36688 | 151  | 
|
| 
51983
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
152  | 
/* CharSequence */  | 
| 34141 | 153  | 
|
154  | 
class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence  | 
|
155  | 
  {
 | 
|
156  | 
require(0 <= start && start <= end && end <= text.length)  | 
|
157  | 
||
158  | 
def this(text: CharSequence) = this(text, 0, text.length)  | 
|
159  | 
||
160  | 
def length: Int = end - start  | 
|
161  | 
def charAt(i: Int): Char = text.charAt(end - i - 1)  | 
|
162  | 
||
163  | 
def subSequence(i: Int, j: Int): CharSequence =  | 
|
164  | 
if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i)  | 
|
165  | 
else throw new IndexOutOfBoundsException  | 
|
166  | 
||
167  | 
override def toString: String =  | 
|
168  | 
    {
 | 
|
169  | 
val buf = new StringBuilder(length)  | 
|
170  | 
for (i <- 0 until length)  | 
|
171  | 
buf.append(charAt(i))  | 
|
172  | 
buf.toString  | 
|
173  | 
}  | 
|
174  | 
}  | 
|
175  | 
||
| 
51983
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
176  | 
class Line_Termination(text: CharSequence) extends CharSequence  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
177  | 
  {
 | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
178  | 
def length: Int = text.length + 1  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
179  | 
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: 
51981 
diff
changeset
 | 
180  | 
def subSequence(i: Int, j: Int): CharSequence =  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
181  | 
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: 
51981 
diff
changeset
 | 
182  | 
else text.subSequence(i, j)  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
183  | 
override def toString: String = text.toString + "\n"  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
184  | 
}  | 
| 
 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 
wenzelm 
parents: 
51981 
diff
changeset
 | 
185  | 
|
| 34141 | 186  | 
|
| 56686 | 187  | 
/* canonical list operations */  | 
188  | 
||
| 56688 | 189  | 
def member[A, B](xs: List[A])(x: B): Boolean = xs.exists(_ == x)  | 
190  | 
def insert[A](x: A)(xs: List[A]): List[A] = if (xs.contains(x)) xs else x :: xs  | 
|
191  | 
def remove[A, B](x: B)(xs: List[A]): List[A] = if (member(xs)(x)) xs.filterNot(_ == x) else xs  | 
|
192  | 
def update[A](x: A)(xs: List[A]): List[A] = x :: remove(x)(xs)  | 
|
| 34136 | 193  | 
}  | 
| 43652 | 194  | 
|
195  | 
||
196  | 
class Basic_Library  | 
|
197  | 
{
 | 
|
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
198  | 
val ERROR = Library.ERROR  | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
199  | 
val error = Library.error _  | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
200  | 
val cat_error = Library.cat_error _  | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
201  | 
|
| 43652 | 202  | 
val space_explode = Library.space_explode _  | 
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
203  | 
val split_lines = Library.split_lines _  | 
| 46196 | 204  | 
val cat_lines = Library.cat_lines _  | 
| 43652 | 205  | 
val quote = Library.quote _  | 
206  | 
val commas = Library.commas _  | 
|
207  | 
val commas_quote = Library.commas_quote _  | 
|
208  | 
}  |