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