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