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