src/Pure/Isar/runtime.ML
author wenzelm
Wed, 18 Sep 2013 13:18:51 +0200
changeset 53709 84522727f9d3
parent 52686 f4871fe80410
child 56265 785569927666
permissions -rw-r--r--
improved printing of exception trace in Poly/ML 5.5.1;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/runtime.ML
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     3
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     4
Isar toplevel runtime support.
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     5
*)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     6
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     7
signature RUNTIME =
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     8
sig
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
     9
  exception UNDEF
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    10
  exception TERMINATE
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    11
  exception EXCURSION_FAIL of exn * string
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    12
  exception TOPLEVEL_ERROR
39513
fce2202892c4 discontinued Output.debug, which belongs to early PGIP experiments (b6788dbd2ef9) and causes just too many problems (like spamming the message channel if it is used by more than one module);
wenzelm
parents: 39439
diff changeset
    13
  val debug: bool Unsynchronized.ref
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    14
  val exn_context: Proof.context option -> exn -> exn
51639
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
    15
  type exn_info = {exn_position: exn -> Position.T, pretty_exn: exn -> Pretty.T}
51285
0859bd338c9b tuned signature;
wenzelm
parents: 51242
diff changeset
    16
  type error = ((serial * string) * string option)
51639
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
    17
  val exn_messages_ids: exn_info -> exn -> error list
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
    18
  val exn_messages: exn_info -> exn -> (serial * string) list
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
    19
  val exn_message: exn_info -> exn -> string
53709
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 52686
diff changeset
    20
  val debugging: (exn -> string) -> ('a -> 'b) -> 'a -> 'b
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 52686
diff changeset
    21
  val controlled_execution: (exn -> string) -> ('a -> 'b) -> 'a -> 'b
33603
3713a5208671 generalized Runtime.toplevel_error wrt. output function;
wenzelm
parents: 32091
diff changeset
    22
  val toplevel_error: (exn -> unit) -> ('a -> 'b) -> 'a -> 'b
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    23
end;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    24
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    25
structure Runtime: RUNTIME =
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    26
struct
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    27
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    28
(** exceptions **)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    29
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    30
exception UNDEF;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    31
exception TERMINATE;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    32
exception EXCURSION_FAIL of exn * string;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    33
exception TOPLEVEL_ERROR;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    34
39513
fce2202892c4 discontinued Output.debug, which belongs to early PGIP experiments (b6788dbd2ef9) and causes just too many problems (like spamming the message channel if it is used by more than one module);
wenzelm
parents: 39439
diff changeset
    35
val debug = Unsynchronized.ref false;
fce2202892c4 discontinued Output.debug, which belongs to early PGIP experiments (b6788dbd2ef9) and causes just too many problems (like spamming the message channel if it is used by more than one module);
wenzelm
parents: 39439
diff changeset
    36
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    37
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    38
(* exn_context *)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    39
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    40
exception CONTEXT of Proof.context * exn;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    41
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    42
fun exn_context NONE exn = exn
39285
85728a4b5620 avoid extra wrapping for interrupts;
wenzelm
parents: 39238
diff changeset
    43
  | exn_context (SOME ctxt) exn = if Exn.is_interrupt exn then exn else CONTEXT (ctxt, exn);
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    44
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    45
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    46
(* exn_message *)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    47
51639
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
    48
type exn_info = {exn_position: exn -> Position.T, pretty_exn: exn -> Pretty.T};
51285
0859bd338c9b tuned signature;
wenzelm
parents: 51242
diff changeset
    49
type error = ((serial * string) * string option);
0859bd338c9b tuned signature;
wenzelm
parents: 51242
diff changeset
    50
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    51
local
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    52
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    53
fun robust f x =
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    54
  (case try f x of
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    55
    SOME s => s
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    56
  | NONE => Markup.markup Markup.intensify "<malformed>");
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    57
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    58
fun robust2 f x y = robust (fn () => f x y) ();
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    59
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    60
fun robust_context NONE _ _ = []
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    61
  | robust_context (SOME ctxt) f xs = map (robust2 f ctxt) xs;
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    62
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    63
fun identify exn =
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    64
  let
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    65
    val exn' = Par_Exn.identify [] exn;
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    66
    val exec_id = Properties.get (Exn_Properties.get exn') Markup.exec_idN;
51242
a8e664e4fb5f identify exceptions more robustly, to allow SML/NJ report toplevel errors without crash;
wenzelm
parents: 50914
diff changeset
    67
    val i = Par_Exn.the_serial exn' handle Option.Option => serial ();
a8e664e4fb5f identify exceptions more robustly, to allow SML/NJ report toplevel errors without crash;
wenzelm
parents: 50914
diff changeset
    68
  in ((i, exn'), exec_id) end;
50911
ee7fe4230642 more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
wenzelm
parents: 50505
diff changeset
    69
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    70
fun flatten _ (CONTEXT (ctxt, exn)) = flatten (SOME ctxt) exn
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    71
  | flatten context (Exn.EXCEPTIONS exns) = maps (flatten context) exns
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    72
  | flatten context exn =
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    73
      (case Par_Exn.dest exn of
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    74
        SOME exns => maps (flatten context) exns
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    75
      | NONE => [(context, identify exn)]);
50911
ee7fe4230642 more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
wenzelm
parents: 50505
diff changeset
    76
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    77
in
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    78
51639
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
    79
fun exn_messages_ids {exn_position, pretty_exn} e =
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    80
  let
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    81
    fun raised exn name msgs =
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 45666
diff changeset
    82
      let val pos = Position.here (exn_position exn) in
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    83
        (case msgs of
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    84
          [] => "exception " ^ name ^ " raised" ^ pos
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    85
        | [msg] => "exception " ^ name ^ " raised" ^ pos ^ ": " ^ msg
51653
97de25c51b2d tuned message;
wenzelm
parents: 51640
diff changeset
    86
        | _ =>
97de25c51b2d tuned message;
wenzelm
parents: 51640
diff changeset
    87
            cat_lines (("exception " ^ name ^ " raised" ^ pos ^ ":") ::
97de25c51b2d tuned message;
wenzelm
parents: 51640
diff changeset
    88
              map (Markup.markup Markup.item) msgs))
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    89
      end;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    90
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    91
    fun exn_msgs (context, ((i, exn), id)) =
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
    92
      (case exn of
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
    93
        EXCURSION_FAIL (exn, loc) =>
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    94
          map (fn ((i, msg), id) => ((i, msg ^ Markup.markup Markup.no_report ("\n" ^ loc)), id))
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
    95
            (sorted_msgs context exn)
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
    96
      | _ =>
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
    97
        let
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
    98
          val msg =
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
    99
            (case exn of
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   100
              TERMINATE => "Exit"
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   101
            | TimeLimit.TimeOut => "Timeout"
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   102
            | TOPLEVEL_ERROR => "Error"
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   103
            | ERROR msg => msg
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   104
            | Fail msg => raised exn "Fail" [msg]
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   105
            | THEORY (msg, thys) =>
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   106
                raised exn "THEORY" (msg :: map (robust Context.str_of_thy) thys)
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   107
            | Ast.AST (msg, asts) =>
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   108
                raised exn "AST" (msg :: map (robust (Pretty.string_of o Ast.pretty_ast)) asts)
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   109
            | TYPE (msg, Ts, ts) =>
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   110
                raised exn "TYPE" (msg ::
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   111
                  (robust_context context Syntax.string_of_typ Ts @
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   112
                    robust_context context Syntax.string_of_term ts))
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   113
            | TERM (msg, ts) =>
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   114
                raised exn "TERM" (msg :: robust_context context Syntax.string_of_term ts)
51640
d022e8bd2375 improved printing of exception CTERM (see also d0f0f37ec346);
wenzelm
parents: 51639
diff changeset
   115
            | CTERM (msg, cts) =>
d022e8bd2375 improved printing of exception CTERM (see also d0f0f37ec346);
wenzelm
parents: 51639
diff changeset
   116
                raised exn "CTERM"
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   117
                  (msg :: robust_context context Syntax.string_of_term (map term_of cts))
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   118
            | THM (msg, i, thms) =>
45487
ae60518ac054 simplified Runtime.exn_messages: always print detailed version of low-level exeptions, which should not occur in regular user errors anyway;
wenzelm
parents: 45486
diff changeset
   119
                raised exn ("THM " ^ string_of_int i)
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   120
                  (msg :: robust_context context Display.string_of_thm thms)
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   121
            | _ => raised exn (robust (Pretty.string_of o pretty_exn) exn) []);
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   122
        in [((i, msg), id)] end)
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   123
      and sorted_msgs context exn =
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   124
        sort_distinct (int_ord o pairself (fst o fst)) (maps exn_msgs (flatten context exn));
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   125
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   126
  in sorted_msgs NONE e end;
38874
4a4d34d2f97b more careful treatment of nested multi-exceptions, collapsing cumulatively empty list to Interrupt;
wenzelm
parents: 37370
diff changeset
   127
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   128
end;
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   129
51639
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
   130
fun exn_messages exn_info exn = map #1 (exn_messages_ids exn_info exn);
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   131
51639
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
   132
fun exn_message exn_info exn =
b7f908c99546 prefer pretty_exn where possible -- NB: low-level General.exnMessage may still be used elsewhere (e.g. by the ML compiler itself);
wenzelm
parents: 51285
diff changeset
   133
  (case exn_messages exn_info exn of
38875
c7a66b584147 tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents: 38874
diff changeset
   134
    [] => "Interrupt"
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   135
  | msgs => cat_lines (map snd msgs));
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   136
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   137
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   138
(** controlled execution **)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   139
53709
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 52686
diff changeset
   140
fun debugging exn_message f x =
45486
600682331b79 more detailed exception_trace: in Poly/ML 5.4.x intermediate handlers reset the trace (cf. Poly/ML SVN 1110);
wenzelm
parents: 45197
diff changeset
   141
  if ! debug
53709
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 52686
diff changeset
   142
  then print_exception_trace exn_message (fn () => f x)
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   143
  else f x;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   144
53709
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 52686
diff changeset
   145
fun controlled_execution exn_message f x =
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 52686
diff changeset
   146
  (f |> debugging exn_message |> Future.interruptible_task) x;
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   147
39238
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   148
fun toplevel_error output_exn f x = f x
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   149
  handle exn =>
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   150
    if Exn.is_interrupt exn then reraise exn
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   151
    else
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   152
      let
45197
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   153
        val opt_ctxt =
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   154
          (case Context.thread_data () of
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   155
            NONE => NONE
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   156
          | SOME context => try Context.proof_of context);
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   157
        val _ = output_exn (exn_context opt_ctxt exn);
39238
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   158
      in raise TOPLEVEL_ERROR end;
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   159
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   160
end;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   161