| author | wenzelm | 
| Sun, 18 Sep 2011 00:05:22 +0200 | |
| changeset 44960 | 640c2b957f16 | 
| parent 44617 | 5db68864a967 | 
| child 45249 | b769a3a370ad | 
| permissions | -rw-r--r-- | 
| 34136 | 1  | 
/* Title: Pure/library.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Basic library.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
| 38258 | 9  | 
|
| 
38636
 
b7647ca7de5a
module for simplified thread operations (Scala version);
 
wenzelm 
parents: 
38583 
diff
changeset
 | 
10  | 
import java.lang.System  | 
| 34216 | 11  | 
import java.awt.Component  | 
12  | 
import javax.swing.JOptionPane  | 
|
| 34136 | 13  | 
|
| 37018 | 14  | 
import scala.swing.ComboBox  | 
15  | 
import scala.swing.event.SelectionChanged  | 
|
| 43598 | 16  | 
import scala.collection.mutable  | 
| 44617 | 17  | 
import scala.math.Ordering  | 
| 44609 | 18  | 
import scala.util.Sorting  | 
| 37018 | 19  | 
|
20  | 
||
| 34136 | 21  | 
object Library  | 
22  | 
{
 | 
|
| 43652 | 23  | 
/* user errors */  | 
24  | 
||
25  | 
object ERROR  | 
|
26  | 
  {
 | 
|
27  | 
def apply(message: String): Throwable = new RuntimeException(message)  | 
|
28  | 
def unapply(exn: Throwable): Option[String] =  | 
|
29  | 
      exn match {
 | 
|
| 44158 | 30  | 
case e: RuntimeException => Some(Exn.message(e))  | 
| 43652 | 31  | 
case _ => None  | 
32  | 
}  | 
|
33  | 
}  | 
|
34  | 
||
35  | 
def error(message: String): Nothing = throw ERROR(message)  | 
|
36  | 
||
37  | 
def cat_error(msg1: String, msg2: String): Nothing =  | 
|
38  | 
if (msg1 == "") error(msg1)  | 
|
39  | 
else error(msg1 + "\n" + msg2)  | 
|
40  | 
||
41  | 
||
| 43598 | 42  | 
/* lists */  | 
| 36688 | 43  | 
|
44  | 
def separate[A](s: A, list: List[A]): List[A] =  | 
|
45  | 
    list match {
 | 
|
46  | 
case x :: xs if !xs.isEmpty => x :: s :: separate(s, xs)  | 
|
47  | 
case _ => list  | 
|
48  | 
}  | 
|
49  | 
||
| 43598 | 50  | 
def space_explode(sep: Char, str: String): List[String] =  | 
51  | 
if (str.isEmpty) Nil  | 
|
52  | 
    else {
 | 
|
53  | 
val result = new mutable.ListBuffer[String]  | 
|
54  | 
var start = 0  | 
|
55  | 
var finished = false  | 
|
56  | 
      while (!finished) {
 | 
|
57  | 
val i = str.indexOf(sep, start)  | 
|
58  | 
        if (i == -1) { result += str.substring(start); finished = true }
 | 
|
59  | 
        else { result += str.substring(start, i); start = i + 1 }
 | 
|
60  | 
}  | 
|
61  | 
result.toList  | 
|
62  | 
}  | 
|
63  | 
||
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
64  | 
  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
 | 
65  | 
|
| 44617 | 66  | 
def sort_wrt[A](key: A => String, args: Seq[A]): List[A] =  | 
| 44609 | 67  | 
  {
 | 
| 44617 | 68  | 
val ordering: Ordering[A] =  | 
69  | 
      new Ordering[A] { def compare(x: A, y: A): Int = key(x) compare key(y) }
 | 
|
70  | 
val a = (new Array[Any](args.length)).asInstanceOf[Array[A]]  | 
|
71  | 
for ((x, i) <- args.iterator zipWithIndex) a(i) = x  | 
|
72  | 
Sorting.quickSort[A](a)(ordering)  | 
|
| 44609 | 73  | 
a.toList  | 
74  | 
}  | 
|
75  | 
||
| 44617 | 76  | 
def sort_strings(args: Seq[String]): List[String] = sort_wrt((x: String) => x, args)  | 
77  | 
||
| 43598 | 78  | 
|
| 43652 | 79  | 
/* iterate over chunks (cf. space_explode) */  | 
80  | 
||
81  | 
def chunks(source: CharSequence, sep: Char = '\n') = new Iterator[CharSequence]  | 
|
82  | 
  {
 | 
|
83  | 
private val end = source.length  | 
|
84  | 
private def next_chunk(i: Int): Option[(CharSequence, Int)] =  | 
|
85  | 
    {
 | 
|
86  | 
      if (i < end) {
 | 
|
87  | 
var j = i; do j += 1 while (j < end && source.charAt(j) != sep)  | 
|
88  | 
Some((source.subSequence(i + 1, j), j))  | 
|
89  | 
}  | 
|
90  | 
else None  | 
|
91  | 
}  | 
|
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  | 
}  | 
|
100  | 
}  | 
|
101  | 
||
102  | 
def first_line(source: CharSequence): String =  | 
|
103  | 
  {
 | 
|
104  | 
val lines = chunks(source)  | 
|
105  | 
if (lines.hasNext) lines.next.toString  | 
|
106  | 
else ""  | 
|
107  | 
}  | 
|
108  | 
||
109  | 
||
| 43598 | 110  | 
/* strings */  | 
111  | 
||
112  | 
def quote(s: String): String = "\"" + s + "\""  | 
|
113  | 
  def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | 
|
114  | 
  def commas_quote(ss: Iterable[String]): String = ss.iterator.mkString("\"", ", ", "\"")
 | 
|
115  | 
||
| 36688 | 116  | 
|
| 34141 | 117  | 
/* reverse CharSequence */  | 
118  | 
||
119  | 
class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence  | 
|
120  | 
  {
 | 
|
121  | 
require(0 <= start && start <= end && end <= text.length)  | 
|
122  | 
||
123  | 
def this(text: CharSequence) = this(text, 0, text.length)  | 
|
124  | 
||
125  | 
def length: Int = end - start  | 
|
126  | 
def charAt(i: Int): Char = text.charAt(end - i - 1)  | 
|
127  | 
||
128  | 
def subSequence(i: Int, j: Int): CharSequence =  | 
|
129  | 
if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i)  | 
|
130  | 
else throw new IndexOutOfBoundsException  | 
|
131  | 
||
132  | 
override def toString: String =  | 
|
133  | 
    {
 | 
|
134  | 
val buf = new StringBuilder(length)  | 
|
135  | 
for (i <- 0 until length)  | 
|
136  | 
buf.append(charAt(i))  | 
|
137  | 
buf.toString  | 
|
138  | 
}  | 
|
139  | 
}  | 
|
140  | 
||
141  | 
||
| 44960 | 142  | 
/* graph traversal */  | 
143  | 
||
144  | 
def topological_order[A](next: A => Iterable[A], starts: Iterable[A]): List[A] =  | 
|
145  | 
  {
 | 
|
146  | 
type Reached = (List[A], Set[A])  | 
|
147  | 
def reach(reached: Reached, x: A): Reached =  | 
|
148  | 
    {
 | 
|
149  | 
val (rs, r_set) = reached  | 
|
150  | 
if (r_set(x)) reached  | 
|
151  | 
      else {
 | 
|
152  | 
val (rs1, r_set1) = reachs((rs, r_set + x), next(x))  | 
|
153  | 
(x :: rs1, r_set1)  | 
|
154  | 
}  | 
|
155  | 
}  | 
|
156  | 
def reachs(reached: Reached, xs: Iterable[A]): Reached = (reached /: xs)(reach)  | 
|
157  | 
||
158  | 
reachs((Nil, Set.empty), starts)._1.reverse  | 
|
159  | 
}  | 
|
160  | 
||
161  | 
||
| 34216 | 162  | 
/* simple dialogs */  | 
163  | 
||
164  | 
private def simple_dialog(kind: Int, default_title: String)  | 
|
165  | 
(parent: Component, title: String, message: Any*)  | 
|
166  | 
  {
 | 
|
| 36791 | 167  | 
    Swing_Thread.now {
 | 
| 
38232
 
00b72526dc64
simple_dialog: allow scala.swing.Component as well;
 
wenzelm 
parents: 
37686 
diff
changeset
 | 
168  | 
      val java_message = message map { case x: scala.swing.Component => x.peer case x => x }
 | 
| 36791 | 169  | 
JOptionPane.showMessageDialog(parent,  | 
| 
38232
 
00b72526dc64
simple_dialog: allow scala.swing.Component as well;
 
wenzelm 
parents: 
37686 
diff
changeset
 | 
170  | 
java_message.toArray.asInstanceOf[Array[AnyRef]],  | 
| 36791 | 171  | 
if (title == null) default_title else title, kind)  | 
172  | 
}  | 
|
| 34216 | 173  | 
}  | 
174  | 
||
175  | 
def dialog = simple_dialog(JOptionPane.PLAIN_MESSAGE, null) _  | 
|
176  | 
def warning_dialog = simple_dialog(JOptionPane.WARNING_MESSAGE, "Warning") _  | 
|
177  | 
def error_dialog = simple_dialog(JOptionPane.ERROR_MESSAGE, "Error") _  | 
|
178  | 
||
| 
44573
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
179  | 
def confirm_dialog(parent: Component, title: String, option_type: Int, message: Any*): Int =  | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
180  | 
    Swing_Thread.now {
 | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
181  | 
      val java_message = message map { case x: scala.swing.Component => x.peer case x => x }
 | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
182  | 
JOptionPane.showConfirmDialog(parent,  | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
183  | 
java_message.toArray.asInstanceOf[Array[AnyRef]], title,  | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
184  | 
option_type, JOptionPane.QUESTION_MESSAGE)  | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
185  | 
}  | 
| 
 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 
wenzelm 
parents: 
44158 
diff
changeset
 | 
186  | 
|
| 34216 | 187  | 
|
| 37018 | 188  | 
/* zoom box */  | 
189  | 
||
| 37048 | 190  | 
class Zoom_Box(apply_factor: Int => Unit) extends ComboBox[String](  | 
191  | 
    List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%"))
 | 
|
192  | 
  {
 | 
|
193  | 
val Factor = "([0-9]+)%?"r  | 
|
194  | 
def parse(text: String): Int =  | 
|
195  | 
      text match {
 | 
|
196  | 
case Factor(s) =>  | 
|
197  | 
val i = Integer.parseInt(s)  | 
|
198  | 
if (10 <= i && i <= 1000) i else 100  | 
|
199  | 
case _ => 100  | 
|
200  | 
}  | 
|
201  | 
def print(i: Int): String = i.toString + "%"  | 
|
| 37018 | 202  | 
|
| 37048 | 203  | 
makeEditable()(c => new ComboBox.BuiltInEditor(c)(text => print(parse(text)), x => x))  | 
204  | 
    reactions += {
 | 
|
205  | 
case SelectionChanged(_) => apply_factor(parse(selection.item))  | 
|
| 37018 | 206  | 
}  | 
| 37048 | 207  | 
listenTo(selection)  | 
208  | 
selection.index = 3  | 
|
209  | 
    prototypeDisplayValue = Some("00000%")
 | 
|
210  | 
}  | 
|
| 37018 | 211  | 
|
212  | 
||
| 34136 | 213  | 
/* timing */  | 
214  | 
||
| 34314 | 215  | 
def timeit[A](message: String)(e: => A) =  | 
| 34136 | 216  | 
  {
 | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
40475 
diff
changeset
 | 
217  | 
val start = System.currentTimeMillis()  | 
| 34136 | 218  | 
val result = Exn.capture(e)  | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
40475 
diff
changeset
 | 
219  | 
val stop = System.currentTimeMillis()  | 
| 34314 | 220  | 
System.err.println(  | 
| 37686 | 221  | 
(if (message == null || message.isEmpty) "" else message + ": ") +  | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
40475 
diff
changeset
 | 
222  | 
new Time(stop - start).message + " elapsed time")  | 
| 34136 | 223  | 
Exn.release(result)  | 
224  | 
}  | 
|
225  | 
}  | 
|
| 43652 | 226  | 
|
227  | 
||
228  | 
class Basic_Library  | 
|
229  | 
{
 | 
|
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
230  | 
val ERROR = Library.ERROR  | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
231  | 
val error = Library.error _  | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
232  | 
val cat_error = Library.cat_error _  | 
| 
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
233  | 
|
| 43652 | 234  | 
val space_explode = Library.space_explode _  | 
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
235  | 
val split_lines = Library.split_lines _  | 
| 43652 | 236  | 
|
237  | 
val quote = Library.quote _  | 
|
238  | 
val commas = Library.commas _  | 
|
239  | 
val commas_quote = Library.commas_quote _  | 
|
240  | 
}  |