src/Pure/Isar/toplevel.ML
author wenzelm
Thu, 02 Oct 2008 14:22:40 +0200
changeset 28458 0966ac3f4a40
parent 28453 06702e7acd1d
child 28463 b8f16c92122a
permissions -rw-r--r--
simplified Exn.EXCEPTIONS;
     1 (*  Title:      Pure/Isar/toplevel.ML
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Muenchen
     4 
     5 Isabelle/Isar toplevel transactions.
     6 *)
     7 
     8 signature TOPLEVEL =
     9 sig
    10   exception UNDEF
    11   type generic_theory
    12   type node
    13   val theory_node: node -> generic_theory option
    14   val proof_node: node -> ProofNode.T option
    15   val cases_node: (generic_theory -> 'a) -> (Proof.state -> 'a) -> node -> 'a
    16   val presentation_context: node option -> xstring option -> Proof.context
    17   type state
    18   val toplevel: state
    19   val is_toplevel: state -> bool
    20   val is_theory: state -> bool
    21   val is_proof: state -> bool
    22   val level: state -> int
    23   val previous_node_of: state -> node option
    24   val node_of: state -> node
    25   val node_case: (generic_theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a
    26   val context_of: state -> Proof.context
    27   val generic_theory_of: state -> generic_theory
    28   val theory_of: state -> theory
    29   val proof_of: state -> Proof.state
    30   val proof_position_of: state -> int
    31   val enter_proof_body: state -> Proof.state
    32   val print_state_context: state -> unit
    33   val print_state: bool -> state -> unit
    34   val quiet: bool ref
    35   val debug: bool ref
    36   val interact: bool ref
    37   val timing: bool ref
    38   val profiling: int ref
    39   val skip_proofs: bool ref
    40   exception TERMINATE
    41   exception TOPLEVEL_ERROR
    42   exception CONTEXT of Proof.context * exn
    43   val exn_message: exn -> string
    44   val program: (unit -> 'a) -> 'a
    45   type transition
    46   val empty: transition
    47   val init_of: transition -> string option
    48   val name_of: transition -> string
    49   val pos_of: transition -> Position.T
    50   val str_of: transition -> string
    51   val name: string -> transition -> transition
    52   val position: Position.T -> transition -> transition
    53   val interactive: bool -> transition -> transition
    54   val print: transition -> transition
    55   val no_timing: transition -> transition
    56   val init_theory: string -> (bool -> theory) -> (theory -> unit) -> transition -> transition
    57   val exit: transition -> transition
    58   val keep: (state -> unit) -> transition -> transition
    59   val keep': (bool -> state -> unit) -> transition -> transition
    60   val imperative: (unit -> unit) -> transition -> transition
    61   val ignored: Position.T -> transition
    62   val malformed: Position.T -> string -> transition
    63   val theory: (theory -> theory) -> transition -> transition
    64   val generic_theory: (generic_theory -> generic_theory) -> transition -> transition
    65   val theory': (bool -> theory -> theory) -> transition -> transition
    66   val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
    67   val end_local_theory: transition -> transition
    68   val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition
    69   val present_local_theory: xstring option -> (node -> unit) -> transition -> transition
    70   val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) ->
    71     transition -> transition
    72   val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) ->
    73     transition -> transition
    74   val theory_to_proof: (theory -> Proof.state) -> transition -> transition
    75   val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition
    76   val forget_proof: transition -> transition
    77   val present_proof: (node -> unit) -> transition -> transition
    78   val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition
    79   val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition
    80   val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition
    81   val proof: (Proof.state -> Proof.state) -> transition -> transition
    82   val actual_proof: (ProofNode.T -> ProofNode.T) -> transition -> transition
    83   val skip_proof: (int -> int) -> transition -> transition
    84   val skip_proof_to_theory: (int -> bool) -> transition -> transition
    85   val get_id: transition -> string option
    86   val put_id: string -> transition -> transition
    87   val unknown_theory: transition -> transition
    88   val unknown_proof: transition -> transition
    89   val unknown_context: transition -> transition
    90   val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
    91   val status: transition -> Markup.T -> unit
    92   val error_msg: transition -> exn * string -> unit
    93   val add_hook: (transition -> state -> state -> unit) -> unit
    94   val transition: bool -> transition -> state -> (state * (exn * string) option) option
    95   val commit_exit: Position.T -> transition
    96   val command: transition -> state -> state
    97   val excursion: (transition * transition list) list -> (transition * state) list * (unit -> unit)
    98 end;
    99 
   100 structure Toplevel: TOPLEVEL =
   101 struct
   102 
   103 (** toplevel state **)
   104 
   105 exception UNDEF;
   106 
   107 
   108 (* local theory wrappers *)
   109 
   110 type generic_theory = Context.generic;    (*theory or local_theory*)
   111 
   112 val loc_init = TheoryTarget.context;
   113 val loc_exit = LocalTheory.exit_global;
   114 
   115 fun loc_begin loc (Context.Theory thy) = loc_init (the_default "-" loc) thy
   116   | loc_begin NONE (Context.Proof lthy) = lthy
   117   | loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy);
   118 
   119 fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit
   120   | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore
   121   | loc_finish (SOME _) (Context.Proof lthy) = fn lthy' =>
   122       Context.Proof (LocalTheory.reinit (LocalTheory.raw_theory (K (loc_exit lthy')) lthy));
   123 
   124 
   125 (* datatype node *)
   126 
   127 datatype node =
   128   Theory of generic_theory * Proof.context option
   129     (*theory with presentation context*) |
   130   Proof of ProofNode.T * ((Proof.context -> generic_theory) * generic_theory)
   131     (*proof node, finish, original theory*) |
   132   SkipProof of int * (generic_theory * generic_theory);
   133     (*proof depth, resulting theory, original theory*)
   134 
   135 val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF;
   136 val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
   137 val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
   138 
   139 fun cases_node f _ (Theory (gthy, _)) = f gthy
   140   | cases_node _ g (Proof (prf, _)) = g (ProofNode.current prf)
   141   | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy;
   142 
   143 fun presentation_context (SOME (Theory (_, SOME ctxt))) NONE = ctxt
   144   | presentation_context (SOME node) NONE = cases_node Context.proof_of Proof.context_of node
   145   | presentation_context (SOME node) (SOME loc) =
   146       loc_init loc (cases_node Context.theory_of Proof.theory_of node)
   147   | presentation_context NONE _ = raise UNDEF;
   148 
   149 fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE)
   150   | reset_presentation node = node;
   151 
   152 
   153 (* datatype state *)
   154 
   155 type node_info = node * (theory -> unit);  (*node with exit operation*)
   156 datatype state = State of node_info option * node_info option;  (*current, previous*)
   157 
   158 val toplevel = State (NONE, NONE);
   159 
   160 fun is_toplevel (State (NONE, _)) = true
   161   | is_toplevel _ = false;
   162 
   163 fun level (State (NONE, _)) = 0
   164   | level (State (SOME (Theory _, _), _)) = 0
   165   | level (State (SOME (Proof (prf, _), _), _)) = Proof.level (ProofNode.current prf)
   166   | level (State (SOME (SkipProof (d, _), _), _)) = d + 1;   (*different notion of proof depth!*)
   167 
   168 fun str_of_state (State (NONE, _)) = "at top level"
   169   | str_of_state (State (SOME (Theory (Context.Theory _, _), _), _)) = "in theory mode"
   170   | str_of_state (State (SOME (Theory (Context.Proof _, _), _), _)) = "in local theory mode"
   171   | str_of_state (State (SOME (Proof _, _), _)) = "in proof mode"
   172   | str_of_state (State (SOME (SkipProof _, _), _)) = "in skipped proof mode";
   173 
   174 
   175 (* current node *)
   176 
   177 fun previous_node_of (State (_, prev)) = Option.map #1 prev;
   178 
   179 fun node_of (State (NONE, _)) = raise UNDEF
   180   | node_of (State (SOME (node, _), _)) = node;
   181 
   182 fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state));
   183 fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state));
   184 
   185 fun node_case f g state = cases_node f g (node_of state);
   186 
   187 val context_of = node_case Context.proof_of Proof.context_of;
   188 val generic_theory_of = node_case I (Context.Proof o Proof.context_of);
   189 val theory_of = node_case Context.theory_of Proof.theory_of;
   190 val proof_of = node_case (fn _ => raise UNDEF) I;
   191 
   192 fun proof_position_of state =
   193   (case node_of state of
   194     Proof (prf, _) => ProofNode.position prf
   195   | _ => raise UNDEF);
   196 
   197 val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward;
   198 
   199 
   200 (* print state *)
   201 
   202 val pretty_context = LocalTheory.pretty o Context.cases (TheoryTarget.init NONE) I;
   203 
   204 fun print_state_context state =
   205   (case try node_of state of
   206     NONE => []
   207   | SOME (Theory (gthy, _)) => pretty_context gthy
   208   | SOME (Proof (_, (_, gthy))) => pretty_context gthy
   209   | SOME (SkipProof (_, (gthy, _))) => pretty_context gthy)
   210   |> Pretty.chunks |> Pretty.writeln;
   211 
   212 fun print_state prf_only state =
   213   (case try node_of state of
   214     NONE => []
   215   | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy
   216   | SOME (Proof (prf, _)) =>
   217       Proof.pretty_state (ProofNode.position prf) (ProofNode.current prf)
   218   | SOME (SkipProof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)])
   219   |> Pretty.markup_chunks Markup.state |> Pretty.writeln;
   220 
   221 
   222 
   223 (** toplevel transitions **)
   224 
   225 val quiet = ref false;
   226 val debug = Output.debugging;
   227 val interact = ref false;
   228 val timing = Output.timing;
   229 val profiling = ref 0;
   230 val skip_proofs = ref false;
   231 
   232 exception TERMINATE;
   233 exception EXCURSION_FAIL of exn * string;
   234 exception FAILURE of state * exn;
   235 exception TOPLEVEL_ERROR;
   236 
   237 
   238 (* print exceptions *)
   239 
   240 exception CONTEXT of Proof.context * exn;
   241 
   242 fun exn_context NONE exn = exn
   243   | exn_context (SOME ctxt) exn = CONTEXT (ctxt, exn);
   244 
   245 local
   246 
   247 fun if_context NONE _ _ = []
   248   | if_context (SOME ctxt) f xs = map (f ctxt) xs;
   249 
   250 fun raised name [] = "exception " ^ name ^ " raised"
   251   | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg
   252   | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);
   253 
   254 in
   255 
   256 fun exn_message e =
   257   let
   258     val detailed = ! debug;
   259 
   260     fun exn_msg _ (CONTEXT (ctxt, exn)) = exn_msg (SOME ctxt) exn
   261       | exn_msg ctxt (Exn.EXCEPTIONS exns) = cat_lines (map (exn_msg ctxt) exns)
   262       | exn_msg ctxt (EXCURSION_FAIL (exn, loc)) =
   263           exn_msg ctxt exn ^ Markup.markup Markup.location ("\n" ^ loc)
   264       | exn_msg _ TERMINATE = "Exit."
   265       | exn_msg _ Exn.Interrupt = "Interrupt."
   266       | exn_msg _ TimeLimit.TimeOut = "Timeout."
   267       | exn_msg _ TOPLEVEL_ERROR = "Error."
   268       | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg
   269       | exn_msg _ (ERROR msg) = msg
   270       | exn_msg _ (Fail msg) = raised "Fail" [msg]
   271       | exn_msg _ (THEORY (msg, thys)) =
   272           raised "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else []))
   273       | exn_msg _ (Syntax.AST (msg, asts)) = raised "AST" (msg ::
   274             (if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else []))
   275       | exn_msg ctxt (TYPE (msg, Ts, ts)) = raised "TYPE" (msg ::
   276             (if detailed then
   277               if_context ctxt Syntax.string_of_typ Ts @ if_context ctxt Syntax.string_of_term ts
   278              else []))
   279       | exn_msg ctxt (TERM (msg, ts)) = raised "TERM" (msg ::
   280             (if detailed then if_context ctxt Syntax.string_of_term ts else []))
   281       | exn_msg ctxt (THM (msg, i, thms)) = raised ("THM " ^ string_of_int i) (msg ::
   282             (if detailed then if_context ctxt ProofContext.string_of_thm thms else []))
   283       | exn_msg _ exn = raised (General.exnMessage exn) []
   284   in exn_msg NONE e end;
   285 
   286 end;
   287 
   288 
   289 (* controlled execution *)
   290 
   291 local
   292 
   293 fun debugging f x =
   294   if ! debug then exception_trace (fn () => f x)
   295   else f x;
   296 
   297 fun toplevel_error f x =
   298   let val ctxt = try ML_Context.the_local_context () in
   299     f x handle exn =>
   300       (Output.error_msg (exn_message (exn_context ctxt exn)); raise TOPLEVEL_ERROR)
   301   end;
   302 
   303 in
   304 
   305 fun controlled_execution f =
   306   f
   307   |> debugging
   308   |> interruptible;
   309 
   310 fun program f =
   311  (f
   312   |> debugging
   313   |> toplevel_error) ();
   314 
   315 end;
   316 
   317 
   318 (* node transactions -- maintaining stable checkpoints *)
   319 
   320 local
   321 
   322 fun is_draft_theory (Theory (gthy, _)) = Context.is_draft (Context.theory_of gthy)
   323   | is_draft_theory _ = false;
   324 
   325 fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false;
   326 
   327 fun stale_error NONE = SOME (ERROR "Stale theory encountered after successful execution!")
   328   | stale_error some = some;
   329 
   330 fun map_theory f (Theory (gthy, ctxt)) =
   331       Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, ctxt)
   332   | map_theory _ node = node;
   333 
   334 in
   335 
   336 fun apply_transaction pos f (node, exit) =
   337   let
   338     val _ = is_draft_theory node andalso error "Illegal draft theory in toplevel state";
   339     val cont_node = reset_presentation node;
   340     val back_node = map_theory (Theory.checkpoint o Theory.copy) cont_node;
   341     fun state_error e nd = (State (SOME (nd, exit), SOME (node, exit)), e);
   342 
   343     val (result, err) =
   344       cont_node
   345       |> controlled_execution f
   346       |> map_theory Theory.checkpoint
   347       |> state_error NONE
   348       handle exn => state_error (SOME exn) cont_node;
   349 
   350     val (result', err') =
   351       if is_stale result then state_error (stale_error err) back_node
   352       else (result, err);
   353   in
   354     (case err' of
   355       NONE => result'
   356     | SOME exn => raise FAILURE (result', exn))
   357   end;
   358 
   359 end;
   360 
   361 
   362 (* primitive transitions *)
   363 
   364 datatype trans =
   365   Init of string * (bool -> theory) * (theory -> unit) | (*theory name, init, exit*)
   366   Exit |                                         (*formal exit of theory -- without committing*)
   367   CommitExit |                                   (*exit and commit final theory*)
   368   Keep of bool -> state -> unit |                (*peek at state*)
   369   Transaction of bool -> node -> node;           (*node transaction*)
   370 
   371 local
   372 
   373 fun apply_tr int _ (Init (_, f, exit)) (State (NONE, _)) =
   374       State (SOME (Theory (Context.Theory (Theory.checkpoint (f int)), NONE), exit), NONE)
   375   | apply_tr _ _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) =
   376       State (NONE, prev)
   377   | apply_tr _ _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) =
   378       (controlled_execution exit thy; toplevel)
   379   | apply_tr int _ (Keep f) state =
   380       controlled_execution (fn x => tap (f int) x) state
   381   | apply_tr int pos (Transaction f) (State (SOME state, _)) =
   382       apply_transaction pos (fn x => f int x) state
   383   | apply_tr _ _ _ _ = raise UNDEF;
   384 
   385 fun apply_union _ _ [] state = raise FAILURE (state, UNDEF)
   386   | apply_union int pos (tr :: trs) state =
   387       apply_union int pos trs state
   388         handle UNDEF => apply_tr int pos tr state
   389           | FAILURE (alt_state, UNDEF) => apply_tr int pos tr alt_state
   390           | exn as FAILURE _ => raise exn
   391           | exn => raise FAILURE (state, exn);
   392 
   393 in
   394 
   395 fun apply_trans int pos trs state = (apply_union int pos trs state, NONE)
   396   handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
   397 
   398 end;
   399 
   400 
   401 (* datatype transition *)
   402 
   403 datatype transition = Transition of
   404  {name: string,              (*command name*)
   405   pos: Position.T,           (*source position*)
   406   int_only: bool,            (*interactive-only*)
   407   print: bool,               (*print result state*)
   408   no_timing: bool,           (*suppress timing*)
   409   trans: trans list};        (*primitive transitions (union)*)
   410 
   411 fun make_transition (name, pos, int_only, print, no_timing, trans) =
   412   Transition {name = name, pos = pos, int_only = int_only, print = print, no_timing = no_timing,
   413     trans = trans};
   414 
   415 fun map_transition f (Transition {name, pos, int_only, print, no_timing, trans}) =
   416   make_transition (f (name, pos, int_only, print, no_timing, trans));
   417 
   418 val empty = make_transition ("", Position.none, false, false, false, []);
   419 
   420 
   421 (* diagnostics *)
   422 
   423 fun init_of (Transition {trans = [Init (name, _, _)], ...}) = SOME name
   424   | init_of _ = NONE;
   425 
   426 fun name_of (Transition {name, ...}) = name;
   427 fun pos_of (Transition {pos, ...}) = pos;
   428 fun str_of tr = quote (name_of tr) ^ Position.str_of (pos_of tr);
   429 
   430 fun command_msg msg tr = msg ^ "command " ^ str_of tr;
   431 fun at_command tr = command_msg "At " tr ^ ".";
   432 
   433 fun type_error tr state =
   434   ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state);
   435 
   436 
   437 (* modify transitions *)
   438 
   439 fun name name = map_transition (fn (_, pos, int_only, print, no_timing, trans) =>
   440   (name, pos, int_only, print, no_timing, trans));
   441 
   442 fun position pos = map_transition (fn (name, _, int_only, print, no_timing, trans) =>
   443   (name, pos, int_only, print, no_timing, trans));
   444 
   445 fun interactive int_only = map_transition (fn (name, pos, _, print, no_timing, trans) =>
   446   (name, pos, int_only, print, no_timing, trans));
   447 
   448 val no_timing = map_transition (fn (name, pos, int_only, print, _, trans) =>
   449   (name, pos, int_only, print, true, trans));
   450 
   451 fun add_trans tr = map_transition (fn (name, pos, int_only, print, no_timing, trans) =>
   452   (name, pos, int_only, print, no_timing, tr :: trans));
   453 
   454 val reset_trans = map_transition (fn (name, pos, int_only, print, no_timing, _) =>
   455   (name, pos, int_only, print, no_timing, []));
   456 
   457 fun set_print print = map_transition (fn (name, pos, int_only, _, no_timing, trans) =>
   458   (name, pos, int_only, print, no_timing, trans));
   459 
   460 val print = set_print true;
   461 
   462 
   463 (* basic transitions *)
   464 
   465 fun init_theory name f exit = add_trans (Init (name, f, exit));
   466 val exit = add_trans Exit;
   467 val keep' = add_trans o Keep;
   468 fun transaction f = add_trans (Transaction f);
   469 
   470 fun keep f = add_trans (Keep (fn _ => f));
   471 fun imperative f = keep (fn _ => f ());
   472 
   473 fun ignored pos = empty |> name "<ignored>" |> position pos |> imperative I;
   474 fun malformed pos msg =
   475   empty |> name "<malformed>" |> position pos |> imperative (fn () => error msg);
   476 
   477 val unknown_theory = imperative (fn () => warning "Unknown theory context");
   478 val unknown_proof = imperative (fn () => warning "Unknown proof context");
   479 val unknown_context = imperative (fn () => warning "Unknown context");
   480 
   481 
   482 (* theory transitions *)
   483 
   484 fun generic_theory f = transaction (fn _ =>
   485   (fn Theory (gthy, _) => Theory (f gthy, NONE)
   486     | _ => raise UNDEF));
   487 
   488 fun theory' f = transaction (fn int =>
   489   (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE)
   490     | _ => raise UNDEF));
   491 
   492 fun theory f = theory' (K f);
   493 
   494 fun begin_local_theory begin f = transaction (fn _ =>
   495   (fn Theory (Context.Theory thy, _) =>
   496         let
   497           val lthy = f thy;
   498           val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy);
   499         in Theory (gthy, SOME lthy) end
   500     | _ => raise UNDEF));
   501 
   502 val end_local_theory = transaction (fn _ =>
   503   (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy)
   504     | _ => raise UNDEF));
   505 
   506 local
   507 
   508 fun local_theory_presentation loc f g = transaction (fn int =>
   509   (fn Theory (gthy, _) =>
   510         let
   511           val finish = loc_finish loc gthy;
   512           val lthy' = f (loc_begin loc gthy);
   513         in Theory (finish lthy', SOME lthy') end
   514     | _ => raise UNDEF) #> tap g);
   515 
   516 in
   517 
   518 fun local_theory loc f = local_theory_presentation loc f (K I);
   519 fun present_local_theory loc g = local_theory_presentation loc I g;
   520 
   521 end;
   522 
   523 
   524 (* proof transitions *)
   525 
   526 fun end_proof f = transaction (fn int =>
   527   (fn Proof (prf, (finish, _)) =>
   528         let val state = ProofNode.current prf in
   529           if can (Proof.assert_bottom true) state then
   530             let
   531               val ctxt' = f int state;
   532               val gthy' = finish ctxt';
   533             in Theory (gthy', SOME ctxt') end
   534           else raise UNDEF
   535         end
   536     | SkipProof (0, (gthy, _)) => Theory (gthy, NONE)
   537     | _ => raise UNDEF));
   538 
   539 local
   540 
   541 fun begin_proof init finish = transaction (fn int =>
   542   (fn Theory (gthy, _) =>
   543     let
   544       val prf = init int gthy;
   545       val skip = ! skip_proofs;
   546       val schematic = Proof.schematic_goal prf;
   547     in
   548       if skip andalso schematic then
   549         warning "Cannot skip proof of schematic goal statement"
   550       else ();
   551       if skip andalso not schematic then
   552         SkipProof (0, (finish gthy (Proof.global_skip_proof int prf), gthy))
   553       else Proof (ProofNode.init prf, (finish gthy, gthy))
   554     end
   555   | _ => raise UNDEF));
   556 
   557 in
   558 
   559 fun local_theory_to_proof' loc f = begin_proof
   560   (fn int => fn gthy => f int (loc_begin loc gthy))
   561   (loc_finish loc);
   562 
   563 fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f);
   564 
   565 fun theory_to_proof f = begin_proof
   566   (K (fn Context.Theory thy => f thy | _ => raise UNDEF))
   567   (K (Context.Theory o ProofContext.theory_of));
   568 
   569 end;
   570 
   571 val forget_proof = transaction (fn _ =>
   572   (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
   573     | SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
   574     | _ => raise UNDEF));
   575 
   576 fun present_proof f = transaction (fn _ =>
   577   (fn Proof (prf, x) => Proof (ProofNode.apply I prf, x)
   578     | skip as SkipProof _ => skip
   579     | _ => raise UNDEF) #> tap f);
   580 
   581 fun proofs' f = transaction (fn int =>
   582   (fn Proof (prf, x) => Proof (ProofNode.applys (f int) prf, x)
   583     | skip as SkipProof _ => skip
   584     | _ => raise UNDEF));
   585 
   586 fun proof' f = proofs' (Seq.single oo f);
   587 val proofs = proofs' o K;
   588 val proof = proof' o K;
   589 
   590 fun actual_proof f = transaction (fn _ =>
   591   (fn Proof (prf, x) => Proof (f prf, x)
   592     | _ => raise UNDEF));
   593 
   594 fun skip_proof f = transaction (fn _ =>
   595   (fn SkipProof (h, x) => SkipProof (f h, x)
   596     | _ => raise UNDEF));
   597 
   598 fun skip_proof_to_theory pred = transaction (fn _ =>
   599   (fn SkipProof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF
   600   | _ => raise UNDEF));
   601 
   602 
   603 
   604 (** toplevel transactions **)
   605 
   606 (* identification *)
   607 
   608 fun get_id (Transition {pos, ...}) = Position.get_id pos;
   609 fun put_id id (tr as Transition {pos, ...}) = position (Position.put_id id pos) tr;
   610 
   611 
   612 (* thread position *)
   613 
   614 fun setmp_thread_position (Transition {pos, ...}) f x =
   615   Position.setmp_thread_data pos f x;
   616 
   617 fun status tr m =
   618   setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) ();
   619 
   620 fun error_msg tr exn_info =
   621   setmp_thread_position tr (fn () => Output.error_msg (exn_message (EXCURSION_FAIL exn_info))) ();
   622 
   623 
   624 (* post-transition hooks *)
   625 
   626 local val hooks = ref ([]: (transition -> state -> state -> unit) list) in
   627 
   628 fun add_hook f = CRITICAL (fn () => change hooks (cons f));
   629 fun get_hooks () = CRITICAL (fn () => ! hooks);
   630 
   631 end;
   632 
   633 
   634 (* apply transitions *)
   635 
   636 local
   637 
   638 fun app int (tr as Transition {trans, pos, int_only, print, no_timing, ...}) =
   639   setmp_thread_position tr (fn state =>
   640     let
   641       val _ = if not int andalso int_only then warning (command_msg "Interactive-only " tr) else ();
   642 
   643       fun do_timing f x = (warning (command_msg "" tr); timeap f x);
   644       fun do_profiling f x = profile (! profiling) f x;
   645 
   646       val (result, status) =
   647          state |> (apply_trans int pos trans
   648           |> (if ! profiling > 0 andalso not no_timing then do_profiling else I)
   649           |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I));
   650 
   651       val _ = if int andalso not (! quiet) andalso print then print_state false result else ();
   652     in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) status) end);
   653 
   654 in
   655 
   656 fun transition int tr st =
   657   let
   658     val hooks = get_hooks ();
   659     fun apply_hooks st' = hooks |> List.app (fn f => (try (fn () => f tr st st') (); ()));
   660 
   661     val ctxt = try context_of st;
   662     val res =
   663       (case app int tr st of
   664         (_, SOME TERMINATE) => NONE
   665       | (st', SOME (EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info)
   666       | (st', SOME exn) => SOME (st', SOME (exn_context ctxt exn, at_command tr))
   667       | (st', NONE) => SOME (st', NONE));
   668     val _ = (case res of SOME (st', NONE) => apply_hooks st' | _ => ());
   669   in res end;
   670 
   671 end;
   672 
   673 
   674 (* commit final exit *)
   675 
   676 fun commit_exit pos =
   677   name "end" empty
   678   |> position pos
   679   |> add_trans CommitExit
   680   |> imperative (fn () => warning "Expected to find finished theory");
   681 
   682 
   683 (* nested commands *)
   684 
   685 fun command tr st =
   686   (case transition (! interact) tr st of
   687     SOME (st', NONE) => st'
   688   | SOME (_, SOME exn_info) => raise EXCURSION_FAIL exn_info
   689   | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));
   690 
   691 
   692 (* excursion of units, consisting of commands with proof *)
   693 
   694 fun command_result tr st =
   695   let val st' = command tr st
   696   in (st', st') end;
   697 
   698 fun unit_result immediate (tr, proof_trs) st =
   699   let val st' = command tr st in
   700     if immediate orelse null proof_trs orelse
   701       not (can proof_of st') orelse Proof.is_relevant (proof_of st')
   702     then
   703       let val (states, st'') = fold_map command_result proof_trs st'
   704       in (fn () => (tr, st') :: (proof_trs ~~ states), st'') end
   705     else
   706       let
   707         val (body_trs, end_tr) = split_last proof_trs;
   708         val body_states = ref ([]: state list);
   709         val finish = Context.Theory o ProofContext.theory_of;
   710         fun make_result prf =
   711           let val (states, State (result_node, _)) =
   712             (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
   713               => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev))
   714             |> fold_map command_result body_trs
   715             ||> command (end_tr |> set_print false)
   716           in body_states := states; presentation_context (Option.map #1 result_node) NONE end;
   717         val st'' = st'
   718           |> command (end_tr |> reset_trans |> end_proof (K (Proof.future_proof make_result)));
   719       in (fn () => (tr, st') :: (body_trs ~~ ! body_states) @ [(end_tr, st'')], st'') end
   720   end;
   721 
   722 fun excursion input =
   723   let
   724     val end_pos = if null input then error "No input" else pos_of (fst (List.last input));
   725 
   726     val immediate = not (! future_scheduler) orelse Multithreading.max_threads_value () <= 1;
   727     val (future_results, end_state) = fold_map (unit_result immediate) input toplevel;
   728     val _ =
   729       (case end_state of
   730         State (NONE, SOME (Theory (Context.Theory thy, _), _)) => Thm.join_futures thy
   731       | _ => error "Unfinished development at end of input");
   732     val results = maps (fn res => res ()) future_results;
   733   in (results, fn () => ignore (command (commit_exit end_pos) end_state)) end;
   734 
   735 end;