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