src/Pure/Isar/runtime.ML
author wenzelm
Mon, 21 Jul 2025 16:21:37 +0200
changeset 82892 45107da819fc
parent 82092 93195437fdee
permissions -rw-r--r--
eliminate odd Unicode characters (amending e9f3b94eb6a0, b69e4da2604b, 8f0b2daa7eaa, 8d1e295aab70);
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 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
    11
  exception TOPLEVEL_ERROR
62516
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    12
  val pretty_exn: exn -> Pretty.T
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    13
  val exn_context: Proof.context option -> exn -> exn
61077
06cca32aa519 thread context for exceptions from forks, e.g. relevant when printing errors;
wenzelm
parents: 59582
diff changeset
    14
  val thread_context: exn -> exn
51285
0859bd338c9b tuned signature;
wenzelm
parents: 51242
diff changeset
    15
  type error = ((serial * string) * string option)
65948
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
    16
  val exn_messages: exn -> error list
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
    17
  val exn_message_list: exn -> string list
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
    18
  val exn_message: exn -> string
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
    19
  val exn_error_message: exn -> unit
57975
c657c68a60ab explicit system message for protocol failure -- show on Syslog panel instead of Raw Output;
wenzelm
parents: 57831
diff changeset
    20
  val exn_system_message: exn -> unit
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
    21
  val exn_trace: (unit -> 'a) -> 'a
59056
cbe9563c03d1 even more exception traces for Document.update, which goes through additional execution wrappers;
wenzelm
parents: 59055
diff changeset
    22
  val exn_trace_system: (unit -> 'a) -> 'a
62498
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
    23
  val exn_debugger: (unit -> 'a) -> 'a
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
    24
  val exn_debugger_system: (unit -> 'a) -> 'a
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
    25
  val debugging: Context.generic option -> ('a -> 'b) -> 'a -> 'b
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
    26
  val controlled_execution: Context.generic option -> ('a -> 'b) -> 'a -> 'b
33603
3713a5208671 generalized Runtime.toplevel_error wrt. output function;
wenzelm
parents: 32091
diff changeset
    27
  val toplevel_error: (exn -> unit) -> ('a -> 'b) -> 'a -> 'b
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
    28
  val toplevel_program: (unit -> 'a) -> 'a
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    29
end;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    30
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    31
structure Runtime: RUNTIME =
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    32
struct
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    33
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    34
(** exceptions **)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    35
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    36
exception UNDEF;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    37
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
    38
exception TOPLEVEL_ERROR;
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
62516
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    41
(* pretty *)
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    42
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    43
fun pretty_exn (exn: exn) =
80809
4a64fc4d1cde clarified signature: type ML_Pretty.pretty coincides with PolyML.pretty;
wenzelm
parents: 78759
diff changeset
    44
  Pretty.from_ML (ML_system_pretty (exn, FixedInt.fromInt (ML_Print_Depth.get_print_depth ())));
62516
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    45
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    46
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    47
(* exn_context *)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    48
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    49
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
    50
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    51
fun exn_context NONE exn = exn
39285
85728a4b5620 avoid extra wrapping for interrupts;
wenzelm
parents: 39238
diff changeset
    52
  | 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
    53
61077
06cca32aa519 thread context for exceptions from forks, e.g. relevant when printing errors;
wenzelm
parents: 59582
diff changeset
    54
fun thread_context exn =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
    55
  exn_context (Option.map Context.proof_of (Context.get_generic_context ())) exn;
61077
06cca32aa519 thread context for exceptions from forks, e.g. relevant when printing errors;
wenzelm
parents: 59582
diff changeset
    56
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    57
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    58
(* exn_message *)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    59
51285
0859bd338c9b tuned signature;
wenzelm
parents: 51242
diff changeset
    60
type error = ((serial * string) * string option);
0859bd338c9b tuned signature;
wenzelm
parents: 51242
diff changeset
    61
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    62
local
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    63
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    64
fun robust f x =
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    65
  (case try f x of
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    66
    SOME s => s
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    67
  | NONE => Markup.markup Markup.intensify "<malformed>");
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    68
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    69
fun robust2 f x y = robust (fn () => f x y) ();
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    70
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    71
fun robust_context NONE _ _ = []
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
    72
  | 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
    73
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    74
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
    75
  let
78679
dc7455787a8e clarified modules;
wenzelm
parents: 73841
diff changeset
    76
    val exn' = Exn_Properties.identify [] exn;
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    77
    val exec_id = Properties.get (Exn_Properties.get exn') Markup.exec_idN;
78679
dc7455787a8e clarified modules;
wenzelm
parents: 73841
diff changeset
    78
    val i = Exn_Properties.the_serial exn' handle Option.Option => serial ();
51242
a8e664e4fb5f identify exceptions more robustly, to allow SML/NJ report toplevel errors without crash;
wenzelm
parents: 50914
diff changeset
    79
  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
    80
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    81
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
    82
  | 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
    83
  | 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
    84
      (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
    85
        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
    86
      | 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
    87
61878
fa4dbb82732f tuned signature;
wenzelm
parents: 61268
diff changeset
    88
val print_thy = Pretty.unformatted_string_of o Context.pretty_abbrev_thy;
fa4dbb82732f tuned signature;
wenzelm
parents: 61268
diff changeset
    89
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    90
in
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    91
65948
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
    92
fun exn_messages e =
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
    93
  let
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    94
    fun raised exn name msgs =
62516
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    95
      let val pos = Position.here (Exn_Properties.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
    96
        (case msgs of
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
    97
          [] => "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
    98
        | [msg] => "exception " ^ name ^ " raised" ^ pos ^ ": " ^ msg
51653
97de25c51b2d tuned message;
wenzelm
parents: 51640
diff changeset
    99
        | _ =>
97de25c51b2d tuned message;
wenzelm
parents: 51640
diff changeset
   100
            cat_lines (("exception " ^ name ^ " raised" ^ pos ^ ":") ::
97de25c51b2d tuned message;
wenzelm
parents: 51640
diff changeset
   101
              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
   102
      end;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   103
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   104
    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
   105
      (case exn of
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   106
        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
   107
          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
   108
            (sorted_msgs context exn)
81957
adda8961df7b tuned output;
wenzelm
parents: 80809
diff changeset
   109
      | Morphism.MORPHISM (name, exn) =>
adda8961df7b tuned output;
wenzelm
parents: 80809
diff changeset
   110
          map (fn ((i, msg), id) => ((i, "MORPHISM " ^ quote name ^ "\n" ^ msg), id))
adda8961df7b tuned output;
wenzelm
parents: 80809
diff changeset
   111
            (sorted_msgs context exn)
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   112
      | _ =>
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   113
        let
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   114
          val msg =
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   115
            (case exn of
82092
93195437fdee clarified signature;
wenzelm
parents: 81957
diff changeset
   116
              Timeout.TIMEOUT t => Timeout.message t
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   117
            | TOPLEVEL_ERROR => "Error"
57831
885888a880fb more general notion of "user error" including empty message -- NB: Output.error_message needs non-empty string to emit anything;
wenzelm
parents: 57666
diff changeset
   118
            | ERROR "" => "Error"
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   119
            | ERROR msg => msg
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   120
            | Fail msg => raised exn "Fail" [msg]
61878
fa4dbb82732f tuned signature;
wenzelm
parents: 61268
diff changeset
   121
            | THEORY (msg, thys) => raised exn "THEORY" (msg :: map (robust print_thy) thys)
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   122
            | Ast.AST (msg, asts) =>
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   123
                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
   124
            | TYPE (msg, Ts, ts) =>
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   125
                raised exn "TYPE" (msg ::
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   126
                  (robust_context context Syntax.string_of_typ Ts @
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   127
                    robust_context context Syntax.string_of_term ts))
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   128
            | TERM (msg, ts) =>
52686
f4871fe80410 more robust exn_messages_ids;
wenzelm
parents: 51653
diff changeset
   129
                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
   130
            | CTERM (msg, cts) =>
d022e8bd2375 improved printing of exception CTERM (see also d0f0f37ec346);
wenzelm
parents: 51639
diff changeset
   131
                raised exn "CTERM"
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   132
                  (msg :: robust_context context Syntax.string_of_term (map Thm.term_of cts))
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents: 43761
diff changeset
   133
            | 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
   134
                raised exn ("THM " ^ string_of_int i)
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61077
diff changeset
   135
                  (msg :: robust_context context Thm.string_of_thm thms)
62516
5732f1c31566 tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
   136
            | _ => 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
   137
        in [((i, msg), id)] end)
44270
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   138
      and sorted_msgs context exn =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 59056
diff changeset
   139
        sort_distinct (int_ord o apply2 (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
   140
3eaad39e520c more careful treatment of exception serial numbers, with propagation to message channel;
wenzelm
parents: 44264
diff changeset
   141
  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
   142
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   143
end;
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   144
65948
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
   145
fun exn_message_list exn =
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
   146
  (case exn_messages exn of
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
   147
    [] => ["Interrupt"]
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
   148
  | msgs => map (#2 o #1) msgs);
50914
fe4714886d92 identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
wenzelm
parents: 50911
diff changeset
   149
65948
de7888573ed7 clarified build errors;
wenzelm
parents: 62889
diff changeset
   150
val exn_message = cat_lines o exn_message_list;
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   151
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   152
val exn_error_message = Output.error_message o exn_message;
57975
c657c68a60ab explicit system message for protocol failure -- show on Syslog panel instead of Raw Output;
wenzelm
parents: 57831
diff changeset
   153
val exn_system_message = Output.system_message o exn_message;
62505
9e2a65912111 clarified modules;
wenzelm
parents: 62498
diff changeset
   154
fun exn_trace e = Exn.trace exn_message tracing e;
9e2a65912111 clarified modules;
wenzelm
parents: 62498
diff changeset
   155
fun exn_trace_system e = Exn.trace exn_message Output.system_message e;
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   156
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   157
62498
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   158
(* exception debugger *)
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   159
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   160
local
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   161
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   162
fun print_entry (name, loc) =
62821
48c24d0b6d85 tuned signature;
wenzelm
parents: 62819
diff changeset
   163
  Markup.markup Markup.item (name ^ Position.here (Exn_Properties.position_of_polyml_location loc));
62498
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   164
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   165
fun exception_debugger output e =
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   166
  let
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   167
    val (trace, result) = Exn_Debugger.capture_exception_trace e;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   168
    val _ =
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   169
      (case (trace, result) of
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   170
        (_ :: _, Exn.Exn exn) =>
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   171
          output (cat_lines ("Exception trace - " ^ exn_message exn :: map print_entry trace))
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   172
      | _ => ());
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   173
  in Exn.release result end;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   174
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   175
in
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   176
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   177
fun exn_debugger e = exception_debugger tracing e;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   178
fun exn_debugger_system e = exception_debugger Output.system_message e;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   179
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   180
end;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   181
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   182
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   183
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   184
(** controlled execution **)
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   185
62498
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   186
fun debugging1 opt_context f x =
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   187
  if ML_Options.exception_trace_enabled opt_context
59055
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 58894
diff changeset
   188
  then exn_trace (fn () => f x) else f x;
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   189
62498
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   190
fun debugging2 opt_context f x =
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   191
  if ML_Options.exception_debugger_enabled opt_context
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   192
  then exn_debugger (fn () => f x) else f x;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   193
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   194
fun debugging opt_context f =
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   195
  f |> debugging1 opt_context |> debugging2 opt_context;
5dfcc9697f29 support for ML_exception_debugger;
wenzelm
parents: 61878
diff changeset
   196
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   197
fun controlled_execution opt_context f x =
73841
95484bd7e1ec proper profiling within command execution: messages require PIDE id;
wenzelm
parents: 70603
diff changeset
   198
  (f |> debugging opt_context |> Future.interruptible_task
95484bd7e1ec proper profiling within command execution: messages require PIDE id;
wenzelm
parents: 70603
diff changeset
   199
    |> ML_Profiling.profile (Options.default_string "profiling")) x;
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   200
78759
461e924cc825 proper Exn.capture / Isabelle_Thread.try_catch;
wenzelm
parents: 78679
diff changeset
   201
fun toplevel_error output_exn f x =
461e924cc825 proper Exn.capture / Isabelle_Thread.try_catch;
wenzelm
parents: 78679
diff changeset
   202
  Isabelle_Thread.try_catch (fn () => f x)
461e924cc825 proper Exn.capture / Isabelle_Thread.try_catch;
wenzelm
parents: 78679
diff changeset
   203
    (fn exn =>
39238
7189a138dd6c refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
wenzelm
parents: 39232
diff changeset
   204
      let
45197
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   205
        val opt_ctxt =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   206
          (case Context.get_generic_context () of
45197
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   207
            NONE => NONE
b6c527c64789 more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
wenzelm
parents: 44295
diff changeset
   208
          | 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
   209
        val _ = output_exn (exn_context opt_ctxt exn);
78759
461e924cc825 proper Exn.capture / Isabelle_Thread.try_catch;
wenzelm
parents: 78679
diff changeset
   210
      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
   211
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   212
fun toplevel_program body =
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   213
  (body |> controlled_execution NONE |> toplevel_error exn_error_message) ();
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56279
diff changeset
   214
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
diff changeset
   215
end;