| author | haftmann | 
| Fri, 03 Jul 2009 23:29:03 +0200 | |
| changeset 31938 | f193d95b4632 | 
| parent 31476 | c5d2899b6de9 | 
| child 32091 | 30e2ffbba718 | 
| 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 | 
| 
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 | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 14 | 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 | 15 |   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 | 16 |   val controlled_execution: ('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 | 17 |   val toplevel_error: (exn -> string) -> ('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 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 19 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 20 | structure Runtime: RUNTIME = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 21 | struct | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 22 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 23 | (** exceptions **) | 
| 
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 | exception UNDEF; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 26 | exception TERMINATE; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 27 | 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 | 28 | exception TOPLEVEL_ERROR; | 
| 
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 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 31 | (* exn_context *) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 32 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 33 | 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 | 34 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 35 | fun exn_context NONE exn = exn | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 36 | | exn_context (SOME ctxt) exn = CONTEXT (ctxt, exn); | 
| 
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 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 39 | (* exn_message *) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 40 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 41 | 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 | 42 | | 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 | 43 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 44 | fun exn_message exn_position e = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 45 | let | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 46 | 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 | 47 | 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 | 48 | (case msgs of | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 49 | [] => "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 | 50 | | [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 | 51 |         | _ => 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 | 52 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 53 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 54 | val detailed = ! Output.debugging; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 55 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 56 | fun exn_msg _ (CONTEXT (ctxt, exn)) = exn_msg (SOME ctxt) exn | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 57 | | exn_msg ctxt (Exn.EXCEPTIONS exns) = cat_lines (map (exn_msg ctxt) exns) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 58 | | exn_msg ctxt (EXCURSION_FAIL (exn, loc)) = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 59 |           exn_msg ctxt exn ^ Markup.markup Markup.location ("\n" ^ loc)
 | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 60 | | exn_msg _ TERMINATE = "Exit." | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 61 | | exn_msg _ Exn.Interrupt = "Interrupt." | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 62 | | exn_msg _ TimeLimit.TimeOut = "Timeout." | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 63 | | exn_msg _ TOPLEVEL_ERROR = "Error." | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 64 | | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 65 | | exn_msg _ (ERROR msg) = msg | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 66 | | exn_msg _ (exn as Fail msg) = raised exn "Fail" [msg] | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 67 | | exn_msg _ (exn as THEORY (msg, thys)) = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 68 | raised exn "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else [])) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 69 | | exn_msg _ (exn as Syntax.AST (msg, asts)) = raised exn "AST" (msg :: | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 70 | (if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else [])) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 71 | | exn_msg ctxt (exn as TYPE (msg, Ts, ts)) = raised exn "TYPE" (msg :: | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 72 | (if detailed then | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 73 | if_context ctxt Syntax.string_of_typ Ts @ if_context ctxt Syntax.string_of_term ts | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 74 | else [])) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 75 | | exn_msg ctxt (exn as TERM (msg, ts)) = raised exn "TERM" (msg :: | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 76 | (if detailed then if_context ctxt Syntax.string_of_term ts else [])) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 77 |       | exn_msg ctxt (exn as THM (msg, i, thms)) = raised exn ("THM " ^ string_of_int i) (msg ::
 | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 78 | (if detailed then if_context ctxt ProofContext.string_of_thm thms else [])) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 79 | | exn_msg _ exn = raised exn (General.exnMessage exn) []; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 80 | in exn_msg NONE e end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 81 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 82 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 83 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 84 | (** controlled execution **) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 85 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 86 | fun debugging f x = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 87 | if ! Output.debugging then | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 88 | Exn.release (exception_trace (fn () => | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 89 | Exn.Result (f x) handle | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 90 | exn as UNDEF => Exn.Exn exn | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 91 | | exn as EXCURSION_FAIL _ => Exn.Exn exn)) | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 92 | else f x; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 93 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 94 | fun controlled_execution f = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 95 | f | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 96 | |> debugging | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 97 | |> Future.interruptible_task; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 98 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 99 | fun toplevel_error exn_message f x = | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 100 | let val ctxt = Option.map Context.proof_of (Context.thread_data ()) in | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 101 | f x handle exn => | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 102 | (Output.error_msg (exn_message (exn_context ctxt exn)); | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 103 | raise TOPLEVEL_ERROR) | 
| 
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 | |
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 106 | end; | 
| 
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
 wenzelm parents: diff
changeset | 107 |