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