| author | wenzelm | 
| Thu, 12 Nov 2020 11:46:53 +0100 | |
| changeset 72595 | c806eeb9138c | 
| parent 72214 | 5924c1da3c45 | 
| child 73120 | c3589f2dff31 | 
| 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 | |
| 34136 | 15 | object Library | 
| 16 | {
 | |
| 63789 | 17 | /* resource management */ | 
| 18 | ||
| 69393 
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
 wenzelm parents: 
68715diff
changeset | 19 | def using[A <: AutoCloseable, B](a: A)(f: A => B): B = | 
| 63789 | 20 |   {
 | 
| 69393 
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
 wenzelm parents: 
68715diff
changeset | 21 |     try { f(a) }
 | 
| 
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
 wenzelm parents: 
68715diff
changeset | 22 |     finally { if (a != null) a.close() }
 | 
| 63789 | 23 | } | 
| 24 | ||
| 25 | ||
| 57909 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 26 | /* integers */ | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 27 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 28 | private val small_int = 10000 | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 29 | private lazy val small_int_table = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 30 |   {
 | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 31 | val array = new Array[String](small_int) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 32 | for (i <- 0 until small_int) array(i) = i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 33 | array | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 34 | } | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 35 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 36 | def is_small_int(s: String): Boolean = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 37 |   {
 | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 38 | val len = s.length | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 39 | 1 <= len && len <= 4 && | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 40 | s.forall(c => '0' <= c && c <= '9') && | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 41 | (len == 1 || s(0) != '0') | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 42 | } | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 43 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 44 | def signed_string_of_long(i: Long): String = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 45 | if (0 <= i && i < small_int) small_int_table(i.toInt) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 46 | else i.toString | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 47 | |
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 48 | def signed_string_of_int(i: Int): String = | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 49 | if (0 <= i && i < small_int) small_int_table(i) | 
| 
0fb331032f02
more compact representation of special string values;
 wenzelm parents: 
57831diff
changeset | 50 | else i.toString | 
| 
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 | |
| 48996 | 53 | /* separated chunks */ | 
| 36688 | 54 | |
| 55 | 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 | 56 |   {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 57 | val result = new mutable.ListBuffer[A] | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 58 | var first = true | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 59 |     for (x <- list) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 60 |       if (first) {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 61 | first = false | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 62 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 63 | } | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 64 |       else {
 | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 65 | result += s | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 66 | result += x | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 67 | } | 
| 36688 | 68 | } | 
| 51981 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 69 | result.toList | 
| 
a8ffd3692f57
more scalable Library.separate -- NB: JVM has tiny fixed-size stack;
 wenzelm parents: 
51616diff
changeset | 70 | } | 
| 36688 | 71 | |
| 56600 | 72 | def separated_chunks(sep: Char => Boolean, source: CharSequence): Iterator[CharSequence] = | 
| 48996 | 73 |     new Iterator[CharSequence] {
 | 
| 74 | private val end = source.length | |
| 75 | private def next_chunk(i: Int): Option[(CharSequence, Int)] = | |
| 76 |       {
 | |
| 77 |         if (i < end) {
 | |
| 56600 | 78 | var j = i; do j += 1 while (j < end && !sep(source.charAt(j))) | 
| 48996 | 79 | Some((source.subSequence(i + 1, j), j)) | 
| 80 | } | |
| 81 | else None | |
| 43598 | 82 | } | 
| 48996 | 83 | private var state: Option[(CharSequence, Int)] = if (end == 0) None else next_chunk(-1) | 
| 84 | ||
| 85 | def hasNext(): Boolean = state.isDefined | |
| 86 | def next(): CharSequence = | |
| 87 |         state match {
 | |
| 60215 | 88 | case Some((s, i)) => state = next_chunk(i); s | 
| 48996 | 89 | case None => Iterator.empty.next() | 
| 90 | } | |
| 43598 | 91 | } | 
| 92 | ||
| 48996 | 93 | def space_explode(sep: Char, str: String): List[String] = | 
| 56600 | 94 | separated_chunks(_ == sep, str).map(_.toString).toList | 
| 48996 | 95 | |
| 96 | ||
| 97 | /* lines */ | |
| 98 | ||
| 64002 | 99 |   def terminate_lines(lines: TraversableOnce[String]): String = lines.mkString("", "\n", "\n")
 | 
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 100 | |
| 48996 | 101 |   def cat_lines(lines: TraversableOnce[String]): String = lines.mkString("\n")
 | 
| 102 | ||
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43652diff
changeset | 103 |   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 | 104 | |
| 62590 | 105 | def prefix_lines(prfx: String, str: String): String = | 
| 71164 | 106 | if (str == "") str else cat_lines(split_lines(str).map(prfx + _)) | 
| 62590 | 107 | |
| 48996 | 108 | def first_line(source: CharSequence): String = | 
| 109 |   {
 | |
| 56600 | 110 | val lines = separated_chunks(_ == '\n', source) | 
| 48996 | 111 | if (lines.hasNext) lines.next.toString | 
| 112 | else "" | |
| 113 | } | |
| 114 | ||
| 65932 | 115 |   def encode_lines(s: String): String = s.replace('\n', '\u000b')
 | 
| 116 |   def decode_lines(s: String): String = s.replace('\u000b', '\n')
 | |
| 117 | ||
| 50847 | 118 | |
| 119 | /* strings */ | |
| 120 | ||
| 64355 | 121 | def make_string(f: StringBuilder => Unit): String = | 
| 122 |   {
 | |
| 123 | val s = new StringBuilder | |
| 124 | f(s) | |
| 125 | s.toString | |
| 126 | } | |
| 127 | ||
| 50847 | 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 | ||
| 65606 | 134 | def perhaps_unprefix(prfx: String, s: String): String = try_unprefix(prfx, s) getOrElse s | 
| 135 | def perhaps_unsuffix(sffx: String, s: String): String = try_unsuffix(sffx, s) getOrElse s | |
| 136 | ||
| 52444 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 137 | def trim_line(s: String): String = | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 138 |     if (s.endsWith("\r\n")) s.substring(0, s.length - 2)
 | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 139 |     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: 
51983diff
changeset | 140 | else s | 
| 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
51983diff
changeset | 141 | |
| 64063 | 142 | def trim_split_lines(s: String): List[String] = | 
| 71601 | 143 | split_lines(trim_line(s)).map(trim_line) | 
| 64063 | 144 | |
| 65903 | 145 | def isolate_substring(s: String): String = new String(s.toCharArray) | 
| 64820 
00488a8c042f
Line.Document consists of independently allocated strings;
 wenzelm parents: 
64370diff
changeset | 146 | |
| 71864 | 147 | def strip_ansi_color(s: String): String = | 
| 148 |     s.replaceAll("""\u001b\[\d+m""", "")
 | |
| 149 | ||
| 43598 | 150 | |
| 48996 | 151 | /* quote */ | 
| 46196 | 152 | |
| 67820 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
67436diff
changeset | 153 | def single_quote(s: String): String = "'" + s + "'" | 
| 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
67436diff
changeset | 154 | |
| 43598 | 155 | def quote(s: String): String = "\"" + s + "\"" | 
| 56843 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 156 | |
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 157 | def try_unquote(s: String): Option[String] = | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 158 |     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 | 159 | else None | 
| 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56730diff
changeset | 160 | |
| 58592 | 161 | def perhaps_unquote(s: String): String = try_unquote(s) getOrElse s | 
| 162 | ||
| 43598 | 163 |   def commas(ss: Iterable[String]): String = ss.iterator.mkString(", ")
 | 
| 48362 | 164 |   def commas_quote(ss: Iterable[String]): String = ss.iterator.map(quote).mkString(", ")
 | 
| 43598 | 165 | |
| 36688 | 166 | |
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 167 | /* CharSequence */ | 
| 34141 | 168 | |
| 169 | class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence | |
| 170 |   {
 | |
| 171 | require(0 <= start && start <= end && end <= text.length) | |
| 172 | ||
| 173 | def this(text: CharSequence) = this(text, 0, text.length) | |
| 174 | ||
| 175 | def length: Int = end - start | |
| 176 | def charAt(i: Int): Char = text.charAt(end - i - 1) | |
| 177 | ||
| 178 | def subSequence(i: Int, j: Int): CharSequence = | |
| 179 | if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i) | |
| 180 | else throw new IndexOutOfBoundsException | |
| 181 | ||
| 182 | override def toString: String = | |
| 183 |     {
 | |
| 184 | val buf = new StringBuilder(length) | |
| 185 | for (i <- 0 until length) | |
| 186 | buf.append(charAt(i)) | |
| 187 | buf.toString | |
| 188 | } | |
| 189 | } | |
| 190 | ||
| 51983 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 191 | class Line_Termination(text: CharSequence) extends CharSequence | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 192 |   {
 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 193 | def length: Int = text.length + 1 | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 194 | 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 | 195 | 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 | 196 | 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 | 197 | else text.subSequence(i, j) | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 198 | 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 | 199 | } | 
| 
32692ce4c61a
more frugal line termination, to cope with huge log files (see also 016cb7d8f297);
 wenzelm parents: 
51981diff
changeset | 200 | |
| 34141 | 201 | |
| 59224 | 202 | /* regular expressions */ | 
| 203 | ||
| 204 | def make_regex(s: String): Option[Regex] = | |
| 205 |     try { Some(new Regex(s)) } catch { case ERROR(_) => None }
 | |
| 206 | ||
| 64871 | 207 |   def is_regex_meta(c: Char): Boolean = """()[]{}\^$|?*+.<>-=!""".contains(c)
 | 
| 208 | ||
| 209 | def escape_regex(s: String): String = | |
| 71601 | 210 |     if (s.exists(is_regex_meta)) {
 | 
| 64871 | 211 | (for (c <- s.iterator) | 
| 212 |        yield { if (is_regex_meta(c)) "\\" + c.toString else c.toString }).mkString
 | |
| 213 | } | |
| 214 | else s | |
| 215 | ||
| 59224 | 216 | |
| 61883 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 217 | /* lists */ | 
| 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 218 | |
| 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 wenzelm parents: 
60215diff
changeset | 219 | 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 | 220 | (xs.takeWhile(pred), xs.dropWhile(pred)) | 
| 56686 | 221 | |
| 67434 | 222 | def take_suffix[A](pred: A => Boolean, xs: List[A]): (List[A], List[A]) = | 
| 223 |   {
 | |
| 224 | val rev_xs = xs.reverse | |
| 225 | (rev_xs.dropWhile(pred).reverse, rev_xs.takeWhile(pred).reverse) | |
| 226 | } | |
| 227 | ||
| 228 | def trim[A](pred: A => Boolean, xs: List[A]): List[A] = | |
| 229 | take_suffix(pred, take_prefix(pred, xs)._2)._1 | |
| 230 | ||
| 60215 | 231 | def member[A, B](xs: List[A])(x: B): Boolean = xs.contains(x) | 
| 56688 | 232 | def insert[A](x: A)(xs: List[A]): List[A] = if (xs.contains(x)) xs else x :: xs | 
| 233 | def remove[A, B](x: B)(xs: List[A]): List[A] = if (member(xs)(x)) xs.filterNot(_ == x) else xs | |
| 234 | 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 | 235 | |
| 63867 | 236 | def merge[A](xs: List[A], ys: List[A]): List[A] = | 
| 237 | if (xs.eq(ys)) xs | |
| 238 | else if (xs.isEmpty) ys | |
| 239 | else ys.foldRight(xs)(Library.insert(_)(_)) | |
| 240 | ||
| 64207 | 241 | 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 | 242 |   {
 | 
| 
133e3e84e6fb
some support for merge of Isabelle/jEdit shortcuts wrt. jEdit keymap;
 wenzelm parents: 
62590diff
changeset | 243 | val result = new mutable.ListBuffer[A] | 
| 64207 | 244 | 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 | 245 | result.toList | 
| 
133e3e84e6fb
some support for merge of Isabelle/jEdit shortcuts wrt. jEdit keymap;
 wenzelm parents: 
62590diff
changeset | 246 | } | 
| 63781 | 247 | |
| 64207 | 248 | def duplicates[A](lst: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] = | 
| 63781 | 249 |   {
 | 
| 250 | val result = new mutable.ListBuffer[A] | |
| 251 | @tailrec def dups(rest: List[A]): Unit = | |
| 252 |       rest match {
 | |
| 253 | case Nil => | |
| 254 | case x :: xs => | |
| 64207 | 255 | if (!result.exists(y => eq(x, y)) && xs.exists(y => eq(x, y))) result += x | 
| 63781 | 256 | dups(xs) | 
| 257 | } | |
| 258 | dups(lst) | |
| 259 | result.toList | |
| 260 | } | |
| 65761 | 261 | |
| 68715 | 262 | def replicate[A](n: Int, a: A): List[A] = | 
| 263 | if (n < 0) throw new IllegalArgumentException | |
| 264 | else if (n == 0) Nil | |
| 265 |     else {
 | |
| 266 | val res = new mutable.ListBuffer[A] | |
| 267 | (1 to n).foreach(_ => res += a) | |
| 268 | res.toList | |
| 269 | } | |
| 270 | ||
| 65761 | 271 | |
| 272 | /* proper values */ | |
| 273 | ||
| 274 | def proper_string(s: String): Option[String] = | |
| 275 | if (s == null || s == "") None else Some(s) | |
| 276 | ||
| 277 | def proper_list[A](list: List[A]): Option[List[A]] = | |
| 278 | if (list == null || list.isEmpty) None else Some(list) | |
| 72214 | 279 | |
| 280 | ||
| 281 | /* reflection */ | |
| 282 | ||
| 283 | def is_subclass[A, B](a: Class[A], b: Class[B]): Boolean = | |
| 284 |   {
 | |
| 285 | @tailrec def subclass(c: Class[_]): Boolean = | |
| 286 |     {
 | |
| 287 | c == b || | |
| 288 |         {
 | |
| 289 | val d = c.getSuperclass | |
| 290 | d != null && subclass(d) | |
| 291 | } | |
| 292 | } | |
| 293 | subclass(a) | |
| 294 | } | |
| 34136 | 295 | } |