| author | wenzelm | 
| Fri, 27 Jul 2012 14:15:04 +0200 | |
| changeset 48549 | cc7990d6eb38 | 
| parent 45666 | d83797ef0d2d | 
| child 48992 | 0518bf89c777 | 
| 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 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: 
39439diff
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 | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 15 | val exn_messages: (exn -> Position.T) -> exn -> (serial * string) list | 
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 16 | val exn_message: (exn -> Position.T) -> exn -> string | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 17 |   val debugging: ('a -> 'b) -> 'a -> 'b
 | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 18 |   val controlled_execution: ('a -> 'b) -> 'a -> 'b
 | 
| 33603 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 wenzelm parents: 
32091diff
changeset | 19 |   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 | 20 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 21 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 22 | structure Runtime: RUNTIME = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 23 | struct | 
| 
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 | (** exceptions **) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 26 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 27 | exception UNDEF; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 28 | exception TERMINATE; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 29 | 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 | 30 | exception TOPLEVEL_ERROR; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 31 | |
| 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: 
39439diff
changeset | 32 | 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: 
39439diff
changeset | 33 | |
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 34 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 35 | (* exn_context *) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 36 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 37 | 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 | 38 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 39 | fun exn_context NONE exn = exn | 
| 39285 | 40 | | 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 | 41 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 42 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 43 | (* exn_message *) | 
| 
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 | fun if_context NONE _ _ = [] | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 46 | | if_context (SOME ctxt) f xs = map (f ctxt) xs; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 47 | |
| 38874 
4a4d34d2f97b
more careful treatment of nested multi-exceptions, collapsing cumulatively empty list to Interrupt;
 wenzelm parents: 
37370diff
changeset | 48 | fun exn_messages exn_position e = | 
| 31476 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 49 | let | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 50 | fun raised exn name msgs = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 51 | let val pos = Position.str_of (exn_position exn) in | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 52 | (case msgs of | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 53 | [] => "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 | 54 | | [msg] => "exception " ^ name ^ " raised" ^ pos ^ ": " ^ msg | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 55 |         | _ => cat_lines (("exception " ^ name ^ " raised" ^ pos ^ ":") :: msgs))
 | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 56 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 57 | |
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 58 | fun flatten _ (CONTEXT (ctxt, exn)) = flatten (SOME ctxt) exn | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 59 | | flatten context (Exn.EXCEPTIONS exns) = maps (flatten context) exns | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 60 | | flatten context exn = | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 61 | (case Par_Exn.dest exn of | 
| 44271 
89f40a5939b2
more precise treatment of exception nesting and serial numbers;
 wenzelm parents: 
44270diff
changeset | 62 | SOME exns => maps (flatten context) exns | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 63 | | NONE => [(context, Par_Exn.serial exn)]); | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 64 | |
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 65 | fun exn_msgs (context, (i, exn)) = | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 66 | (case exn of | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 67 | EXCURSION_FAIL (exn, loc) => | 
| 45666 | 68 |           map (apsnd (fn msg => msg ^ Markup.markup Isabelle_Markup.no_report ("\n" ^ loc)))
 | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 69 | (sorted_msgs context exn) | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 70 | | _ => | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 71 | let | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 72 | val msg = | 
| 44247 | 73 | (case exn of | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 74 | TERMINATE => "Exit" | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 75 | | TimeLimit.TimeOut => "Timeout" | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 76 | | TOPLEVEL_ERROR => "Error" | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 77 | | ERROR msg => msg | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 78 | | Fail msg => raised exn "Fail" [msg] | 
| 44247 | 79 | | THEORY (msg, thys) => | 
| 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 | 80 | raised exn "THEORY" (msg :: map Context.str_of_thy thys) | 
| 44247 | 81 | | Ast.AST (msg, asts) => | 
| 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 | 82 | raised exn "AST" (msg :: map (Pretty.string_of o Ast.pretty_ast) asts) | 
| 44247 | 83 | | TYPE (msg, Ts, ts) => | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 84 | raised exn "TYPE" (msg :: | 
| 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 | 85 | (if_context context Syntax.string_of_typ Ts @ | 
| 
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 | 86 | if_context context Syntax.string_of_term ts)) | 
| 44247 | 87 | | TERM (msg, ts) => | 
| 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 | 88 | raised exn "TERM" (msg :: if_context context Syntax.string_of_term ts) | 
| 44247 | 89 | | 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 | 90 |                 raised exn ("THM " ^ string_of_int i)
 | 
| 
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 | 91 | (msg :: if_context context Display.string_of_thm thms) | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 92 | | _ => raised exn (General.exnMessage exn) []); | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 93 | in [(i, msg)] end) | 
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 94 | and sorted_msgs context exn = | 
| 44271 
89f40a5939b2
more precise treatment of exception nesting and serial numbers;
 wenzelm parents: 
44270diff
changeset | 95 | sort_distinct (int_ord o pairself 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 | 96 | |
| 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 97 | 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 | 98 | |
| 
4a4d34d2f97b
more careful treatment of nested multi-exceptions, collapsing cumulatively empty list to Interrupt;
 wenzelm parents: 
37370diff
changeset | 99 | fun exn_message exn_position exn = | 
| 
4a4d34d2f97b
more careful treatment of nested multi-exceptions, collapsing cumulatively empty list to Interrupt;
 wenzelm parents: 
37370diff
changeset | 100 | (case exn_messages exn_position exn of | 
| 38875 
c7a66b584147
tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
 wenzelm parents: 
38874diff
changeset | 101 | [] => "Interrupt" | 
| 44270 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 wenzelm parents: 
44264diff
changeset | 102 | | 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 | 103 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 104 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 105 | (** controlled execution **) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 106 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 107 | fun debugging 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: 
45197diff
changeset | 108 | if ! debug | 
| 
600682331b79
more detailed exception_trace: in Poly/ML 5.4.x intermediate handlers reset the trace (cf. Poly/ML SVN 1110);
 wenzelm parents: 
45197diff
changeset | 109 | then exception_trace (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 | 110 | else f x; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 111 | |
| 41715 
22f8c2483bd2
explicit Multithreading.interrupted to ensure that interrupts stay within the boundaries of managed evaluation blocks;
 wenzelm parents: 
40318diff
changeset | 112 | fun controlled_execution f x = | 
| 44295 
e43f0ea90c9a
more focused use of Multithreading.interrupted: retain interrupts within task group boundary, without loss of information;
 wenzelm parents: 
44271diff
changeset | 113 | (f |> debugging |> 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 | 114 | |
| 39238 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 115 | 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 | 116 | handle exn => | 
| 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 117 | if Exn.is_interrupt exn then reraise exn | 
| 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 118 | else | 
| 
7189a138dd6c
refined Runtime.toplevel_error/Document.run_command: let interrupts pass unhindered;
 wenzelm parents: 
39232diff
changeset | 119 | let | 
| 45197 
b6c527c64789
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
 wenzelm parents: 
44295diff
changeset | 120 | val opt_ctxt = | 
| 
b6c527c64789
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
 wenzelm parents: 
44295diff
changeset | 121 | (case Context.thread_data () of | 
| 
b6c527c64789
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
 wenzelm parents: 
44295diff
changeset | 122 | NONE => NONE | 
| 
b6c527c64789
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);
 wenzelm parents: 
44295diff
changeset | 123 | | 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 | 124 | 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 | 125 | 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 | 126 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 127 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 128 |