| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 15 Mar 2024 17:57:03 +0100 | |
| changeset 79899 | c73a36081b1c | 
| parent 79870 | 510fe8c3d9b8 | 
| child 80270 | 1d4300506338 | 
| 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 | |
| 63781 | 10 | import scala.annotation.tailrec | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 11 | import scala.collection.mutable | 
| 59224 | 12 | import scala.util.matching.Regex | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 13 | |
| 37018 | 14 | |
| 75393 | 15 | object Library {
 | 
| 63789 | 16 | /* resource management */ | 
| 17 | ||
| 75393 | 18 |   def using[A <: AutoCloseable, B](a: A)(f: A => B): B = {
 | 
| 69393 
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
 wenzelm parents: 
68715diff
changeset | 19 |     try { f(a) }
 | 
| 
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
 wenzelm parents: 
68715diff
changeset | 20 |     finally { if (a != null) a.close() }
 | 
| 63789 | 21 | } | 
| 22 | ||
| 77515 | 23 | def using_option[A <: AutoCloseable, B](opt: Option[A])(f: A => B): Option[B] = | 
| 24 | opt.map(a => using(a)(f)) | |
| 25 | ||
| 78198 | 26 |   def using_optional[A <: AutoCloseable, B](opt: Option[A])(f: Option[A] => B): B = {
 | 
| 27 |     try { f(opt) }
 | |
| 28 |     finally {
 | |
| 29 |       opt match {
 | |
| 30 | case Some(a) if a != null => a.close() | |
| 31 | case _ => | |
| 32 | } | |
| 33 | } | |
| 34 | } | |
| 35 | ||
| 63789 | 36 | |
| 57909 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 37 | /* integers */ | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 38 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 39 | private val small_int = 10000 | 
| 75393 | 40 |   private lazy val small_int_table = {
 | 
| 57909 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 41 | val array = new Array[String](small_int) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 42 | for (i <- 0 until small_int) array(i) = i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 43 | array | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 44 | } | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 45 | |
| 75393 | 46 |   def is_small_int(s: String): Boolean = {
 | 
| 57909 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 47 | val len = s.length | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 48 | 1 <= len && len <= 4 && | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 49 | s.forall(c => '0' <= c && c <= '9') && | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 50 | (len == 1 || s(0) != '0') | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 51 | } | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 52 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 53 | def signed_string_of_long(i: Long): String = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 54 | if (0 <= i && i < small_int) small_int_table(i.toInt) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 55 | else i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 56 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 57 | def signed_string_of_int(i: Int): String = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 58 | if (0 <= i && i < small_int) small_int_table(i) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 59 | else i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 60 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 61 | |
| 48996 | 62 | /* separated chunks */ | 
| 36688 | 63 | |
| 75393 | 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: 
51616diff
changeset | 65 | val result = new mutable.ListBuffer[A] | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 66 | var first = true | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 67 |     for (x <- list) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 68 |       if (first) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 69 | first = false | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 70 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 71 | } | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 72 |       else {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 73 | result += s | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 74 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 75 | } | 
| 36688 | 76 | } | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 77 | result.toList | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 78 | } | 
| 36688 | 79 | |
| 56600 | 80 | def separated_chunks(sep: Char => Boolean, source: CharSequence): Iterator[CharSequence] = | 
| 48996 | 81 |     new Iterator[CharSequence] {
 | 
| 82 | private val end = source.length | |
| 75393 | 83 |       private def next_chunk(i: Int): Option[(CharSequence, Int)] = {
 | 
| 48996 | 84 |         if (i < end) {
 | 
| 75382 
81673c441ce3
tuned: eliminted do-while for the sake of scala3;
 wenzelm parents: 
75295diff
changeset | 85 | var j = i | 
| 75709 | 86 |           while ({
 | 
| 75382 
81673c441ce3
tuned: eliminted do-while for the sake of scala3;
 wenzelm parents: 
75295diff
changeset | 87 | j += 1 | 
| 75709 | 88 | j < end && !sep(source.charAt(j)) | 
| 89 | }) () | |
| 48996 | 90 | Some((source.subSequence(i + 1, j), j)) | 
| 91 | } | |
| 92 | else None | |
| 43598 | 93 | } | 
| 48996 | 94 | private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1) | 
| 95 | ||
| 73337 | 96 | def hasNext: Boolean = state.isDefined | 
| 48996 | 97 | def next(): CharSequence = | 
| 98 |         state match {
 | |
| 60215 | 99 | case Some((s, i)) => state = next_chunk(i); s | 
| 48996 | 100 | case None => Iterator.empty.next() | 
| 101 | } | |
| 43598 | 102 | } | 
| 103 | ||
| 48996 | 104 | def space_explode(sep: Char, str: String): List[String] = | 
| 56600 | 105 | separated_chunks(_ == sep, str).map(_.toString).toList | 
| 48996 | 106 | |
| 107 | ||
| 108 | /* lines */ | |
| 109 | ||
| 77007 | 110 | def count_newlines(str: String): Int = str.count(_ == '\n') | 
| 111 | ||
| 75859 
7164f537370f
proper treatment of empty lines (amending 08f89f0e8a62);
 wenzelm parents: 
75709diff
changeset | 112 |   def terminate_lines(lines: IterableOnce[String]): String = {
 | 
| 
7164f537370f
proper treatment of empty lines (amending 08f89f0e8a62);
 wenzelm parents: 
75709diff
changeset | 113 | val it = lines.iterator | 
| 
7164f537370f
proper treatment of empty lines (amending 08f89f0e8a62);
 wenzelm parents: 
75709diff
changeset | 114 |     if (it.isEmpty) "" else it.mkString("", "\n", "\n")
 | 
| 
7164f537370f
proper treatment of empty lines (amending 08f89f0e8a62);
 wenzelm parents: 
75709diff
changeset | 115 | } | 
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 116 | |
| 73362 | 117 | def cat_lines(lines: IterableOnce[String]): String = | 
| 118 |     lines.iterator.mkString("\n")
 | |
| 48996 | 119 | |
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 120 |   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 | 121 | |
| 62590 | 122 | def prefix_lines(prfx: String, str: String): String = | 
| 73963 | 123 | isabelle.setup.Library.prefix_lines(prfx, str) | 
| 62590 | 124 | |
| 73736 | 125 | def indent_lines(n: Int, str: String): String = | 
| 126 | prefix_lines(Symbol.spaces(n), str) | |
| 127 | ||
| 75393 | 128 |   def first_line(source: CharSequence): String = {
 | 
| 56600 | 129 | val lines = separated_chunks(_ == '\n', source) | 
| 73344 | 130 | if (lines.hasNext) lines.next().toString | 
| 48996 | 131 | else "" | 
| 132 | } | |
| 133 | ||
| 73332 | 134 | def trim_line(s: String): String = | 
| 73963 | 135 | isabelle.setup.Library.trim_line(s) | 
| 73332 | 136 | |
| 137 | def trim_split_lines(s: String): List[String] = | |
| 138 | split_lines(trim_line(s)).map(trim_line) | |
| 139 | ||
| 65932 | 140 |   def encode_lines(s: String): String = s.replace('\n', '\u000b')
 | 
| 141 |   def decode_lines(s: String): String = s.replace('\u000b', '\n')
 | |
| 142 | ||
| 50847 | 143 | |
| 144 | /* strings */ | |
| 145 | ||
| 75393 | 146 |   def make_string(f: StringBuilder => Unit, capacity: Int = 16): String = {
 | 
| 74794 
c606fddc5b05
slightly faster XML output: avoid too much regrowing of StringBuilder;
 wenzelm parents: 
73963diff
changeset | 147 | val s = new StringBuilder(capacity) | 
| 64355 | 148 | f(s) | 
| 149 | s.toString | |
| 150 | } | |
| 151 | ||
| 50847 | 152 | def try_unprefix(prfx: String, s: String): Option[String] = | 
| 153 | if (s.startsWith(prfx)) Some(s.substring(prfx.length)) else None | |
| 154 | ||
| 55033 | 155 | def try_unsuffix(sffx: String, s: String): Option[String] = | 
| 156 | if (s.endsWith(sffx)) Some(s.substring(0, s.length - sffx.length)) else None | |
| 157 | ||
| 65606 | 158 | def perhaps_unprefix(prfx: String, s: String): String = try_unprefix(prfx, s) getOrElse s | 
| 159 | def perhaps_unsuffix(sffx: String, s: String): String = try_unsuffix(sffx, s) getOrElse s | |
| 160 | ||
| 65903 | 161 | def isolate_substring(s: String): String = new String(s.toCharArray) | 
| 64820 
00488a8c042f
Line.Document consists of independently allocated strings;
 wenzelm parents: 
64370diff
changeset | 162 | |
| 71864 | 163 | def strip_ansi_color(s: String): String = | 
| 73355 | 164 |     s.replaceAll("\u001b\\[\\d+m", "")
 | 
| 71864 | 165 | |
| 43598 | 166 | |
| 48996 | 167 | /* quote */ | 
| 46196 | 168 | |
| 67820 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
67436diff
changeset | 169 | def single_quote(s: String): String = "'" + s + "'" | 
| 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
67436diff
changeset | 170 | |
| 43598 | 171 | def quote(s: String): String = "\"" + s + "\"" | 
| 56843 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 172 | |
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 173 | def try_unquote(s: String): Option[String] = | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 174 |     if (s.startsWith("\"") && s.endsWith("\"")) Some(s.substring(1, s.length - 1))
 | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 175 | else None | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 176 | |
| 58592 | 177 | def perhaps_unquote(s: String): String = try_unquote(s) getOrElse s | 
| 178 | ||
| 43598 | 179 |   def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | 
| 48362 | 180 |   def commas_quote(ss: Iterable[String]): String = ss.iterator.map(quote).mkString(", ")
 | 
| 43598 | 181 | |
| 36688 | 182 | |
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 183 | /* CharSequence */ | 
| 34141 | 184 | |
| 75393 | 185 |   class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence {
 | 
| 73120 
c3589f2dff31
more informative errors: simplify diagnosis of spurious failures reported by users;
 wenzelm parents: 
72214diff
changeset | 186 | require(0 <= start && start <= end && end <= text.length, "bad reverse range") | 
| 34141 | 187 | |
| 188 | def this(text: CharSequence) = this(text, 0, text.length) | |
| 189 | ||
| 190 | def length: Int = end - start | |
| 191 | def charAt(i: Int): Char = text.charAt(end - i - 1) | |
| 192 | ||
| 193 | def subSequence(i: Int, j: Int): CharSequence = | |
| 194 | if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i) | |
| 195 | else throw new IndexOutOfBoundsException | |
| 196 | ||
| 75393 | 197 |     override def toString: String = {
 | 
| 34141 | 198 | val buf = new StringBuilder(length) | 
| 199 | for (i <- 0 until length) | |
| 200 | buf.append(charAt(i)) | |
| 201 | buf.toString | |
| 202 | } | |
| 203 | } | |
| 204 | ||
| 75393 | 205 |   class Line_Termination(text: CharSequence) extends CharSequence {
 | 
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 206 | def length: Int = text.length + 1 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 207 | 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 | 208 | 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 | 209 | 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 | 210 | else text.subSequence(i, j) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 211 | 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 | 212 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 213 | |
| 34141 | 214 | |
| 59224 | 215 | /* regular expressions */ | 
| 216 | ||
| 217 | def make_regex(s: String): Option[Regex] = | |
| 218 |     try { Some(new Regex(s)) } catch { case ERROR(_) => None }
 | |
| 219 | ||
| 64871 | 220 |   def is_regex_meta(c: Char): Boolean = """()[]{}\^$|?*+.<>-=!""".contains(c)
 | 
| 221 | ||
| 222 | def escape_regex(s: String): String = | |
| 71601 | 223 |     if (s.exists(is_regex_meta)) {
 | 
| 64871 | 224 | (for (c <- s.iterator) | 
| 225 |        yield { if (is_regex_meta(c)) "\\" + c.toString else c.toString }).mkString
 | |
| 226 | } | |
| 227 | else s | |
| 228 | ||
| 59224 | 229 | |
| 61883 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 230 | /* lists */ | 
| 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 231 | |
| 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 232 | def take_prefix[A](pred: A => Boolean, xs: List[A]): (List[A], List[A]) = | 
| 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 233 | (xs.takeWhile(pred), xs.dropWhile(pred)) | 
| 56686 | 234 | |
| 75393 | 235 |   def take_suffix[A](pred: A => Boolean, xs: List[A]): (List[A], List[A]) = {
 | 
| 67434 | 236 | val rev_xs = xs.reverse | 
| 237 | (rev_xs.dropWhile(pred).reverse, rev_xs.takeWhile(pred).reverse) | |
| 238 | } | |
| 239 | ||
| 240 | def trim[A](pred: A => Boolean, xs: List[A]): List[A] = | |
| 241 | take_suffix(pred, take_prefix(pred, xs)._2)._1 | |
| 242 | ||
| 60215 | 243 | def member[A, B](xs: List[A])(x: B): Boolean = xs.contains(x) | 
| 56688 | 244 | def insert[A](x: A)(xs: List[A]): List[A] = if (xs.contains(x)) xs else x :: xs | 
| 245 | def remove[A, B](x: B)(xs: List[A]): List[A] = if (member(xs)(x)) xs.filterNot(_ == x) else xs | |
| 246 | def update[A](x: A)(xs: List[A]): List[A] = x :: remove(x)(xs) | |
| 63734 
133e3e84e6fb
some support for merge of Isabelle/jEdit shortcuts wrt. jEdit keymap;
 wenzelm parents: 
62590diff
changeset | 247 | |
| 63867 | 248 | def merge[A](xs: List[A], ys: List[A]): List[A] = | 
| 249 | if (xs.eq(ys)) xs | |
| 250 | else if (xs.isEmpty) ys | |
| 251 | else ys.foldRight(xs)(Library.insert(_)(_)) | |
| 252 | ||
| 75393 | 253 |   def distinct[A](xs: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] = {
 | 
| 63734 
133e3e84e6fb
some support for merge of Isabelle/jEdit shortcuts wrt. jEdit keymap;
 wenzelm parents: 
62590diff
changeset | 254 | val result = new mutable.ListBuffer[A] | 
| 64207 | 255 | xs.foreach(x => if (!result.exists(y => eq(x, y))) result += x) | 
| 63734 
133e3e84e6fb
some support for merge of Isabelle/jEdit shortcuts wrt. jEdit keymap;
 wenzelm parents: 
62590diff
changeset | 256 | result.toList | 
| 
133e3e84e6fb
some support for merge of Isabelle/jEdit shortcuts wrt. jEdit keymap;
 wenzelm parents: 
62590diff
changeset | 257 | } | 
| 63781 | 258 | |
| 75393 | 259 |   def duplicates[A](lst: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] = {
 | 
| 63781 | 260 | val result = new mutable.ListBuffer[A] | 
| 261 | @tailrec def dups(rest: List[A]): Unit = | |
| 262 |       rest match {
 | |
| 263 | case Nil => | |
| 264 | case x :: xs => | |
| 64207 | 265 | if (!result.exists(y => eq(x, y)) && xs.exists(y => eq(x, y))) result += x | 
| 63781 | 266 | dups(xs) | 
| 267 | } | |
| 268 | dups(lst) | |
| 269 | result.toList | |
| 270 | } | |
| 65761 | 271 | |
| 68715 | 272 | def replicate[A](n: Int, a: A): List[A] = | 
| 273 | if (n < 0) throw new IllegalArgumentException | |
| 274 | else if (n == 0) Nil | |
| 275 |     else {
 | |
| 276 | val res = new mutable.ListBuffer[A] | |
| 277 | (1 to n).foreach(_ => res += a) | |
| 278 | res.toList | |
| 279 | } | |
| 280 | ||
| 78943 | 281 | def the_single[A](xs: List[A], message: => String = "Single argument expected"): A = | 
| 73571 | 282 |     xs match {
 | 
| 283 | case List(x) => x | |
| 78943 | 284 | case _ => error(message) | 
| 73571 | 285 | } | 
| 286 | ||
| 65761 | 287 | |
| 79870 | 288 | /* named items */ | 
| 289 | ||
| 290 |   trait Named { def name: String }
 | |
| 291 | ||
| 292 | ||
| 79834 | 293 | /* data update */ | 
| 294 | ||
| 295 |   object Update {
 | |
| 296 | type Data[A] = Map[String, A] | |
| 297 | ||
| 79869 | 298 | sealed abstract class Op[A] | 
| 299 | case class Delete[A](name: String) extends Op[A] | |
| 300 | case class Insert[A](item: A) extends Op[A] | |
| 301 | ||
| 79870 | 302 | def data[A <: Named](old_data: Data[A], updates: List[Op[A]]): Data[A] = | 
| 303 |       updates.foldLeft(old_data) {
 | |
| 304 | case (map, Delete(name)) => map - name | |
| 305 | case (map, Insert(item)) => map + (item.name -> item) | |
| 306 | } | |
| 307 | ||
| 79834 | 308 | val empty: Update = Update() | 
| 309 | ||
| 79842 | 310 | def make[A](a: Data[A], b: Data[A], kind: Int = 0): Update = | 
| 311 | if (a eq b) empty | |
| 79834 | 312 |       else {
 | 
| 79842 | 313 | val delete = List.from(for ((x, y) <- a.iterator if !b.get(x).contains(y)) yield x) | 
| 314 | val insert = List.from(for ((x, y) <- b.iterator if !a.get(x).contains(y)) yield x) | |
| 79839 | 315 | Update(delete = delete, insert = insert, kind = kind) | 
| 79834 | 316 | } | 
| 317 | } | |
| 318 | ||
| 319 | sealed case class Update( | |
| 320 | delete: List[String] = Nil, | |
| 79839 | 321 | insert: List[String] = Nil, | 
| 322 | kind: Int = 0 | |
| 79834 | 323 |   ) {
 | 
| 324 | def deletes: Boolean = delete.nonEmpty | |
| 325 | def inserts: Boolean = insert.nonEmpty | |
| 326 | def defined: Boolean = deletes || inserts | |
| 79836 | 327 | lazy val domain: Set[String] = delete.toSet ++ insert | 
| 79834 | 328 | } | 
| 329 | ||
| 330 | ||
| 65761 | 331 | /* proper values */ | 
| 332 | ||
| 75295 
38398766be6b
command-line arguments for "isabelle vscode", similar to "isabelle jedit";
 wenzelm parents: 
74794diff
changeset | 333 | def proper_bool(b: Boolean): Option[Boolean] = | 
| 
38398766be6b
command-line arguments for "isabelle vscode", similar to "isabelle jedit";
 wenzelm parents: 
74794diff
changeset | 334 | if (!b) None else Some(b) | 
| 
38398766be6b
command-line arguments for "isabelle vscode", similar to "isabelle jedit";
 wenzelm parents: 
74794diff
changeset | 335 | |
| 65761 | 336 | def proper_string(s: String): Option[String] = | 
| 337 | if (s == null || s == "") None else Some(s) | |
| 338 | ||
| 339 | def proper_list[A](list: List[A]): Option[List[A]] = | |
| 340 | if (list == null || list.isEmpty) None else Some(list) | |
| 72214 | 341 | |
| 77367 | 342 | def if_proper[A](x: Iterable[A], body: => String): String = | 
| 343 | if (x == null || x.isEmpty) "" else body | |
| 344 | ||
| 77614 | 345 | def if_proper(b: Boolean, body: => String): String = | 
| 346 | if (!b) "" else body | |
| 347 | ||
| 72214 | 348 | |
| 349 | /* reflection */ | |
| 350 | ||
| 75393 | 351 |   def is_subclass[A, B](a: Class[A], b: Class[B]): Boolean = {
 | 
| 73339 | 352 | import scala.language.existentials | 
| 75393 | 353 |     @tailrec def subclass(c: Class[_]): Boolean = {
 | 
| 354 |       c == b || { val d = c.getSuperclass; d != null && subclass(d) }
 | |
| 72214 | 355 | } | 
| 356 | subclass(a) | |
| 357 | } | |
| 76788 | 358 | |
| 359 | def as_subclass[C](c: Class[C])(x: AnyRef): Option[C] = | |
| 360 | if (x == null || is_subclass(x.getClass, c)) Some(x.asInstanceOf[C]) else None | |
| 34136 | 361 | } |