| author | wenzelm | 
| Mon, 02 May 2011 21:59:47 +0200 | |
| changeset 42631 | 028f94955436 | 
| parent 40848 | 8662b9b1f123 | 
| child 43442 | e1fff67b23ac | 
| 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  | 
|
16  | 
||
17  | 
||
| 34136 | 18  | 
object Library  | 
19  | 
{
 | 
|
| 37035 | 20  | 
/* partial functions */  | 
21  | 
||
22  | 
  def undefined[A, B] = new PartialFunction[A, B] {
 | 
|
23  | 
    def apply(x: A): B = throw new NoSuchElementException("undefined")
 | 
|
24  | 
def isDefinedAt(x: A) = false  | 
|
25  | 
}  | 
|
26  | 
||
27  | 
||
| 36688 | 28  | 
/* separate */  | 
29  | 
||
30  | 
def separate[A](s: A, list: List[A]): List[A] =  | 
|
31  | 
    list match {
 | 
|
32  | 
case x :: xs if !xs.isEmpty => x :: s :: separate(s, xs)  | 
|
33  | 
case _ => list  | 
|
34  | 
}  | 
|
35  | 
||
36  | 
||
| 34141 | 37  | 
/* reverse CharSequence */  | 
38  | 
||
39  | 
class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence  | 
|
40  | 
  {
 | 
|
41  | 
require(0 <= start && start <= end && end <= text.length)  | 
|
42  | 
||
43  | 
def this(text: CharSequence) = this(text, 0, text.length)  | 
|
44  | 
||
45  | 
def length: Int = end - start  | 
|
46  | 
def charAt(i: Int): Char = text.charAt(end - i - 1)  | 
|
47  | 
||
48  | 
def subSequence(i: Int, j: Int): CharSequence =  | 
|
49  | 
if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i)  | 
|
50  | 
else throw new IndexOutOfBoundsException  | 
|
51  | 
||
52  | 
override def toString: String =  | 
|
53  | 
    {
 | 
|
54  | 
val buf = new StringBuilder(length)  | 
|
55  | 
for (i <- 0 until length)  | 
|
56  | 
buf.append(charAt(i))  | 
|
57  | 
buf.toString  | 
|
58  | 
}  | 
|
59  | 
}  | 
|
60  | 
||
61  | 
||
| 36685 | 62  | 
/* iterate over chunks (cf. space_explode/split_lines in ML) */  | 
63  | 
||
64  | 
def chunks(source: CharSequence, sep: Char = '\n') = new Iterator[CharSequence]  | 
|
65  | 
  {
 | 
|
66  | 
private val end = source.length  | 
|
67  | 
private def next_chunk(i: Int): Option[(CharSequence, Int)] =  | 
|
68  | 
    {
 | 
|
69  | 
      if (i < end) {
 | 
|
70  | 
var j = i; do j += 1 while (j < end && source.charAt(j) != sep)  | 
|
71  | 
Some((source.subSequence(i + 1, j), j))  | 
|
72  | 
}  | 
|
73  | 
else None  | 
|
74  | 
}  | 
|
75  | 
private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1)  | 
|
76  | 
||
77  | 
def hasNext(): Boolean = state.isDefined  | 
|
78  | 
def next(): CharSequence =  | 
|
79  | 
      state match {
 | 
|
80  | 
        case Some((s, i)) => { state = next_chunk(i); s }
 | 
|
| 38583 | 81  | 
case None => Iterator.empty.next()  | 
| 36685 | 82  | 
}  | 
83  | 
}  | 
|
84  | 
||
| 40475 | 85  | 
def first_line(source: CharSequence): String =  | 
86  | 
  {
 | 
|
87  | 
val lines = chunks(source)  | 
|
88  | 
if (lines.hasNext) lines.next.toString  | 
|
89  | 
else ""  | 
|
90  | 
}  | 
|
91  | 
||
| 36685 | 92  | 
|
| 34216 | 93  | 
/* simple dialogs */  | 
94  | 
||
95  | 
private def simple_dialog(kind: Int, default_title: String)  | 
|
96  | 
(parent: Component, title: String, message: Any*)  | 
|
97  | 
  {
 | 
|
| 36791 | 98  | 
    Swing_Thread.now {
 | 
| 
38232
 
00b72526dc64
simple_dialog: allow scala.swing.Component as well;
 
wenzelm 
parents: 
37686 
diff
changeset
 | 
99  | 
      val java_message = message map { case x: scala.swing.Component => x.peer case x => x }
 | 
| 36791 | 100  | 
JOptionPane.showMessageDialog(parent,  | 
| 
38232
 
00b72526dc64
simple_dialog: allow scala.swing.Component as well;
 
wenzelm 
parents: 
37686 
diff
changeset
 | 
101  | 
java_message.toArray.asInstanceOf[Array[AnyRef]],  | 
| 36791 | 102  | 
if (title == null) default_title else title, kind)  | 
103  | 
}  | 
|
| 34216 | 104  | 
}  | 
105  | 
||
106  | 
def dialog = simple_dialog(JOptionPane.PLAIN_MESSAGE, null) _  | 
|
107  | 
def warning_dialog = simple_dialog(JOptionPane.WARNING_MESSAGE, "Warning") _  | 
|
108  | 
def error_dialog = simple_dialog(JOptionPane.ERROR_MESSAGE, "Error") _  | 
|
109  | 
||
110  | 
||
| 37018 | 111  | 
/* zoom box */  | 
112  | 
||
| 37048 | 113  | 
class Zoom_Box(apply_factor: Int => Unit) extends ComboBox[String](  | 
114  | 
    List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%"))
 | 
|
115  | 
  {
 | 
|
116  | 
val Factor = "([0-9]+)%?"r  | 
|
117  | 
def parse(text: String): Int =  | 
|
118  | 
      text match {
 | 
|
119  | 
case Factor(s) =>  | 
|
120  | 
val i = Integer.parseInt(s)  | 
|
121  | 
if (10 <= i && i <= 1000) i else 100  | 
|
122  | 
case _ => 100  | 
|
123  | 
}  | 
|
124  | 
def print(i: Int): String = i.toString + "%"  | 
|
| 37018 | 125  | 
|
| 37048 | 126  | 
makeEditable()(c => new ComboBox.BuiltInEditor(c)(text => print(parse(text)), x => x))  | 
127  | 
    reactions += {
 | 
|
128  | 
case SelectionChanged(_) => apply_factor(parse(selection.item))  | 
|
| 37018 | 129  | 
}  | 
| 37048 | 130  | 
listenTo(selection)  | 
131  | 
selection.index = 3  | 
|
132  | 
    prototypeDisplayValue = Some("00000%")
 | 
|
133  | 
}  | 
|
| 37018 | 134  | 
|
135  | 
||
| 34136 | 136  | 
/* timing */  | 
137  | 
||
| 34314 | 138  | 
def timeit[A](message: String)(e: => A) =  | 
| 34136 | 139  | 
  {
 | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
40475 
diff
changeset
 | 
140  | 
val start = System.currentTimeMillis()  | 
| 34136 | 141  | 
val result = Exn.capture(e)  | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
40475 
diff
changeset
 | 
142  | 
val stop = System.currentTimeMillis()  | 
| 34314 | 143  | 
System.err.println(  | 
| 37686 | 144  | 
(if (message == null || message.isEmpty) "" else message + ": ") +  | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
40475 
diff
changeset
 | 
145  | 
new Time(stop - start).message + " elapsed time")  | 
| 34136 | 146  | 
Exn.release(result)  | 
147  | 
}  | 
|
148  | 
}  |