| author | nipkow | 
| Mon, 02 Apr 2012 20:12:10 +0200 | |
| changeset 47302 | 70239da25ef6 | 
| parent 46997 | 395b7277ed76 | 
| child 47867 | dd9cbe708e6b | 
| 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 | |
| 38636 
b7647ca7de5a
module for simplified thread operations (Scala version);
 wenzelm parents: 
38583diff
changeset | 11 | import java.lang.System | 
| 34216 | 12 | import java.awt.Component | 
| 13 | import javax.swing.JOptionPane | |
| 34136 | 14 | |
| 37018 | 15 | import scala.swing.ComboBox | 
| 16 | import scala.swing.event.SelectionChanged | |
| 43598 | 17 | import scala.collection.mutable | 
| 37018 | 18 | |
| 19 | ||
| 34136 | 20 | object Library | 
| 21 | {
 | |
| 43652 | 22 | /* user errors */ | 
| 23 | ||
| 24 | object ERROR | |
| 25 |   {
 | |
| 26 | def apply(message: String): Throwable = new RuntimeException(message) | |
| 27 | def unapply(exn: Throwable): Option[String] = | |
| 28 |       exn match {
 | |
| 44158 | 29 | case e: RuntimeException => Some(Exn.message(e)) | 
| 43652 | 30 | case _ => None | 
| 31 | } | |
| 32 | } | |
| 33 | ||
| 34 | def error(message: String): Nothing = throw ERROR(message) | |
| 35 | ||
| 36 | def cat_error(msg1: String, msg2: String): Nothing = | |
| 37 | if (msg1 == "") error(msg1) | |
| 38 | else error(msg1 + "\n" + msg2) | |
| 39 | ||
| 40 | ||
| 43598 | 41 | /* lists */ | 
| 36688 | 42 | |
| 43 | def separate[A](s: A, list: List[A]): List[A] = | |
| 44 |     list match {
 | |
| 45 | case x :: xs if !xs.isEmpty => x :: s :: separate(s, xs) | |
| 46 | case _ => list | |
| 47 | } | |
| 48 | ||
| 43598 | 49 | def space_explode(sep: Char, str: String): List[String] = | 
| 50 | if (str.isEmpty) Nil | |
| 51 |     else {
 | |
| 52 | val result = new mutable.ListBuffer[String] | |
| 53 | var start = 0 | |
| 54 | var finished = false | |
| 55 |       while (!finished) {
 | |
| 56 | val i = str.indexOf(sep, start) | |
| 57 |         if (i == -1) { result += str.substring(start); finished = true }
 | |
| 58 |         else { result += str.substring(start, i); start = i + 1 }
 | |
| 59 | } | |
| 60 | result.toList | |
| 61 | } | |
| 62 | ||
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 63 |   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 | 64 | |
| 43598 | 65 | |
| 43652 | 66 | /* iterate over chunks (cf. space_explode) */ | 
| 67 | ||
| 68 | def chunks(source: CharSequence, sep: Char = '\n') = new Iterator[CharSequence] | |
| 69 |   {
 | |
| 70 | private val end = source.length | |
| 71 | private def next_chunk(i: Int): Option[(CharSequence, Int)] = | |
| 72 |     {
 | |
| 73 |       if (i < end) {
 | |
| 74 | var j = i; do j += 1 while (j < end && source.charAt(j) != sep) | |
| 75 | Some((source.subSequence(i + 1, j), j)) | |
| 76 | } | |
| 77 | else None | |
| 78 | } | |
| 79 | private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1) | |
| 80 | ||
| 81 | def hasNext(): Boolean = state.isDefined | |
| 82 | def next(): CharSequence = | |
| 83 |       state match {
 | |
| 84 |         case Some((s, i)) => { state = next_chunk(i); s }
 | |
| 85 | case None => Iterator.empty.next() | |
| 86 | } | |
| 87 | } | |
| 88 | ||
| 89 | def first_line(source: CharSequence): String = | |
| 90 |   {
 | |
| 91 | val lines = chunks(source) | |
| 92 | if (lines.hasNext) lines.next.toString | |
| 93 | else "" | |
| 94 | } | |
| 95 | ||
| 96 | ||
| 43598 | 97 | /* strings */ | 
| 98 | ||
| 46196 | 99 |   def cat_lines(lines: TraversableOnce[String]): String = lines.mkString("\n")
 | 
| 100 | ||
| 43598 | 101 | def quote(s: String): String = "\"" + s + "\"" | 
| 102 |   def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | |
| 103 |   def commas_quote(ss: Iterable[String]): String = ss.iterator.mkString("\"", ", ", "\"")
 | |
| 104 | ||
| 36688 | 105 | |
| 34141 | 106 | /* reverse CharSequence */ | 
| 107 | ||
| 108 | class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence | |
| 109 |   {
 | |
| 110 | require(0 <= start && start <= end && end <= text.length) | |
| 111 | ||
| 112 | def this(text: CharSequence) = this(text, 0, text.length) | |
| 113 | ||
| 114 | def length: Int = end - start | |
| 115 | def charAt(i: Int): Char = text.charAt(end - i - 1) | |
| 116 | ||
| 117 | def subSequence(i: Int, j: Int): CharSequence = | |
| 118 | if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i) | |
| 119 | else throw new IndexOutOfBoundsException | |
| 120 | ||
| 121 | override def toString: String = | |
| 122 |     {
 | |
| 123 | val buf = new StringBuilder(length) | |
| 124 | for (i <- 0 until length) | |
| 125 | buf.append(charAt(i)) | |
| 126 | buf.toString | |
| 127 | } | |
| 128 | } | |
| 129 | ||
| 130 | ||
| 34216 | 131 | /* simple dialogs */ | 
| 132 | ||
| 46997 | 133 | private def simple_dialog(kind: Int, default_title: String, | 
| 134 | parent: Component, title: String, message: Seq[Any]) | |
| 34216 | 135 |   {
 | 
| 36791 | 136 |     Swing_Thread.now {
 | 
| 38232 
00b72526dc64
simple_dialog: allow scala.swing.Component as well;
 wenzelm parents: 
37686diff
changeset | 137 |       val java_message = message map { case x: scala.swing.Component => x.peer case x => x }
 | 
| 36791 | 138 | JOptionPane.showMessageDialog(parent, | 
| 38232 
00b72526dc64
simple_dialog: allow scala.swing.Component as well;
 wenzelm parents: 
37686diff
changeset | 139 | java_message.toArray.asInstanceOf[Array[AnyRef]], | 
| 36791 | 140 | if (title == null) default_title else title, kind) | 
| 141 | } | |
| 34216 | 142 | } | 
| 143 | ||
| 46997 | 144 | def dialog(parent: Component, title: String, message: Any*) = | 
| 145 | simple_dialog(JOptionPane.PLAIN_MESSAGE, null, parent, title, message) | |
| 146 | ||
| 147 | def warning_dialog(parent: Component, title: String, message: Any*) = | |
| 148 | simple_dialog(JOptionPane.WARNING_MESSAGE, "Warning", parent, title, message) | |
| 149 | ||
| 150 | def error_dialog(parent: Component, title: String, message: Any*) = | |
| 151 | simple_dialog(JOptionPane.ERROR_MESSAGE, "Error", parent, title, message) | |
| 34216 | 152 | |
| 44573 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 wenzelm parents: 
44158diff
changeset | 153 | 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: 
44158diff
changeset | 154 |     Swing_Thread.now {
 | 
| 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 wenzelm parents: 
44158diff
changeset | 155 |       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: 
44158diff
changeset | 156 | JOptionPane.showConfirmDialog(parent, | 
| 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 wenzelm parents: 
44158diff
changeset | 157 | java_message.toArray.asInstanceOf[Array[AnyRef]], title, | 
| 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 wenzelm parents: 
44158diff
changeset | 158 | option_type, JOptionPane.QUESTION_MESSAGE) | 
| 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 wenzelm parents: 
44158diff
changeset | 159 | } | 
| 
51f8895b9ad9
some dialog for auto loading of required files (still inactive);
 wenzelm parents: 
44158diff
changeset | 160 | |
| 34216 | 161 | |
| 37018 | 162 | /* zoom box */ | 
| 163 | ||
| 37048 | 164 | class Zoom_Box(apply_factor: Int => Unit) extends ComboBox[String]( | 
| 165 |     List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%"))
 | |
| 166 |   {
 | |
| 167 | val Factor = "([0-9]+)%?"r | |
| 168 | def parse(text: String): Int = | |
| 169 |       text match {
 | |
| 170 | case Factor(s) => | |
| 171 | val i = Integer.parseInt(s) | |
| 172 | if (10 <= i && i <= 1000) i else 100 | |
| 173 | case _ => 100 | |
| 174 | } | |
| 175 | def print(i: Int): String = i.toString + "%" | |
| 37018 | 176 | |
| 37048 | 177 | makeEditable()(c => new ComboBox.BuiltInEditor(c)(text => print(parse(text)), x => x)) | 
| 178 |     reactions += {
 | |
| 179 | case SelectionChanged(_) => apply_factor(parse(selection.item)) | |
| 37018 | 180 | } | 
| 37048 | 181 | listenTo(selection) | 
| 182 | selection.index = 3 | |
| 183 |     prototypeDisplayValue = Some("00000%")
 | |
| 184 | } | |
| 34136 | 185 | } | 
| 43652 | 186 | |
| 187 | ||
| 188 | class Basic_Library | |
| 189 | {
 | |
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 190 | val ERROR = Library.ERROR | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 191 | val error = Library.error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 192 | val cat_error = Library.cat_error _ | 
| 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 193 | |
| 43652 | 194 | val space_explode = Library.space_explode _ | 
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 195 | val split_lines = Library.split_lines _ | 
| 46196 | 196 | val cat_lines = Library.cat_lines _ | 
| 43652 | 197 | val quote = Library.quote _ | 
| 198 | val commas = Library.commas _ | |
| 199 | val commas_quote = Library.commas_quote _ | |
| 200 | } |