wenzelm@5828: (* Title: Pure/Isar/toplevel.ML wenzelm@5828: Author: Markus Wenzel, TU Muenchen wenzelm@5828: wenzelm@26602: Isabelle/Isar toplevel transactions. wenzelm@5828: *) wenzelm@5828: wenzelm@5828: signature TOPLEVEL = wenzelm@5828: sig wenzelm@19063: exception UNDEF wenzelm@5828: type state wenzelm@26602: val toplevel: state wenzelm@7732: val is_toplevel: state -> bool wenzelm@18589: val is_theory: state -> bool wenzelm@18589: val is_proof: state -> bool wenzelm@51555: val is_skipped_proof: state -> bool wenzelm@17076: val level: state -> int wenzelm@30398: val presentation_context_of: state -> Proof.context wenzelm@30801: val previous_context_of: state -> Proof.context option wenzelm@21506: val context_of: state -> Proof.context wenzelm@22089: val generic_theory_of: state -> generic_theory wenzelm@5828: val theory_of: state -> theory wenzelm@5828: val proof_of: state -> Proof.state wenzelm@18589: val proof_position_of: state -> int wenzelm@37953: val end_theory: Position.T -> state -> theory wenzelm@16815: val print_state_context: state -> unit wenzelm@16815: val print_state: bool -> state -> unit wenzelm@37858: val pretty_abstract: state -> Pretty.T wenzelm@32738: val quiet: bool Unsynchronized.ref wenzelm@32738: val debug: bool Unsynchronized.ref wenzelm@32738: val interact: bool Unsynchronized.ref wenzelm@32738: val timing: bool Unsynchronized.ref wenzelm@32738: val profiling: int Unsynchronized.ref wenzelm@20128: val program: (unit -> 'a) -> 'a wenzelm@33604: val thread: bool -> (unit -> unit) -> Thread.thread wenzelm@16682: type transition wenzelm@5828: val empty: transition wenzelm@38888: val print_of: transition -> bool wenzelm@27427: val name_of: transition -> string wenzelm@28105: val pos_of: transition -> Position.T wenzelm@5828: val name: string -> transition -> transition wenzelm@5828: val position: Position.T -> transition -> transition wenzelm@5828: val interactive: bool -> transition -> transition wenzelm@38888: val set_print: bool -> transition -> transition wenzelm@5828: val print: transition -> transition wenzelm@44187: val init_theory: (unit -> theory) -> transition -> transition wenzelm@44187: val is_init: transition -> bool wenzelm@44186: val modify_init: (unit -> theory) -> transition -> transition wenzelm@6689: val exit: transition -> transition wenzelm@5828: val keep: (state -> unit) -> transition -> transition wenzelm@7612: val keep': (bool -> state -> unit) -> transition -> transition wenzelm@5828: val imperative: (unit -> unit) -> transition -> transition wenzelm@27840: val ignored: Position.T -> transition wenzelm@51268: val is_ignored: transition -> bool wenzelm@27840: val malformed: Position.T -> string -> transition wenzelm@48772: val is_malformed: transition -> bool wenzelm@26491: val generic_theory: (generic_theory -> generic_theory) -> transition -> transition wenzelm@7612: val theory': (bool -> theory -> theory) -> transition -> transition wenzelm@49012: val theory: (theory -> theory) -> transition -> transition wenzelm@20985: val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition wenzelm@21007: val end_local_theory: transition -> transition wenzelm@47069: val open_target: (generic_theory -> local_theory) -> transition -> transition wenzelm@47069: val close_target: transition -> transition wenzelm@45488: val local_theory': (xstring * Position.T) option -> (bool -> local_theory -> local_theory) -> wenzelm@45488: transition -> transition wenzelm@45488: val local_theory: (xstring * Position.T) option -> (local_theory -> local_theory) -> wenzelm@29380: transition -> transition wenzelm@45488: val present_local_theory: (xstring * Position.T) option -> (state -> unit) -> berghofe@24453: transition -> transition wenzelm@45488: val local_theory_to_proof': (xstring * Position.T) option -> wenzelm@45488: (bool -> local_theory -> Proof.state) -> transition -> transition wenzelm@45488: val local_theory_to_proof: (xstring * Position.T) option -> wenzelm@45488: (local_theory -> Proof.state) -> transition -> transition wenzelm@17363: val theory_to_proof: (theory -> Proof.state) -> transition -> transition wenzelm@21007: val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition wenzelm@21007: val forget_proof: transition -> transition wenzelm@30366: val present_proof: (state -> unit) -> transition -> transition wenzelm@49863: val proofs': (bool -> Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition wenzelm@17904: val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition wenzelm@49863: val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition wenzelm@21177: val proof: (Proof.state -> Proof.state) -> transition -> transition wenzelm@33390: val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition wenzelm@27564: val skip_proof: (int -> int) -> transition -> transition wenzelm@17904: val skip_proof_to_theory: (int -> bool) -> transition -> transition wenzelm@52536: val exec_id: Document_ID.exec -> transition -> transition wenzelm@9512: val unknown_theory: transition -> transition wenzelm@9512: val unknown_proof: transition -> transition wenzelm@9512: val unknown_context: transition -> transition wenzelm@28425: val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b wenzelm@27606: val status: transition -> Markup.T -> unit wenzelm@28103: val add_hook: (transition -> state -> state -> unit) -> unit wenzelm@51423: val get_timing: transition -> Time.time option wenzelm@51423: val put_timing: Time.time option -> transition -> transition wenzelm@26602: val transition: bool -> transition -> state -> (state * (exn * string) option) option wenzelm@51323: val command_errors: bool -> transition -> state -> Runtime.error list * state option wenzelm@51284: val command_exception: bool -> transition -> state -> state wenzelm@51332: type result wenzelm@51332: val join_results: result -> (transition * state) list wenzelm@51332: val element_result: transition Thy_Syntax.element -> state -> result * state wenzelm@5828: end; wenzelm@5828: wenzelm@6965: structure Toplevel: TOPLEVEL = wenzelm@5828: struct wenzelm@5828: wenzelm@5828: (** toplevel state **) wenzelm@5828: wenzelm@31476: exception UNDEF = Runtime.UNDEF; wenzelm@19063: wenzelm@19063: wenzelm@21294: (* local theory wrappers *) wenzelm@5828: haftmann@38350: val loc_init = Named_Target.context_cmd; haftmann@52118: val loc_exit = Local_Theory.assert_bottom true #> Local_Theory.exit_global; wenzelm@21294: wenzelm@47274: fun loc_begin loc (Context.Theory thy) = wenzelm@47274: (Context.Theory o loc_exit, loc_init (the_default ("-", Position.none) loc) thy) wenzelm@47274: | loc_begin NONE (Context.Proof lthy) = wenzelm@47274: (Context.Proof o Local_Theory.restore, lthy) wenzelm@47274: | loc_begin (SOME loc) (Context.Proof lthy) = haftmann@52118: (Context.Proof o Named_Target.reinit lthy, loc_init loc (loc_exit (Local_Theory.assert_nonbrittle lthy))); wenzelm@21294: wenzelm@21294: wenzelm@21958: (* datatype node *) wenzelm@21294: wenzelm@5828: datatype node = wenzelm@27576: Theory of generic_theory * Proof.context option wenzelm@27576: (*theory with presentation context*) | wenzelm@33390: Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory) wenzelm@27576: (*proof node, finish, original theory*) | wenzelm@51555: Skipped_Proof of int * (generic_theory * generic_theory); wenzelm@27564: (*proof depth, resulting theory, original theory*) wenzelm@5828: wenzelm@20963: val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE; wenzelm@18589: val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; wenzelm@51555: val skipped_proof_node = fn Skipped_Proof _ => true | _ => false; wenzelm@18589: wenzelm@20963: fun cases_node f _ (Theory (gthy, _)) = f gthy wenzelm@33390: | cases_node _ g (Proof (prf, _)) = g (Proof_Node.current prf) wenzelm@51555: | cases_node f _ (Skipped_Proof (_, (gthy, _))) = f gthy; wenzelm@19063: wenzelm@29066: val context_node = cases_node Context.proof_of Proof.context_of; wenzelm@29066: wenzelm@21958: wenzelm@21958: (* datatype state *) wenzelm@21958: wenzelm@37953: datatype state = State of node option * node option; (*current, previous*) wenzelm@5828: wenzelm@27576: val toplevel = State (NONE, NONE); wenzelm@5828: wenzelm@27576: fun is_toplevel (State (NONE, _)) = true wenzelm@7732: | is_toplevel _ = false; wenzelm@7732: wenzelm@27576: fun level (State (NONE, _)) = 0 wenzelm@37953: | level (State (SOME (Theory _), _)) = 0 wenzelm@37953: | level (State (SOME (Proof (prf, _)), _)) = Proof.level (Proof_Node.current prf) wenzelm@51555: | level (State (SOME (Skipped_Proof (d, _)), _)) = d + 1; (*different notion of proof depth!*) wenzelm@17076: wenzelm@27576: fun str_of_state (State (NONE, _)) = "at top level" wenzelm@37953: | str_of_state (State (SOME (Theory (Context.Theory _, _)), _)) = "in theory mode" wenzelm@37953: | str_of_state (State (SOME (Theory (Context.Proof _, _)), _)) = "in local theory mode" wenzelm@37953: | str_of_state (State (SOME (Proof _), _)) = "in proof mode" wenzelm@51555: | str_of_state (State (SOME (Skipped_Proof _), _)) = "in skipped proof mode"; wenzelm@5946: wenzelm@5946: wenzelm@27576: (* current node *) wenzelm@5828: wenzelm@27576: fun node_of (State (NONE, _)) = raise UNDEF wenzelm@37953: | node_of (State (SOME node, _)) = node; wenzelm@5828: wenzelm@18589: fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state)); wenzelm@18589: fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state)); wenzelm@51555: fun is_skipped_proof state = not (is_toplevel state) andalso skipped_proof_node (node_of state); wenzelm@18589: wenzelm@19063: fun node_case f g state = cases_node f g (node_of state); wenzelm@5828: wenzelm@30398: fun presentation_context_of state = wenzelm@30398: (case try node_of state of wenzelm@30398: SOME (Theory (_, SOME ctxt)) => ctxt wenzelm@30398: | SOME node => context_node node wenzelm@30398: | NONE => raise UNDEF); wenzelm@30366: wenzelm@30801: fun previous_context_of (State (_, NONE)) = NONE wenzelm@37953: | previous_context_of (State (_, SOME prev)) = SOME (context_node prev); wenzelm@30801: wenzelm@21506: val context_of = node_case Context.proof_of Proof.context_of; wenzelm@22089: val generic_theory_of = node_case I (Context.Proof o Proof.context_of); wenzelm@20963: val theory_of = node_case Context.theory_of Proof.theory_of; wenzelm@18589: val proof_of = node_case (fn _ => raise UNDEF) I; wenzelm@17208: wenzelm@18589: fun proof_position_of state = wenzelm@18589: (case node_of state of wenzelm@33390: Proof (prf, _) => Proof_Node.position prf wenzelm@18589: | _ => raise UNDEF); wenzelm@6664: wenzelm@43667: fun end_theory _ (State (NONE, SOME (Theory (Context.Theory thy, _)))) = thy wenzelm@48992: | end_theory pos (State (NONE, _)) = error ("Bad theory" ^ Position.here pos) wenzelm@48992: | end_theory pos (State (SOME _, _)) = error ("Unfinished theory" ^ Position.here pos); wenzelm@37953: wenzelm@5828: wenzelm@16815: (* print state *) wenzelm@16815: haftmann@38388: val pretty_context = Local_Theory.pretty o Context.cases (Named_Target.theory_init) I; wenzelm@16815: wenzelm@23640: fun print_state_context state = wenzelm@24795: (case try node_of state of wenzelm@21506: NONE => [] wenzelm@24795: | SOME (Theory (gthy, _)) => pretty_context gthy wenzelm@24795: | SOME (Proof (_, (_, gthy))) => pretty_context gthy wenzelm@51555: | SOME (Skipped_Proof (_, (gthy, _))) => pretty_context gthy) wenzelm@23640: |> Pretty.chunks |> Pretty.writeln; wenzelm@16815: wenzelm@23640: fun print_state prf_only state = wenzelm@23701: (case try node_of state of wenzelm@23701: NONE => [] wenzelm@23701: | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy wenzelm@23701: | SOME (Proof (prf, _)) => wenzelm@33390: Proof.pretty_state (Proof_Node.position prf) (Proof_Node.current prf) wenzelm@51555: | SOME (Skipped_Proof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]) wenzelm@50201: |> Pretty.markup_chunks Markup.state |> Pretty.writeln; wenzelm@16815: wenzelm@37858: fun pretty_abstract state = Pretty.str (""); wenzelm@37858: wenzelm@16815: wenzelm@15668: wenzelm@5828: (** toplevel transitions **) wenzelm@5828: wenzelm@32738: val quiet = Unsynchronized.ref false; wenzelm@39513: val debug = Runtime.debug; wenzelm@32738: val interact = Unsynchronized.ref false; wenzelm@42012: val timing = Unsynchronized.ref false; wenzelm@32738: val profiling = Unsynchronized.ref 0; wenzelm@16682: wenzelm@33604: fun program body = wenzelm@33604: (body wenzelm@31476: |> Runtime.controlled_execution wenzelm@33604: |> Runtime.toplevel_error (Output.error_msg o ML_Compiler.exn_message)) (); wenzelm@33604: wenzelm@33604: fun thread interrupts body = wenzelm@33604: Thread.fork wenzelm@39232: (((fn () => body () handle exn => if Exn.is_interrupt exn then () else reraise exn) wenzelm@33604: |> Runtime.debugging wenzelm@33604: |> Runtime.toplevel_error wenzelm@40132: (fn exn => wenzelm@40132: Output.urgent_message ("## INTERNAL ERROR ##\n" ^ ML_Compiler.exn_message exn))), wenzelm@37216: Simple_Thread.attributes interrupts); wenzelm@20128: wenzelm@5828: wenzelm@27601: (* node transactions -- maintaining stable checkpoints *) wenzelm@7022: wenzelm@31476: exception FAILURE of state * exn; wenzelm@31476: wenzelm@6689: local wenzelm@6689: wenzelm@30366: fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE) wenzelm@30366: | reset_presentation node = node; wenzelm@30366: wenzelm@26624: fun is_draft_theory (Theory (gthy, _)) = Context.is_draft (Context.theory_of gthy) wenzelm@26624: | is_draft_theory _ = false; wenzelm@26624: wenzelm@31476: fun is_stale state = Context.is_stale (theory_of state) handle Runtime.UNDEF => false; wenzelm@27601: wenzelm@26624: fun stale_error NONE = SOME (ERROR "Stale theory encountered after successful execution!") wenzelm@26624: | stale_error some = some; wenzelm@16815: wenzelm@27576: fun map_theory f (Theory (gthy, ctxt)) = wenzelm@33671: Theory (Context.mapping f (Local_Theory.raw_theory f) gthy, ctxt) wenzelm@27576: | map_theory _ node = node; wenzelm@6689: wenzelm@6689: in wenzelm@6689: wenzelm@37953: fun apply_transaction f g node = wenzelm@20128: let wenzelm@27576: val _ = is_draft_theory node andalso error "Illegal draft theory in toplevel state"; wenzelm@27576: val cont_node = reset_presentation node; wenzelm@27576: val back_node = map_theory (Theory.checkpoint o Theory.copy) cont_node; wenzelm@37953: fun state_error e nd = (State (SOME nd, SOME node), e); wenzelm@6689: wenzelm@20128: val (result, err) = wenzelm@20128: cont_node wenzelm@31476: |> Runtime.controlled_execution f wenzelm@26624: |> map_theory Theory.checkpoint wenzelm@26624: |> state_error NONE wenzelm@26624: handle exn => state_error (SOME exn) cont_node; wenzelm@26624: wenzelm@26624: val (result', err') = wenzelm@26624: if is_stale result then state_error (stale_error err) back_node wenzelm@26624: else (result, err); wenzelm@20128: in wenzelm@26624: (case err' of wenzelm@30366: NONE => tap g result' wenzelm@26624: | SOME exn => raise FAILURE (result', exn)) wenzelm@20128: end; wenzelm@6689: wenzelm@43667: val exit_transaction = wenzelm@43667: apply_transaction wenzelm@43667: (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (Theory.end_theory thy), NONE) wenzelm@43667: | node => node) (K ()) wenzelm@43667: #> (fn State (node', _) => State (NONE, node')); wenzelm@43667: wenzelm@6689: end; wenzelm@6689: wenzelm@6689: wenzelm@6689: (* primitive transitions *) wenzelm@6689: wenzelm@5828: datatype trans = wenzelm@44187: Init of unit -> theory | (*init theory*) wenzelm@37953: Exit | (*formal exit of theory*) wenzelm@37953: Keep of bool -> state -> unit | (*peek at state*) wenzelm@30366: Transaction of (bool -> node -> node) * (state -> unit); (*node transaction and presentation*) wenzelm@21958: wenzelm@6689: local wenzelm@5828: wenzelm@44187: fun apply_tr _ (Init f) (State (NONE, _)) = wenzelm@33727: State (SOME (Theory (Context.Theory wenzelm@44186: (Theory.checkpoint (Runtime.controlled_execution f ())), NONE)), NONE) wenzelm@43667: | apply_tr _ Exit (State (SOME (state as Theory (Context.Theory _, _)), _)) = wenzelm@43667: exit_transaction state wenzelm@32792: | apply_tr int (Keep f) state = wenzelm@31476: Runtime.controlled_execution (fn x => tap (f int) x) state wenzelm@32792: | apply_tr int (Transaction (f, g)) (State (SOME state, _)) = wenzelm@32792: apply_transaction (fn x => f int x) g state wenzelm@32792: | apply_tr _ _ _ = raise UNDEF; wenzelm@5828: wenzelm@32792: fun apply_union _ [] state = raise FAILURE (state, UNDEF) wenzelm@32792: | apply_union int (tr :: trs) state = wenzelm@32792: apply_union int trs state wenzelm@32792: handle Runtime.UNDEF => apply_tr int tr state wenzelm@32792: | FAILURE (alt_state, UNDEF) => apply_tr int tr alt_state wenzelm@6689: | exn as FAILURE _ => raise exn wenzelm@6689: | exn => raise FAILURE (state, exn); wenzelm@6689: wenzelm@6689: in wenzelm@6689: wenzelm@32792: fun apply_trans int trs state = (apply_union int trs state, NONE) skalberg@15531: handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); wenzelm@6689: wenzelm@6689: end; wenzelm@5828: wenzelm@5828: wenzelm@5828: (* datatype transition *) wenzelm@5828: wenzelm@5828: datatype transition = Transition of wenzelm@26621: {name: string, (*command name*) wenzelm@26621: pos: Position.T, (*source position*) wenzelm@26621: int_only: bool, (*interactive-only*) wenzelm@26621: print: bool, (*print result state*) wenzelm@51423: timing: Time.time option, (*prescient timing information*) wenzelm@26621: trans: trans list}; (*primitive transitions (union)*) wenzelm@5828: wenzelm@51658: fun make_transition (name, pos, int_only, print, timing, trans) = wenzelm@51217: Transition {name = name, pos = pos, int_only = int_only, print = print, wenzelm@51658: timing = timing, trans = trans}; wenzelm@5828: wenzelm@51658: fun map_transition f (Transition {name, pos, int_only, print, timing, trans}) = wenzelm@51658: make_transition (f (name, pos, int_only, print, timing, trans)); wenzelm@5828: wenzelm@51658: val empty = make_transition ("", Position.none, false, false, NONE, []); wenzelm@5828: wenzelm@5828: wenzelm@5828: (* diagnostics *) wenzelm@5828: wenzelm@38888: fun print_of (Transition {print, ...}) = print; wenzelm@27427: fun name_of (Transition {name, ...}) = name; wenzelm@28105: fun pos_of (Transition {pos, ...}) = pos; wenzelm@5828: wenzelm@48993: fun command_msg msg tr = msg ^ "command " ^ quote (name_of tr) ^ Position.here (pos_of tr); wenzelm@38875: fun at_command tr = command_msg "At " tr; wenzelm@5828: wenzelm@5828: fun type_error tr state = wenzelm@18685: ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state); wenzelm@5828: wenzelm@5828: wenzelm@5828: (* modify transitions *) wenzelm@5828: wenzelm@51658: fun name name = map_transition (fn (_, pos, int_only, print, timing, trans) => wenzelm@51658: (name, pos, int_only, print, timing, trans)); wenzelm@9010: wenzelm@51658: fun position pos = map_transition (fn (name, _, int_only, print, timing, trans) => wenzelm@51658: (name, pos, int_only, print, timing, trans)); wenzelm@5828: wenzelm@51658: fun interactive int_only = map_transition (fn (name, pos, _, print, timing, trans) => wenzelm@51658: (name, pos, int_only, print, timing, trans)); wenzelm@14923: wenzelm@51658: fun add_trans tr = map_transition (fn (name, pos, int_only, print, timing, trans) => wenzelm@51658: (name, pos, int_only, print, timing, tr :: trans)); wenzelm@16607: wenzelm@51658: val reset_trans = map_transition (fn (name, pos, int_only, print, timing, _) => wenzelm@51658: (name, pos, int_only, print, timing, [])); wenzelm@28433: wenzelm@51658: fun set_print print = map_transition (fn (name, pos, int_only, _, timing, trans) => wenzelm@51658: (name, pos, int_only, print, timing, trans)); wenzelm@28453: wenzelm@28453: val print = set_print true; wenzelm@5828: wenzelm@5828: wenzelm@21007: (* basic transitions *) wenzelm@5828: wenzelm@44187: fun init_theory f = add_trans (Init f); wenzelm@37977: wenzelm@44187: fun is_init (Transition {trans = [Init _], ...}) = true wenzelm@44187: | is_init _ = false; wenzelm@44187: wenzelm@44187: fun modify_init f tr = if is_init tr then init_theory f (reset_trans tr) else tr; wenzelm@37977: wenzelm@6689: val exit = add_trans Exit; wenzelm@7612: val keep' = add_trans o Keep; wenzelm@30366: wenzelm@30366: fun present_transaction f g = add_trans (Transaction (f, g)); wenzelm@30366: fun transaction f = present_transaction f (K ()); wenzelm@5828: wenzelm@7612: fun keep f = add_trans (Keep (fn _ => f)); wenzelm@5828: fun imperative f = keep (fn _ => f ()); wenzelm@5828: wenzelm@27840: fun ignored pos = empty |> name "" |> position pos |> imperative I; wenzelm@51268: fun is_ignored tr = name_of tr = ""; wenzelm@48772: wenzelm@48772: val malformed_name = ""; wenzelm@27840: fun malformed pos msg = wenzelm@48772: empty |> name malformed_name |> position pos |> imperative (fn () => error msg); wenzelm@48772: fun is_malformed tr = name_of tr = malformed_name; wenzelm@27840: wenzelm@21007: val unknown_theory = imperative (fn () => warning "Unknown theory context"); wenzelm@21007: val unknown_proof = imperative (fn () => warning "Unknown proof context"); wenzelm@21007: val unknown_context = imperative (fn () => warning "Unknown context"); wenzelm@15668: wenzelm@21007: wenzelm@49012: (* theory transitions *) wenzelm@44304: wenzelm@27601: fun generic_theory f = transaction (fn _ => wenzelm@26491: (fn Theory (gthy, _) => Theory (f gthy, NONE) wenzelm@26491: | _ => raise UNDEF)); wenzelm@26491: wenzelm@27601: fun theory' f = transaction (fn int => wenzelm@33725: (fn Theory (Context.Theory thy, _) => wenzelm@33725: let val thy' = thy wenzelm@49012: |> Sign.new_group wenzelm@49012: |> Theory.checkpoint wenzelm@33725: |> f int wenzelm@33725: |> Sign.reset_group; wenzelm@33725: in Theory (Context.Theory thy', NONE) end wenzelm@20963: | _ => raise UNDEF)); wenzelm@20963: wenzelm@20963: fun theory f = theory' (K f); wenzelm@20963: wenzelm@27601: fun begin_local_theory begin f = transaction (fn _ => wenzelm@20963: (fn Theory (Context.Theory thy, _) => wenzelm@20963: let wenzelm@20985: val lthy = f thy; wenzelm@21294: val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy); wenzelm@21294: in Theory (gthy, SOME lthy) end wenzelm@20963: | _ => raise UNDEF)); wenzelm@17076: wenzelm@27601: val end_local_theory = transaction (fn _ => wenzelm@21294: (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy) wenzelm@21007: | _ => raise UNDEF)); wenzelm@21007: wenzelm@47069: fun open_target f = transaction (fn _ => wenzelm@47069: (fn Theory (gthy, _) => wenzelm@47069: let val lthy = f gthy wenzelm@47069: in Theory (Context.Proof lthy, SOME lthy) end wenzelm@47069: | _ => raise UNDEF)); wenzelm@47069: wenzelm@47069: val close_target = transaction (fn _ => wenzelm@47069: (fn Theory (Context.Proof lthy, _) => wenzelm@47069: (case try Local_Theory.close_target lthy of wenzelm@50739: SOME ctxt' => wenzelm@50739: let wenzelm@50739: val gthy' = wenzelm@50739: if can Local_Theory.assert ctxt' wenzelm@50739: then Context.Proof ctxt' wenzelm@50739: else Context.Theory (Proof_Context.theory_of ctxt'); wenzelm@50739: in Theory (gthy', SOME lthy) end wenzelm@47069: | NONE => raise UNDEF) wenzelm@47069: | _ => raise UNDEF)); wenzelm@47069: wenzelm@47069: wenzelm@21007: local wenzelm@21007: wenzelm@30366: fun local_theory_presentation loc f = present_transaction (fn int => wenzelm@21294: (fn Theory (gthy, _) => wenzelm@21294: let wenzelm@49062: val (finish, lthy) = loc_begin loc gthy; wenzelm@47274: val lthy' = lthy wenzelm@49012: |> Local_Theory.new_group wenzelm@33725: |> f int wenzelm@33725: |> Local_Theory.reset_group; wenzelm@21294: in Theory (finish lthy', SOME lthy') end wenzelm@30366: | _ => raise UNDEF)); wenzelm@15668: wenzelm@21007: in wenzelm@21007: wenzelm@30366: fun local_theory' loc f = local_theory_presentation loc f (K ()); wenzelm@29380: fun local_theory loc f = local_theory' loc (K f); wenzelm@30366: fun present_local_theory loc = local_theory_presentation loc (K I); wenzelm@18955: wenzelm@21007: end; wenzelm@21007: wenzelm@21007: wenzelm@21007: (* proof transitions *) wenzelm@21007: wenzelm@27601: fun end_proof f = transaction (fn int => wenzelm@24795: (fn Proof (prf, (finish, _)) => wenzelm@33390: let val state = Proof_Node.current prf in wenzelm@21007: if can (Proof.assert_bottom true) state then wenzelm@21007: let wenzelm@21007: val ctxt' = f int state; wenzelm@21007: val gthy' = finish ctxt'; wenzelm@21007: in Theory (gthy', SOME ctxt') end wenzelm@21007: else raise UNDEF wenzelm@21007: end wenzelm@51555: | Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE) wenzelm@21007: | _ => raise UNDEF)); wenzelm@21007: wenzelm@21294: local wenzelm@21294: wenzelm@47274: fun begin_proof init = transaction (fn int => wenzelm@21294: (fn Theory (gthy, _) => wenzelm@21294: let wenzelm@47274: val (finish, prf) = init int gthy; wenzelm@52499: val skip = Goal.skip_proofs_enabled (); wenzelm@40960: val (is_goal, no_skip) = wenzelm@40960: (true, Proof.schematic_goal prf) handle ERROR _ => (false, true); wenzelm@47274: val _ = wenzelm@47274: if is_goal andalso skip andalso no_skip then wenzelm@47274: warning "Cannot skip proof of schematic goal statement" wenzelm@47274: else (); wenzelm@21294: in wenzelm@40960: if skip andalso not no_skip then wenzelm@51555: Skipped_Proof (0, (finish (Proof.global_skip_proof true prf), gthy)) wenzelm@47274: else Proof (Proof_Node.init prf, (finish, gthy)) wenzelm@21294: end wenzelm@21294: | _ => raise UNDEF)); wenzelm@21294: wenzelm@21294: in wenzelm@21294: wenzelm@24780: fun local_theory_to_proof' loc f = begin_proof wenzelm@47274: (fn int => fn gthy => wenzelm@49062: let val (finish, lthy) = loc_begin loc gthy wenzelm@49012: in (finish o Local_Theory.reset_group, f int (Local_Theory.new_group lthy)) end); wenzelm@24780: berghofe@24453: fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f); wenzelm@21294: wenzelm@21294: fun theory_to_proof f = begin_proof wenzelm@47274: (fn _ => fn gthy => wenzelm@51560: (Context.Theory o Theory.checkpoint o Sign.reset_group o Proof_Context.theory_of, wenzelm@49062: (case gthy of wenzelm@49012: Context.Theory thy => f (Theory.checkpoint (Sign.new_group thy)) wenzelm@49012: | _ => raise UNDEF))); wenzelm@21294: wenzelm@21294: end; wenzelm@21294: wenzelm@27601: val forget_proof = transaction (fn _ => wenzelm@21007: (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) wenzelm@51555: | Skipped_Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) wenzelm@21007: | _ => raise UNDEF)); wenzelm@21007: wenzelm@30366: val present_proof = present_transaction (fn _ => wenzelm@49062: (fn Proof (prf, x) => Proof (Proof_Node.apply I prf, x) wenzelm@51555: | skip as Skipped_Proof _ => skip wenzelm@30366: | _ => raise UNDEF)); wenzelm@21177: wenzelm@27601: fun proofs' f = transaction (fn int => wenzelm@49062: (fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x) wenzelm@51555: | skip as Skipped_Proof _ => skip wenzelm@16815: | _ => raise UNDEF)); wenzelm@15668: wenzelm@49863: fun proof' f = proofs' ((Seq.single o Seq.Result) oo f); wenzelm@17904: val proofs = proofs' o K; wenzelm@6689: val proof = proof' o K; wenzelm@16815: wenzelm@27601: fun actual_proof f = transaction (fn _ => wenzelm@21007: (fn Proof (prf, x) => Proof (f prf, x) wenzelm@20963: | _ => raise UNDEF)); wenzelm@16815: wenzelm@27601: fun skip_proof f = transaction (fn _ => wenzelm@51555: (fn Skipped_Proof (h, x) => Skipped_Proof (f h, x) wenzelm@18563: | _ => raise UNDEF)); wenzelm@18563: wenzelm@27601: fun skip_proof_to_theory pred = transaction (fn _ => wenzelm@51555: (fn Skipped_Proof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF wenzelm@33725: | _ => raise UNDEF)); wenzelm@5828: wenzelm@5828: wenzelm@5828: wenzelm@5828: (** toplevel transactions **) wenzelm@5828: wenzelm@52527: (* runtime position *) wenzelm@27427: wenzelm@52536: fun exec_id id (tr as Transition {pos, ...}) = wenzelm@52536: position (Position.put_id (Document_ID.print id) pos) tr; wenzelm@25799: wenzelm@25960: fun setmp_thread_position (Transition {pos, ...}) f x = wenzelm@25819: Position.setmp_thread_data pos f x; wenzelm@25799: wenzelm@27606: fun status tr m = wenzelm@43665: setmp_thread_position tr (fn () => Output.status (Markup.markup_only m)) (); wenzelm@27606: wenzelm@25799: wenzelm@28095: (* post-transition hooks *) wenzelm@28095: wenzelm@37905: local wenzelm@37905: val hooks = Unsynchronized.ref ([]: (transition -> state -> state -> unit) list); wenzelm@37905: in wenzelm@28095: wenzelm@32738: fun add_hook f = CRITICAL (fn () => Unsynchronized.change hooks (cons f)); wenzelm@33223: fun get_hooks () = ! hooks; wenzelm@28095: wenzelm@28095: end; wenzelm@28095: wenzelm@28095: wenzelm@5828: (* apply transitions *) wenzelm@5828: wenzelm@51217: fun get_timing (Transition {timing, ...}) = timing; wenzelm@51658: fun put_timing timing = map_transition (fn (name, pos, int_only, print, _, trans) => wenzelm@51658: (name, pos, int_only, print, timing, trans)); wenzelm@51217: wenzelm@6664: local wenzelm@6664: wenzelm@51658: fun app int (tr as Transition {trans, print, ...}) = wenzelm@25819: setmp_thread_position tr (fn state => wenzelm@25799: let wenzelm@51595: val timing_start = Timing.start (); wenzelm@25799: wenzelm@51595: val (result, opt_err) = wenzelm@51658: state |> (apply_trans int trans |> ! profiling > 0 ? profile (! profiling)); wenzelm@26621: val _ = if int andalso not (! quiet) andalso print then print_state false result else (); wenzelm@51216: wenzelm@51595: val timing_result = Timing.result timing_start; wenzelm@51662: val timing_props = wenzelm@51662: Markup.command_timing :: (Markup.nameN, name_of tr) :: Position.properties_of (pos_of tr); wenzelm@51662: val _ = Timing.protocol_message timing_props timing_result; wenzelm@51595: in wenzelm@51595: (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_err) wenzelm@51595: end); wenzelm@6664: wenzelm@6664: in wenzelm@5828: wenzelm@26602: fun transition int tr st = wenzelm@28095: let wenzelm@28095: val hooks = get_hooks (); wenzelm@28103: fun apply_hooks st' = hooks |> List.app (fn f => (try (fn () => f tr st st') (); ())); wenzelm@28095: wenzelm@28095: val ctxt = try context_of st; wenzelm@28095: val res = wenzelm@28095: (case app int tr st of wenzelm@38888: (_, SOME Runtime.TERMINATE) => NONE wenzelm@38888: | (st', SOME (Runtime.EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info) wenzelm@31476: | (st', SOME exn) => SOME (st', SOME (Runtime.exn_context ctxt exn, at_command tr)) wenzelm@28103: | (st', NONE) => SOME (st', NONE)); wenzelm@28103: val _ = (case res of SOME (st', NONE) => apply_hooks st' | _ => ()); wenzelm@28095: in res end; wenzelm@6664: wenzelm@6664: end; wenzelm@5828: wenzelm@5828: wenzelm@51284: (* managed commands *) wenzelm@5828: wenzelm@51323: fun command_errors int tr st = wenzelm@51323: (case transition int tr st of wenzelm@51323: SOME (st', NONE) => ([], SOME st') wenzelm@51323: | SOME (_, SOME (exn, _)) => (ML_Compiler.exn_messages_ids exn, NONE) wenzelm@51323: | NONE => (ML_Compiler.exn_messages_ids Runtime.TERMINATE, NONE)); wenzelm@51323: wenzelm@51284: fun command_exception int tr st = wenzelm@51284: (case transition int tr st of wenzelm@28425: SOME (st', NONE) => st' wenzelm@39285: | SOME (_, SOME (exn, info)) => wenzelm@39285: if Exn.is_interrupt exn then reraise exn else raise Runtime.EXCURSION_FAIL (exn, info) wenzelm@38888: | NONE => raise Runtime.EXCURSION_FAIL (Runtime.TERMINATE, at_command tr)); wenzelm@27576: wenzelm@51323: fun command tr = command_exception (! interact) tr; wenzelm@51284: wenzelm@28433: wenzelm@46959: (* scheduled proof result *) wenzelm@28433: wenzelm@51332: datatype result = wenzelm@51332: Result of transition * state | wenzelm@51332: Result_List of result list | wenzelm@51332: Result_Future of result future; wenzelm@51332: wenzelm@51332: fun join_results (Result x) = [x] wenzelm@51332: | join_results (Result_List xs) = maps join_results xs wenzelm@51332: | join_results (Result_Future x) = join_results (Future.join x); wenzelm@51332: wenzelm@51323: local wenzelm@51323: wenzelm@47417: structure Result = Proof_Data wenzelm@28974: ( wenzelm@51332: type T = result; wenzelm@51332: val empty: T = Result_List []; wenzelm@47417: fun init _ = empty; wenzelm@28974: ); wenzelm@28974: wenzelm@51332: val get_result = Result.get o Proof.context_of; wenzelm@51332: val put_result = Proof.map_context o Result.put; wenzelm@51324: wenzelm@51423: fun timing_estimate include_head elem = wenzelm@51423: let wenzelm@51423: val trs = Thy_Syntax.flat_element elem |> not include_head ? tl; wenzelm@51423: val timings = map get_timing trs; wenzelm@51423: in wenzelm@51423: if forall is_some timings then wenzelm@51423: SOME (fold (curry Time.+ o the) timings Time.zeroTime) wenzelm@51423: else NONE wenzelm@51423: end; wenzelm@51423: wenzelm@51423: fun priority NONE = ~1 wenzelm@51423: | priority (SOME estimate) = wenzelm@51423: Int.min (Real.floor (Real.max (Math.log10 (Time.toReal estimate), ~3.0)) - 3, ~1); wenzelm@51423: wenzelm@51423: fun proof_future_enabled estimate st = wenzelm@51324: (case try proof_of st of wenzelm@51324: NONE => false wenzelm@51324: | SOME state => wenzelm@51324: not (Proof.is_relevant state) andalso wenzelm@51324: (if can (Proof.assert_bottom true) state wenzelm@51324: then Goal.future_enabled () wenzelm@51423: else wenzelm@51423: (case estimate of wenzelm@51423: NONE => Goal.future_enabled_nested 2 wenzelm@51423: | SOME t => Goal.future_enabled_timing t))); wenzelm@51278: wenzelm@51323: fun atom_result tr st = wenzelm@51323: let wenzelm@51323: val st' = wenzelm@51323: if Goal.future_enabled () andalso Keyword.is_diag (name_of tr) then wenzelm@51605: (Goal.fork_params wenzelm@51605: {name = "Toplevel.diag", pos = pos_of tr, wenzelm@51605: pri = priority (timing_estimate true (Thy_Syntax.atom tr))} wenzelm@51605: (fn () => command tr st); st) wenzelm@51323: else command tr st; wenzelm@51332: in (Result (tr, st'), st') end; wenzelm@51323: wenzelm@51323: in wenzelm@51323: wenzelm@51332: fun element_result (Thy_Syntax.Element (tr, NONE)) st = atom_result tr st wenzelm@51423: | element_result (elem as Thy_Syntax.Element (head_tr, SOME element_rest)) st = wenzelm@48633: let wenzelm@51324: val (head_result, st') = atom_result head_tr st; wenzelm@51332: val (body_elems, end_tr) = element_rest; wenzelm@51423: val estimate = timing_estimate false elem; wenzelm@51324: in wenzelm@51423: if not (proof_future_enabled estimate st') wenzelm@51324: then wenzelm@51332: let wenzelm@51332: val proof_trs = maps Thy_Syntax.flat_element body_elems @ [end_tr]; wenzelm@51332: val (proof_results, st'') = fold_map atom_result proof_trs st'; wenzelm@51332: in (Result_List (head_result :: proof_results), st'') end wenzelm@51324: else wenzelm@51324: let wenzelm@51324: val finish = Context.Theory o Proof_Context.theory_of; wenzelm@28974: wenzelm@51605: val future_proof = wenzelm@51605: Proof.future_proof (fn state => wenzelm@51605: Goal.fork_params wenzelm@51605: {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = priority estimate} wenzelm@51605: (fn () => wenzelm@51605: let wenzelm@51605: val State (SOME (Proof (prf, (_, orig_gthy))), prev) = st'; wenzelm@51605: val prf' = Proof_Node.apply (K state) prf; wenzelm@51605: val (result, result_state) = wenzelm@51605: State (SOME (Proof (prf', (finish, orig_gthy))), prev) wenzelm@51605: |> fold_map element_result body_elems ||> command end_tr; wenzelm@51605: in (Result_List result, presentation_context_of result_state) end)) wenzelm@51332: #> (fn (res, state') => state' |> put_result (Result_Future res)); wenzelm@51332: wenzelm@51332: val forked_proof = wenzelm@51332: proof (future_proof #> wenzelm@51332: (fn state => state |> Proof.local_done_proof |> put_result (get_result state))) o wenzelm@51332: end_proof (fn _ => future_proof #> wenzelm@51332: (fn state => state |> Proof.global_done_proof |> Result.put (get_result state))); wenzelm@28974: wenzelm@51324: val st'' = st' wenzelm@51332: |> command (head_tr |> set_print false |> reset_trans |> forked_proof); wenzelm@51332: val end_result = Result (end_tr, st''); wenzelm@51324: val result = wenzelm@51332: Result_List [head_result, Result.get (presentation_context_of st''), end_result]; wenzelm@51324: in (result, st'') end wenzelm@51324: end; wenzelm@28433: wenzelm@6664: end; wenzelm@51323: wenzelm@51323: end;