src/Pure/Isar/toplevel.ML
author wenzelm
Sat, 21 Mar 2009 13:11:12 +0100
changeset 30618 046f4f986fb5
parent 30571 c12484a27367
child 30801 9bdf001bea58
permissions -rw-r--r--
restricted interrupts for tasks running as future worker thread -- attempt to prevent interrupt race conditions;

(*  Title:      Pure/Isar/toplevel.ML
    Author:     Markus Wenzel, TU Muenchen

Isabelle/Isar toplevel transactions.
*)

signature TOPLEVEL =
sig
  exception UNDEF
  type generic_theory
  type node
  val theory_node: node -> generic_theory option
  val proof_node: node -> ProofNode.T option
  val cases_node: (generic_theory -> 'a) -> (Proof.state -> 'a) -> node -> 'a
  val context_node: node -> Proof.context
  type state
  val toplevel: state
  val is_toplevel: state -> bool
  val is_theory: state -> bool
  val is_proof: state -> bool
  val level: state -> int
  val previous_node_of: state -> node option
  val node_of: state -> node
  val node_case: (generic_theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a
  val presentation_context_of: state -> Proof.context
  val context_of: state -> Proof.context
  val generic_theory_of: state -> generic_theory
  val theory_of: state -> theory
  val proof_of: state -> Proof.state
  val proof_position_of: state -> int
  val enter_proof_body: state -> Proof.state
  val print_state_context: state -> unit
  val print_state: bool -> state -> unit
  val quiet: bool ref
  val debug: bool ref
  val interact: bool ref
  val timing: bool ref
  val profiling: int ref
  val skip_proofs: bool ref
  exception TERMINATE
  exception TOPLEVEL_ERROR
  exception CONTEXT of Proof.context * exn
  val exn_message: exn -> string
  val program: (unit -> 'a) -> 'a
  type transition
  val empty: transition
  val init_of: transition -> string option
  val name_of: transition -> string
  val pos_of: transition -> Position.T
  val str_of: transition -> string
  val name: string -> transition -> transition
  val position: Position.T -> transition -> transition
  val interactive: bool -> transition -> transition
  val print: transition -> transition
  val no_timing: transition -> transition
  val init_theory: string -> (bool -> theory) -> (theory -> unit) -> transition -> transition
  val exit: transition -> transition
  val keep: (state -> unit) -> transition -> transition
  val keep': (bool -> state -> unit) -> transition -> transition
  val imperative: (unit -> unit) -> transition -> transition
  val ignored: Position.T -> transition
  val malformed: Position.T -> string -> transition
  val theory: (theory -> theory) -> transition -> transition
  val generic_theory: (generic_theory -> generic_theory) -> transition -> transition
  val theory': (bool -> theory -> theory) -> transition -> transition
  val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
  val end_local_theory: transition -> transition
  val local_theory': xstring option -> (bool -> local_theory -> local_theory) ->
    transition -> transition
  val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition
  val present_local_theory: xstring option -> (state -> unit) -> transition -> transition
  val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) ->
    transition -> transition
  val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) ->
    transition -> transition
  val theory_to_proof: (theory -> Proof.state) -> transition -> transition
  val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition
  val forget_proof: transition -> transition
  val present_proof: (state -> unit) -> transition -> transition
  val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition
  val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition
  val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition
  val proof: (Proof.state -> Proof.state) -> transition -> transition
  val actual_proof: (ProofNode.T -> ProofNode.T) -> transition -> transition
  val skip_proof: (int -> int) -> transition -> transition
  val skip_proof_to_theory: (int -> bool) -> transition -> transition
  val get_id: transition -> string option
  val put_id: string -> transition -> transition
  val unknown_theory: transition -> transition
  val unknown_proof: transition -> transition
  val unknown_context: transition -> transition
  val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
  val status: transition -> Markup.T -> unit
  val error_msg: transition -> exn * string -> unit
  val add_hook: (transition -> state -> state -> unit) -> unit
  val transition: bool -> transition -> state -> (state * (exn * string) option) option
  val commit_exit: Position.T -> transition
  val command: transition -> state -> state
  val run_command: string -> transition -> state -> state option
  val excursion: (transition * transition list) list -> (transition * state) list lazy
end;

structure Toplevel: TOPLEVEL =
struct

(** toplevel state **)

exception UNDEF;


(* local theory wrappers *)

type generic_theory = Context.generic;    (*theory or local_theory*)

val loc_init = TheoryTarget.context;
val loc_exit = LocalTheory.exit_global;

fun loc_begin loc (Context.Theory thy) = loc_init (the_default "-" loc) thy
  | loc_begin NONE (Context.Proof lthy) = lthy
  | loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy);

fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit
  | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore
  | loc_finish (SOME _) (Context.Proof lthy) = fn lthy' =>
      Context.Proof (LocalTheory.reinit (LocalTheory.raw_theory (K (loc_exit lthy')) lthy));


(* datatype node *)

datatype node =
  Theory of generic_theory * Proof.context option
    (*theory with presentation context*) |
  Proof of ProofNode.T * ((Proof.context -> generic_theory) * generic_theory)
    (*proof node, finish, original theory*) |
  SkipProof of int * (generic_theory * generic_theory);
    (*proof depth, resulting theory, original theory*)

val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF;
val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;

fun cases_node f _ (Theory (gthy, _)) = f gthy
  | cases_node _ g (Proof (prf, _)) = g (ProofNode.current prf)
  | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy;

val context_node = cases_node Context.proof_of Proof.context_of;


(* datatype state *)

type node_info = node * (theory -> unit);  (*node with exit operation*)
datatype state = State of node_info option * node_info option;  (*current, previous*)

val toplevel = State (NONE, NONE);

fun is_toplevel (State (NONE, _)) = true
  | is_toplevel _ = false;

fun level (State (NONE, _)) = 0
  | level (State (SOME (Theory _, _), _)) = 0
  | level (State (SOME (Proof (prf, _), _), _)) = Proof.level (ProofNode.current prf)
  | level (State (SOME (SkipProof (d, _), _), _)) = d + 1;   (*different notion of proof depth!*)

fun str_of_state (State (NONE, _)) = "at top level"
  | str_of_state (State (SOME (Theory (Context.Theory _, _), _), _)) = "in theory mode"
  | str_of_state (State (SOME (Theory (Context.Proof _, _), _), _)) = "in local theory mode"
  | str_of_state (State (SOME (Proof _, _), _)) = "in proof mode"
  | str_of_state (State (SOME (SkipProof _, _), _)) = "in skipped proof mode";


(* current node *)

fun previous_node_of (State (_, prev)) = Option.map #1 prev;

fun node_of (State (NONE, _)) = raise UNDEF
  | node_of (State (SOME (node, _), _)) = node;

fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state));
fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state));

fun node_case f g state = cases_node f g (node_of state);

fun presentation_context_of state =
  (case try node_of state of
    SOME (Theory (_, SOME ctxt)) => ctxt
  | SOME node => context_node node
  | NONE => raise UNDEF);

val context_of = node_case Context.proof_of Proof.context_of;
val generic_theory_of = node_case I (Context.Proof o Proof.context_of);
val theory_of = node_case Context.theory_of Proof.theory_of;
val proof_of = node_case (fn _ => raise UNDEF) I;

fun proof_position_of state =
  (case node_of state of
    Proof (prf, _) => ProofNode.position prf
  | _ => raise UNDEF);

val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward;


(* print state *)

val pretty_context = LocalTheory.pretty o Context.cases (TheoryTarget.init NONE) I;

fun print_state_context state =
  (case try node_of state of
    NONE => []
  | SOME (Theory (gthy, _)) => pretty_context gthy
  | SOME (Proof (_, (_, gthy))) => pretty_context gthy
  | SOME (SkipProof (_, (gthy, _))) => pretty_context gthy)
  |> Pretty.chunks |> Pretty.writeln;

fun print_state prf_only state =
  (case try node_of state of
    NONE => []
  | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy
  | SOME (Proof (prf, _)) =>
      Proof.pretty_state (ProofNode.position prf) (ProofNode.current prf)
  | SOME (SkipProof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)])
  |> Pretty.markup_chunks Markup.state |> Pretty.writeln;



(** toplevel transitions **)

val quiet = ref false;
val debug = Output.debugging;
val interact = ref false;
val timing = Output.timing;
val profiling = ref 0;
val skip_proofs = ref false;

exception TERMINATE;
exception EXCURSION_FAIL of exn * string;
exception FAILURE of state * exn;
exception TOPLEVEL_ERROR;


(* print exceptions *)

exception CONTEXT of Proof.context * exn;

fun exn_context NONE exn = exn
  | exn_context (SOME ctxt) exn = CONTEXT (ctxt, exn);

local

fun if_context NONE _ _ = []
  | if_context (SOME ctxt) f xs = map (f ctxt) xs;

fun raised name [] = "exception " ^ name ^ " raised"
  | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg
  | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);

in

fun exn_message e =
  let
    val detailed = ! debug;

    fun exn_msg _ (CONTEXT (ctxt, exn)) = exn_msg (SOME ctxt) exn
      | exn_msg ctxt (Exn.EXCEPTIONS exns) = cat_lines (map (exn_msg ctxt) exns)
      | exn_msg ctxt (EXCURSION_FAIL (exn, loc)) =
          exn_msg ctxt exn ^ Markup.markup Markup.location ("\n" ^ loc)
      | exn_msg _ TERMINATE = "Exit."
      | exn_msg _ Exn.Interrupt = "Interrupt."
      | exn_msg _ TimeLimit.TimeOut = "Timeout."
      | exn_msg _ TOPLEVEL_ERROR = "Error."
      | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg
      | exn_msg _ (ERROR msg) = msg
      | exn_msg _ (Fail msg) = raised "Fail" [msg]
      | exn_msg _ (THEORY (msg, thys)) =
          raised "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else []))
      | exn_msg _ (Syntax.AST (msg, asts)) = raised "AST" (msg ::
            (if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else []))
      | exn_msg ctxt (TYPE (msg, Ts, ts)) = raised "TYPE" (msg ::
            (if detailed then
              if_context ctxt Syntax.string_of_typ Ts @ if_context ctxt Syntax.string_of_term ts
             else []))
      | exn_msg ctxt (TERM (msg, ts)) = raised "TERM" (msg ::
            (if detailed then if_context ctxt Syntax.string_of_term ts else []))
      | exn_msg ctxt (THM (msg, i, thms)) = raised ("THM " ^ string_of_int i) (msg ::
            (if detailed then if_context ctxt ProofContext.string_of_thm thms else []))
      | exn_msg _ exn = raised (General.exnMessage exn) []
  in exn_msg NONE e end;

end;


(* controlled execution *)

local

fun debugging f x =
  if ! debug then
    Exn.release (exception_trace (fn () =>
      Exn.Result (f x) handle
        exn as UNDEF => Exn.Exn exn
      | exn as EXCURSION_FAIL _ => Exn.Exn exn))
  else f x;

fun toplevel_error f x =
  let val ctxt = try ML_Context.the_local_context () in
    f x handle exn =>
      (Output.error_msg (exn_message (exn_context ctxt exn)); raise TOPLEVEL_ERROR)
  end;

in

fun controlled_execution f =
  f
  |> debugging
  |> Future.interruptible_task;

fun program f =
 (f
  |> controlled_execution
  |> toplevel_error) ();

end;


(* node transactions -- maintaining stable checkpoints *)

local

fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE)
  | reset_presentation node = node;

fun is_draft_theory (Theory (gthy, _)) = Context.is_draft (Context.theory_of gthy)
  | is_draft_theory _ = false;

fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false;

fun stale_error NONE = SOME (ERROR "Stale theory encountered after successful execution!")
  | stale_error some = some;

fun map_theory f (Theory (gthy, ctxt)) =
      Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, ctxt)
  | map_theory _ node = node;

in

fun apply_transaction pos f g (node, exit) =
  let
    val _ = is_draft_theory node andalso error "Illegal draft theory in toplevel state";
    val cont_node = reset_presentation node;
    val back_node = map_theory (Theory.checkpoint o Theory.copy) cont_node;
    fun state_error e nd = (State (SOME (nd, exit), SOME (node, exit)), e);

    val (result, err) =
      cont_node
      |> controlled_execution f
      |> map_theory Theory.checkpoint
      |> state_error NONE
      handle exn => state_error (SOME exn) cont_node;

    val (result', err') =
      if is_stale result then state_error (stale_error err) back_node
      else (result, err);
  in
    (case err' of
      NONE => tap g result'
    | SOME exn => raise FAILURE (result', exn))
  end;

end;


(* primitive transitions *)

datatype trans =
  Init of string * (bool -> theory) * (theory -> unit) | (*theory name, init, exit*)
  Exit |                                         (*formal exit of theory -- without committing*)
  CommitExit |                                   (*exit and commit final theory*)
  Keep of bool -> state -> unit |                (*peek at state*)
  Transaction of (bool -> node -> node) * (state -> unit);  (*node transaction and presentation*)

local

fun apply_tr int _ (Init (_, f, exit)) (State (NONE, _)) =
      State (SOME (Theory (Context.Theory (Theory.checkpoint (f int)), NONE), exit), NONE)
  | apply_tr _ _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) =
      State (NONE, prev)
  | apply_tr _ _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) =
      (controlled_execution exit thy; toplevel)
  | apply_tr int _ (Keep f) state =
      controlled_execution (fn x => tap (f int) x) state
  | apply_tr int pos (Transaction (f, g)) (State (SOME state, _)) =
      apply_transaction pos (fn x => f int x) g state
  | apply_tr _ _ _ _ = raise UNDEF;

fun apply_union _ _ [] state = raise FAILURE (state, UNDEF)
  | apply_union int pos (tr :: trs) state =
      apply_union int pos trs state
        handle UNDEF => apply_tr int pos tr state
          | FAILURE (alt_state, UNDEF) => apply_tr int pos tr alt_state
          | exn as FAILURE _ => raise exn
          | exn => raise FAILURE (state, exn);

in

fun apply_trans int pos trs state = (apply_union int pos trs state, NONE)
  handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);

end;


(* datatype transition *)

datatype transition = Transition of
 {name: string,              (*command name*)
  pos: Position.T,           (*source position*)
  int_only: bool,            (*interactive-only*)
  print: bool,               (*print result state*)
  no_timing: bool,           (*suppress timing*)
  trans: trans list};        (*primitive transitions (union)*)

fun make_transition (name, pos, int_only, print, no_timing, trans) =
  Transition {name = name, pos = pos, int_only = int_only, print = print, no_timing = no_timing,
    trans = trans};

fun map_transition f (Transition {name, pos, int_only, print, no_timing, trans}) =
  make_transition (f (name, pos, int_only, print, no_timing, trans));

val empty = make_transition ("", Position.none, false, false, false, []);


(* diagnostics *)

fun init_of (Transition {trans = [Init (name, _, _)], ...}) = SOME name
  | init_of _ = NONE;

fun name_of (Transition {name, ...}) = name;
fun pos_of (Transition {pos, ...}) = pos;
fun str_of tr = quote (name_of tr) ^ Position.str_of (pos_of tr);

fun command_msg msg tr = msg ^ "command " ^ str_of tr;
fun at_command tr = command_msg "At " tr ^ ".";

fun type_error tr state =
  ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state);


(* modify transitions *)

fun name name = map_transition (fn (_, pos, int_only, print, no_timing, trans) =>
  (name, pos, int_only, print, no_timing, trans));

fun position pos = map_transition (fn (name, _, int_only, print, no_timing, trans) =>
  (name, pos, int_only, print, no_timing, trans));

fun interactive int_only = map_transition (fn (name, pos, _, print, no_timing, trans) =>
  (name, pos, int_only, print, no_timing, trans));

val no_timing = map_transition (fn (name, pos, int_only, print, _, trans) =>
  (name, pos, int_only, print, true, trans));

fun add_trans tr = map_transition (fn (name, pos, int_only, print, no_timing, trans) =>
  (name, pos, int_only, print, no_timing, tr :: trans));

val reset_trans = map_transition (fn (name, pos, int_only, print, no_timing, _) =>
  (name, pos, int_only, print, no_timing, []));

fun set_print print = map_transition (fn (name, pos, int_only, _, no_timing, trans) =>
  (name, pos, int_only, print, no_timing, trans));

val print = set_print true;


(* basic transitions *)

fun init_theory name f exit = add_trans (Init (name, f, exit));
val exit = add_trans Exit;
val keep' = add_trans o Keep;

fun present_transaction f g = add_trans (Transaction (f, g));
fun transaction f = present_transaction f (K ());

fun keep f = add_trans (Keep (fn _ => f));
fun imperative f = keep (fn _ => f ());

fun ignored pos = empty |> name "<ignored>" |> position pos |> imperative I;
fun malformed pos msg =
  empty |> name "<malformed>" |> position pos |> imperative (fn () => error msg);

val unknown_theory = imperative (fn () => warning "Unknown theory context");
val unknown_proof = imperative (fn () => warning "Unknown proof context");
val unknown_context = imperative (fn () => warning "Unknown context");


(* theory transitions *)

fun generic_theory f = transaction (fn _ =>
  (fn Theory (gthy, _) => Theory (f gthy, NONE)
    | _ => raise UNDEF));

fun theory' f = transaction (fn int =>
  (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE)
    | _ => raise UNDEF));

fun theory f = theory' (K f);

fun begin_local_theory begin f = transaction (fn _ =>
  (fn Theory (Context.Theory thy, _) =>
        let
          val lthy = f thy;
          val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy);
        in Theory (gthy, SOME lthy) end
    | _ => raise UNDEF));

val end_local_theory = transaction (fn _ =>
  (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy)
    | _ => raise UNDEF));

local

fun local_theory_presentation loc f = present_transaction (fn int =>
  (fn Theory (gthy, _) =>
        let
          val finish = loc_finish loc gthy;
          val lthy' = f int (loc_begin loc gthy);
        in Theory (finish lthy', SOME lthy') end
    | _ => raise UNDEF));

in

fun local_theory' loc f = local_theory_presentation loc f (K ());
fun local_theory loc f = local_theory' loc (K f);
fun present_local_theory loc = local_theory_presentation loc (K I);

end;


(* proof transitions *)

fun end_proof f = transaction (fn int =>
  (fn Proof (prf, (finish, _)) =>
        let val state = ProofNode.current prf in
          if can (Proof.assert_bottom true) state then
            let
              val ctxt' = f int state;
              val gthy' = finish ctxt';
            in Theory (gthy', SOME ctxt') end
          else raise UNDEF
        end
    | SkipProof (0, (gthy, _)) => Theory (gthy, NONE)
    | _ => raise UNDEF));

local

fun begin_proof init finish = transaction (fn int =>
  (fn Theory (gthy, _) =>
    let
      val prf = init int gthy;
      val skip = ! skip_proofs;
      val schematic = Proof.schematic_goal prf;
    in
      if skip andalso schematic then
        warning "Cannot skip proof of schematic goal statement"
      else ();
      if skip andalso not schematic then
        SkipProof (0, (finish gthy (Proof.global_skip_proof int prf), gthy))
      else Proof (ProofNode.init prf, (finish gthy, gthy))
    end
  | _ => raise UNDEF));

in

fun local_theory_to_proof' loc f = begin_proof
  (fn int => fn gthy => f int (loc_begin loc gthy))
  (loc_finish loc);

fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f);

fun theory_to_proof f = begin_proof
  (K (fn Context.Theory thy => f thy | _ => raise UNDEF))
  (K (Context.Theory o ProofContext.theory_of));

end;

val forget_proof = transaction (fn _ =>
  (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
    | SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
    | _ => raise UNDEF));

val present_proof = present_transaction (fn _ =>
  (fn Proof (prf, x) => Proof (ProofNode.apply I prf, x)
    | skip as SkipProof _ => skip
    | _ => raise UNDEF));

fun proofs' f = transaction (fn int =>
  (fn Proof (prf, x) => Proof (ProofNode.applys (f int) prf, x)
    | skip as SkipProof _ => skip
    | _ => raise UNDEF));

fun proof' f = proofs' (Seq.single oo f);
val proofs = proofs' o K;
val proof = proof' o K;

fun actual_proof f = transaction (fn _ =>
  (fn Proof (prf, x) => Proof (f prf, x)
    | _ => raise UNDEF));

fun skip_proof f = transaction (fn _ =>
  (fn SkipProof (h, x) => SkipProof (f h, x)
    | _ => raise UNDEF));

fun skip_proof_to_theory pred = transaction (fn _ =>
  (fn SkipProof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF
  | _ => raise UNDEF));



(** toplevel transactions **)

(* identification *)

fun get_id (Transition {pos, ...}) = Position.get_id pos;
fun put_id id (tr as Transition {pos, ...}) = position (Position.put_id id pos) tr;


(* thread position *)

fun setmp_thread_position (Transition {pos, ...}) f x =
  Position.setmp_thread_data pos f x;

fun status tr m =
  setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) ();

fun error_msg tr exn_info =
  setmp_thread_position tr (fn () => Output.error_msg (exn_message (EXCURSION_FAIL exn_info))) ();


(* post-transition hooks *)

local val hooks = ref ([]: (transition -> state -> state -> unit) list) in

fun add_hook f = CRITICAL (fn () => change hooks (cons f));
fun get_hooks () = CRITICAL (fn () => ! hooks);

end;


(* apply transitions *)

local

fun app int (tr as Transition {trans, pos, print, no_timing, ...}) =
  setmp_thread_position tr (fn state =>
    let
      fun do_timing f x = (warning (command_msg "" tr); timeap f x);
      fun do_profiling f x = profile (! profiling) f x;

      val (result, status) =
         state |> (apply_trans int pos trans
          |> (if ! profiling > 0 andalso not no_timing then do_profiling else I)
          |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I));

      val _ = if int andalso not (! quiet) andalso print then print_state false result else ();
    in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) status) end);

in

fun transition int tr st =
  let
    val hooks = get_hooks ();
    fun apply_hooks st' = hooks |> List.app (fn f => (try (fn () => f tr st st') (); ()));

    val ctxt = try context_of st;
    val res =
      (case app int tr st of
        (_, SOME TERMINATE) => NONE
      | (st', SOME (EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info)
      | (st', SOME exn) => SOME (st', SOME (exn_context ctxt exn, at_command tr))
      | (st', NONE) => SOME (st', NONE));
    val _ = (case res of SOME (st', NONE) => apply_hooks st' | _ => ());
  in res end;

end;


(* commit final exit *)

fun commit_exit pos =
  name "end" empty
  |> position pos
  |> add_trans CommitExit
  |> imperative (fn () => warning "Expected to find finished theory");


(* nested commands *)

fun command tr st =
  (case transition (! interact) tr st of
    SOME (st', NONE) => st'
  | SOME (_, SOME exn_info) => raise EXCURSION_FAIL exn_info
  | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));

fun command_result tr st =
  let val st' = command tr st
  in (st', st') end;

fun run_command thy_name tr st =
  (case
      (case init_of tr of
        SOME name => Exn.capture (fn () => ThyLoad.check_name thy_name name) ()
      | NONE => Exn.Result ()) of
    Exn.Result () =>
      (case transition true tr st of
        SOME (st', NONE) => (status tr Markup.finished; SOME st')
      | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE)
      | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE))
  | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE));


(* excursion of units, consisting of commands with proof *)

structure States = ProofDataFun
(
  type T = state list future option;
  fun init _ = NONE;
);

fun proof_result immediate (tr, proof_trs) st =
  let val st' = command tr st in
    if immediate orelse null proof_trs orelse
      not (can proof_of st') orelse Proof.is_relevant (proof_of st')
    then
      let val (states, st'') = fold_map command_result proof_trs st'
      in (Lazy.value ((tr, st') :: (proof_trs ~~ states)), st'') end
    else
      let
        val (body_trs, end_tr) = split_last proof_trs;
        val finish = Context.Theory o ProofContext.theory_of;

        val future_proof = Proof.global_future_proof
          (fn prf =>
            Future.fork_pri ~1 (fn () =>
              let val (states, result_state) =
                (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
                  => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev))
                |> fold_map command_result body_trs
                ||> command (end_tr |> set_print false);
              in (states, presentation_context_of result_state) end))
          #> (fn (states, ctxt) => States.put (SOME states) ctxt);

        val st'' = st' |> command (end_tr |> reset_trans |> end_proof (K future_proof));

        val states =
          (case States.get (presentation_context_of st'') of
            NONE => sys_error ("No future states for " ^ name_of tr ^ Position.str_of (pos_of tr))
          | SOME states => states);
        val result = Lazy.lazy
          (fn () => (tr, st') :: (body_trs ~~ Future.join states) @ [(end_tr, st'')]);

      in (result, st'') end
  end;

fun excursion input =
  let
    val end_pos = if null input then error "No input" else pos_of (fst (List.last input));

    val immediate = not (Goal.future_enabled ());
    val (results, end_state) = fold_map (proof_result immediate) input toplevel;
    val _ =
      (case end_state of
        State (NONE, SOME (Theory (Context.Theory _, _), _)) =>
          command (commit_exit end_pos) end_state
      | _ => error "Unfinished development at end of input");
  in Lazy.lazy (fn () => maps Lazy.force results) end;

end;