author | wenzelm |
Tue, 19 Sep 2023 19:48:54 +0200 | |
changeset 78674 | 88f47c70187a |
parent 78428 | 48cbee2a6f2e |
child 78683 | cde40295ffd6 |
permissions | -rw-r--r-- |
62508
d0b68218ea55
discontinued RAW session: bootstrap directly from isabelle_process RAW_ML_SYSTEM;
wenzelm
parents:
62492
diff
changeset
|
1 |
/* Title: Pure/General/exn.scala |
34136 | 2 |
Author: Makarius |
3 |
||
43761
e72ba84ae58f
tuned signature -- corresponding to Scala version;
wenzelm
parents:
34313
diff
changeset
|
4 |
Support for exceptions (arbitrary throwables). |
34136 | 5 |
*/ |
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
75393 | 10 |
object Exn { |
62492 | 11 |
/* user errors */ |
12 |
||
75393 | 13 |
class User_Error(message: String) extends RuntimeException(message) { |
62492 | 14 |
override def equals(that: Any): Boolean = |
15 |
that match { |
|
16 |
case other: User_Error => message == other.getMessage |
|
17 |
case _ => false |
|
18 |
} |
|
19 |
override def hashCode: Int = message.hashCode |
|
20 |
||
65828 | 21 |
override def toString: String = "\n" + Output.error_message_text(message) |
62492 | 22 |
} |
23 |
||
75393 | 24 |
object ERROR { |
62492 | 25 |
def apply(message: String): User_Error = new User_Error(message) |
26 |
def unapply(exn: Throwable): Option[String] = user_message(exn) |
|
27 |
} |
|
28 |
||
29 |
def error(message: String): Nothing = throw ERROR(message) |
|
30 |
||
68131 | 31 |
def cat_message(msgs: String*): String = |
32 |
cat_lines(msgs.iterator.filterNot(_ == "")) |
|
62492 | 33 |
|
68131 | 34 |
def cat_error(msgs: String*): Nothing = |
35 |
error(cat_message(msgs:_*)) |
|
62492 | 36 |
|
37 |
||
44158 | 38 |
/* exceptions as values */ |
34136 | 39 |
|
75393 | 40 |
sealed abstract class Result[A] { |
59698
d4ce901f20c5
value-oriented user error, for well-defined Thy_Syntax.chop_common;
wenzelm
parents:
59697
diff
changeset
|
41 |
def user_error: Result[A] = |
d4ce901f20c5
value-oriented user error, for well-defined Thy_Syntax.chop_common;
wenzelm
parents:
59697
diff
changeset
|
42 |
this match { |
d4ce901f20c5
value-oriented user error, for well-defined Thy_Syntax.chop_common;
wenzelm
parents:
59697
diff
changeset
|
43 |
case Exn(ERROR(msg)) => Exn(ERROR(msg)) |
d4ce901f20c5
value-oriented user error, for well-defined Thy_Syntax.chop_common;
wenzelm
parents:
59697
diff
changeset
|
44 |
case _ => this |
d4ce901f20c5
value-oriented user error, for well-defined Thy_Syntax.chop_common;
wenzelm
parents:
59697
diff
changeset
|
45 |
} |
71142 | 46 |
|
47 |
def map[B](f: A => B): Result[B] = |
|
48 |
this match { |
|
49 |
case Res(res) => Res(f(res)) |
|
50 |
case Exn(exn) => Exn(exn) |
|
51 |
} |
|
59698
d4ce901f20c5
value-oriented user error, for well-defined Thy_Syntax.chop_common;
wenzelm
parents:
59697
diff
changeset
|
52 |
} |
43762 | 53 |
case class Res[A](res: A) extends Result[A] |
54 |
case class Exn[A](exn: Throwable) extends Result[A] |
|
34136 | 55 |
|
78674 | 56 |
def is_res[A](result: Result[A]): Boolean = result.isInstanceOf[Res[A]] |
57 |
def is_exn[A](result: Result[A]): Boolean = result.isInstanceOf[Exn[A]] |
|
58 |
||
78428 | 59 |
def the_res[A]: PartialFunction[Result[A], A] = { case Res(res) => res } |
60 |
def the_exn[A]: PartialFunction[Result[A], Throwable] = { case Exn(exn) => exn } |
|
61 |
||
34136 | 62 |
def capture[A](e: => A): Result[A] = |
63 |
try { Res(e) } |
|
34313
2f890016afab
treat *all* JVM throwables as "exceptions", cf. ML version;
wenzelm
parents:
34300
diff
changeset
|
64 |
catch { case exn: Throwable => Exn[A](exn) } |
34136 | 65 |
|
66 |
def release[A](result: Result[A]): A = |
|
67 |
result match { |
|
68 |
case Res(x) => x |
|
69 |
case Exn(exn) => throw exn |
|
70 |
} |
|
44158 | 71 |
|
59136
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
wenzelm
parents:
57831
diff
changeset
|
72 |
def release_first[A](results: List[Result[A]]): List[A] = |
59138 | 73 |
results.find({ case Exn(exn) => !is_interrupt(exn) case _ => false }) match { |
74 |
case Some(Exn(exn)) => throw exn |
|
71601 | 75 |
case _ => results.map(release) |
59138 | 76 |
} |
59136
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
wenzelm
parents:
57831
diff
changeset
|
77 |
|
44158 | 78 |
|
56667
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
79 |
/* interrupts */ |
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
80 |
|
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
81 |
def is_interrupt(exn: Throwable): Boolean = |
73963 | 82 |
isabelle.setup.Exn.is_interrupt(exn) |
56667
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
83 |
|
74068 | 84 |
def failure_rc(exn: Throwable): Int = |
85 |
isabelle.setup.Exn.failure_rc(exn) |
|
86 |
||
76568 | 87 |
def is_interrupt_exn[A](result: Result[A]): Boolean = |
88 |
result match { |
|
89 |
case Exn(exn) if is_interrupt(exn) => true |
|
90 |
case _ => false |
|
91 |
} |
|
92 |
||
68315 | 93 |
def interruptible_capture[A](e: => A): Result[A] = |
94 |
try { Res(e) } |
|
95 |
catch { case exn: Throwable if !is_interrupt(exn) => Exn[A](exn) } |
|
96 |
||
75393 | 97 |
object Interrupt { |
98 |
object ERROR { |
|
72596 | 99 |
def unapply(exn: Throwable): Option[String] = |
100 |
if (is_interrupt(exn)) Some(message(exn)) else user_message(exn) |
|
101 |
} |
|
102 |
||
72335 | 103 |
def apply(): Throwable = new InterruptedException("Interrupt") |
56667
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
104 |
def unapply(exn: Throwable): Boolean = is_interrupt(exn) |
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
105 |
|
73367 | 106 |
def dispose(): Unit = Thread.interrupted() |
107 |
def expose(): Unit = if (Thread.interrupted()) throw apply() |
|
108 |
def impose(): Unit = Thread.currentThread.interrupt() |
|
56667
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
109 |
} |
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
110 |
|
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
111 |
|
44158 | 112 |
/* message */ |
113 |
||
48479
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
wenzelm
parents:
48476
diff
changeset
|
114 |
def user_message(exn: Throwable): Option[String] = |
75393 | 115 |
if (exn.isInstanceOf[User_Error] || exn.getClass == classOf[RuntimeException]) { |
65716 | 116 |
Some(proper_string(exn.getMessage) getOrElse "Error") |
44158 | 117 |
} |
75393 | 118 |
else if (exn.isInstanceOf[java.sql.SQLException]) { |
65716 | 119 |
Some(proper_string(exn.getMessage) getOrElse "SQL error") |
63782 | 120 |
} |
75393 | 121 |
else if (exn.isInstanceOf[java.io.IOException]) { |
59697
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
wenzelm
parents:
59138
diff
changeset
|
122 |
val msg = exn.getMessage |
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
wenzelm
parents:
59138
diff
changeset
|
123 |
Some(if (msg == null || msg == "") "I/O error" else "I/O error: " + msg) |
43e14b0e2ef8
more explicit exception User_Error, with value-oriented equality;
wenzelm
parents:
59138
diff
changeset
|
124 |
} |
48479
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
wenzelm
parents:
48476
diff
changeset
|
125 |
else if (exn.isInstanceOf[RuntimeException]) Some(exn.toString) |
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
wenzelm
parents:
48476
diff
changeset
|
126 |
else None |
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
wenzelm
parents:
48476
diff
changeset
|
127 |
|
819f7a5f3e7f
more general notion of user ERROR (cf. 44f56fe01528);
wenzelm
parents:
48476
diff
changeset
|
128 |
def message(exn: Throwable): String = |
56667
65e84b0ef974
more abstract Exn.Interrupt and POSIX return code;
wenzelm
parents:
52066
diff
changeset
|
129 |
user_message(exn) getOrElse (if (is_interrupt(exn)) "Interrupt" else exn.toString) |
68806 | 130 |
|
131 |
||
73963 | 132 |
/* print */ |
133 |
||
134 |
def debug(): Boolean = isabelle.setup.Exn.debug() |
|
68806 | 135 |
|
73963 | 136 |
def trace(exn: Throwable): String = isabelle.setup.Exn.trace(exn) |
137 |
||
138 |
def print(exn: Throwable): String = |
|
139 |
if (debug()) message(exn) + "\n" + trace(exn) else message(exn) |
|
34136 | 140 |
} |