| author | wenzelm | 
| Wed, 21 Mar 2018 21:50:28 +0100 | |
| changeset 67917 | d13b2dd20f5e | 
| parent 65948 | de7888573ed7 | 
| child 70564 | 2c7c8be65b7d | 
| permissions | -rw-r--r-- | 
| 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 | 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: 
59582diff
changeset | 14 | val thread_context: exn -> exn | 
| 51285 | 15 | type error = ((serial * string) * string option) | 
| 65948 | 16 | val exn_messages: exn -> error list | 
| 17 | val exn_message_list: exn -> string list | |
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 18 | val exn_message: exn -> string | 
| 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
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: 
57831diff
changeset | 20 | val exn_system_message: exn -> unit | 
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
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: 
59055diff
changeset | 22 | val exn_trace_system: (unit -> 'a) -> 'a | 
| 62498 | 23 | val exn_debugger: (unit -> 'a) -> 'a | 
| 24 | val exn_debugger_system: (unit -> 'a) -> 'a | |
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
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: 
56279diff
changeset | 26 |   val controlled_execution: Context.generic option -> ('a -> 'b) -> 'a -> 'b
 | 
| 33603 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 wenzelm parents: 
32091diff
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: 
56279diff
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 | 41 | (* pretty *) | 
| 42 | ||
| 43 | fun pretty_exn (exn: exn) = | |
| 44 | Pretty.from_ML | |
| 45 | (ML_Pretty.from_polyml | |
| 62878 | 46 | (ML_system_pretty (exn, FixedInt.fromInt (ML_Print_Depth.get_print_depth ())))); | 
| 62516 | 47 | |
| 48 | ||
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 49 | (* exn_context *) | 
| 
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 | 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 | 52 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 53 | fun exn_context NONE exn = exn | 
| 39285 | 54 | | 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 | 55 | |
| 61077 
06cca32aa519
thread context for exceptions from forks, e.g. relevant when printing errors;
 wenzelm parents: 
59582diff
changeset | 56 | fun thread_context exn = | 
| 62889 | 57 | 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: 
59582diff
changeset | 58 | |
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 59 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 60 | (* exn_message *) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 61 | |
| 51285 | 62 | type error = ((serial * string) * string option); | 
| 63 | ||
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 64 | local | 
| 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 65 | |
| 52686 | 66 | fun robust f x = | 
| 67 | (case try f x of | |
| 68 | SOME s => s | |
| 69 | | NONE => Markup.markup Markup.intensify "<malformed>"); | |
| 70 | ||
| 71 | fun robust2 f x y = robust (fn () => f x y) (); | |
| 72 | ||
| 73 | fun robust_context NONE _ _ = [] | |
| 74 | | 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 | 75 | |
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 76 | 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 | 77 | let | 
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 78 | 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: 
50911diff
changeset | 79 | 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: 
50914diff
changeset | 80 | 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: 
50914diff
changeset | 81 | 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: 
50505diff
changeset | 82 | |
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 83 | 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: 
50911diff
changeset | 84 | | 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: 
50911diff
changeset | 85 | | flatten context exn = | 
| 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 86 | (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: 
50911diff
changeset | 87 | 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: 
50911diff
changeset | 88 | | NONE => [(context, identify exn)]); | 
| 50911 
ee7fe4230642
more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
 wenzelm parents: 
50505diff
changeset | 89 | |
| 61878 | 90 | val print_thy = Pretty.unformatted_string_of o Context.pretty_abbrev_thy; | 
| 91 | ||
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 92 | in | 
| 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 93 | |
| 65948 | 94 | fun exn_messages e = | 
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 95 | let | 
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 96 | fun raised exn name msgs = | 
| 62516 | 97 | 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 | 98 | (case msgs of | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 99 | [] => "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 | 100 | | [msg] => "exception " ^ name ^ " raised" ^ pos ^ ": " ^ msg | 
| 51653 | 101 | | _ => | 
| 102 |             cat_lines (("exception " ^ name ^ " raised" ^ pos ^ ":") ::
 | |
| 103 | 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 | 104 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 105 | |
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 106 | fun exn_msgs (context, ((i, exn), id)) = | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 107 | (case exn of | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 108 | EXCURSION_FAIL (exn, loc) => | 
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 109 |           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: 
44264diff
changeset | 110 | (sorted_msgs context exn) | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 111 | | _ => | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 112 | let | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 113 | val msg = | 
| 44247 | 114 | (case exn of | 
| 62519 | 115 | Timeout.TIMEOUT t => Timeout.print t | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 116 | | 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: 
57666diff
changeset | 117 | | ERROR "" => "Error" | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 118 | | ERROR msg => msg | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 119 | | Fail msg => raised exn "Fail" [msg] | 
| 61878 | 120 | | THEORY (msg, thys) => raised exn "THEORY" (msg :: map (robust print_thy) thys) | 
| 44247 | 121 | | Ast.AST (msg, asts) => | 
| 52686 | 122 | raised exn "AST" (msg :: map (robust (Pretty.string_of o Ast.pretty_ast)) asts) | 
| 44247 | 123 | | TYPE (msg, Ts, ts) => | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 124 | raised exn "TYPE" (msg :: | 
| 52686 | 125 | (robust_context context Syntax.string_of_typ Ts @ | 
| 126 | robust_context context Syntax.string_of_term ts)) | |
| 44247 | 127 | | TERM (msg, ts) => | 
| 52686 | 128 | raised exn "TERM" (msg :: robust_context context Syntax.string_of_term ts) | 
| 51640 
d022e8bd2375
improved printing of exception CTERM (see also d0f0f37ec346);
 wenzelm parents: 
51639diff
changeset | 129 | | CTERM (msg, cts) => | 
| 
d022e8bd2375
improved printing of exception CTERM (see also d0f0f37ec346);
 wenzelm parents: 
51639diff
changeset | 130 | raised exn "CTERM" | 
| 59582 | 131 | (msg :: robust_context context Syntax.string_of_term (map Thm.term_of cts)) | 
| 44247 | 132 | | 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: 
45486diff
changeset | 133 |                 raised exn ("THM " ^ string_of_int i)
 | 
| 61268 | 134 | (msg :: robust_context context Thm.string_of_thm thms) | 
| 62516 | 135 | | _ => 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: 
50911diff
changeset | 136 | in [((i, msg), id)] end) | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 137 | and sorted_msgs context exn = | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59056diff
changeset | 138 | 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: 
44264diff
changeset | 139 | |
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 140 | in sorted_msgs NONE e end; | 
| 38874 
4a4d34d2f97b
more careful treatment of nested multi-exceptions, collapsing cumulatively empty list to Interrupt;
 wenzelm parents: 
37370diff
changeset | 141 | |
| 50914 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 142 | end; | 
| 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 wenzelm parents: 
50911diff
changeset | 143 | |
| 65948 | 144 | fun exn_message_list exn = | 
| 145 | (case exn_messages exn of | |
| 146 | [] => ["Interrupt"] | |
| 147 | | 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: 
50911diff
changeset | 148 | |
| 65948 | 149 | 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 | 150 | |
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 151 | 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: 
57831diff
changeset | 152 | val exn_system_message = Output.system_message o exn_message; | 
| 62505 | 153 | fun exn_trace e = Exn.trace exn_message tracing e; | 
| 154 | 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: 
56279diff
changeset | 155 | |
| 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 156 | |
| 62498 | 157 | (* exception debugger *) | 
| 158 | ||
| 159 | local | |
| 160 | ||
| 161 | fun print_entry (name, loc) = | |
| 62821 | 162 | Markup.markup Markup.item (name ^ Position.here (Exn_Properties.position_of_polyml_location loc)); | 
| 62498 | 163 | |
| 164 | fun exception_debugger output e = | |
| 165 | let | |
| 166 | val (trace, result) = Exn_Debugger.capture_exception_trace e; | |
| 167 | val _ = | |
| 168 | (case (trace, result) of | |
| 169 | (_ :: _, Exn.Exn exn) => | |
| 170 |           output (cat_lines ("Exception trace - " ^ exn_message exn :: map print_entry trace))
 | |
| 171 | | _ => ()); | |
| 172 | in Exn.release result end; | |
| 173 | ||
| 174 | in | |
| 175 | ||
| 176 | fun exn_debugger e = exception_debugger tracing e; | |
| 177 | fun exn_debugger_system e = exception_debugger Output.system_message e; | |
| 178 | ||
| 179 | end; | |
| 180 | ||
| 181 | ||
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 182 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 183 | (** controlled execution **) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 184 | |
| 62498 | 185 | fun debugging1 opt_context f x = | 
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 186 | if ML_Options.exception_trace_enabled opt_context | 
| 59055 
5a7157b8e870
more informative failure of protocol commands, with exception trace;
 wenzelm parents: 
58894diff
changeset | 187 | 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 | 188 | |
| 62498 | 189 | fun debugging2 opt_context f x = | 
| 190 | if ML_Options.exception_debugger_enabled opt_context | |
| 191 | then exn_debugger (fn () => f x) else f x; | |
| 192 | ||
| 193 | fun debugging opt_context f = | |
| 194 | f |> debugging1 opt_context |> debugging2 opt_context; | |
| 195 | ||
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 196 | fun controlled_execution opt_context f x = | 
| 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 197 | (f |> debugging opt_context |> 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 | 198 | |
| 39238 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 199 | fun toplevel_error output_exn f x = f x | 
| 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 200 | handle exn => | 
| 62505 | 201 | if Exn.is_interrupt exn then Exn.reraise exn | 
| 39238 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 202 | else | 
| 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 203 | let | 
| 45197 
b6c527c64789
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
 wenzelm parents: 
44295diff
changeset | 204 | val opt_ctxt = | 
| 62889 | 205 | (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: 
44295diff
changeset | 206 | NONE => NONE | 
| 
b6c527c64789
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
 wenzelm parents: 
44295diff
changeset | 207 | | 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: 
44295diff
changeset | 208 | val _ = output_exn (exn_context opt_ctxt exn); | 
| 39238 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 209 | 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 | 210 | |
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 211 | fun toplevel_program body = | 
| 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 212 | (body |> controlled_execution NONE |> toplevel_error exn_error_message) (); | 
| 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56279diff
changeset | 213 | |
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 214 | end; |