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