wenzelm@5828: (* Title: Pure/Isar/toplevel.ML wenzelm@5828: ID: $Id$ wenzelm@5828: Author: Markus Wenzel, TU Muenchen wenzelm@5828: wenzelm@5828: The Isabelle/Isar toplevel. wenzelm@5828: *) wenzelm@5828: wenzelm@5828: signature TOPLEVEL = wenzelm@5828: sig wenzelm@19063: exception UNDEF wenzelm@20963: type generic_theory wenzelm@18589: type node wenzelm@20963: val theory_node: node -> generic_theory option wenzelm@18589: val proof_node: node -> ProofHistory.T option wenzelm@20963: val cases_node: (generic_theory -> 'a) -> (Proof.state -> 'a) -> node -> 'a wenzelm@20963: val presentation_context: node option -> xstring option -> Proof.context wenzelm@5828: type state wenzelm@7732: val is_toplevel: state -> bool wenzelm@18589: val is_theory: state -> bool wenzelm@18589: val is_proof: state -> bool wenzelm@17076: val level: state -> int wenzelm@6689: val node_history_of: state -> node History.T wenzelm@5828: val node_of: state -> node wenzelm@20963: val node_case: (generic_theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a 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@21007: val enter_proof_body: state -> Proof.state wenzelm@16815: val print_state_context: state -> unit wenzelm@16815: val print_state: bool -> state -> unit wenzelm@16682: val quiet: bool ref wenzelm@16682: val debug: bool ref wenzelm@17321: val interact: bool ref wenzelm@16682: val timing: bool ref wenzelm@16682: val profiling: int ref wenzelm@16815: val skip_proofs: bool ref wenzelm@5828: exception TERMINATE wenzelm@5990: exception RESTART wenzelm@24055: exception TOPLEVEL_ERROR wenzelm@20128: val exn_message: exn -> string wenzelm@20128: val program: (unit -> 'a) -> 'a wenzelm@16682: type transition wenzelm@6689: val undo_limit: bool -> int option wenzelm@5828: val empty: transition wenzelm@14923: val name_of: transition -> string wenzelm@14923: val source_of: transition -> OuterLex.token list option wenzelm@5828: val name: string -> transition -> transition wenzelm@5828: val position: Position.T -> transition -> transition wenzelm@14923: val source: OuterLex.token list -> transition -> transition wenzelm@5828: val interactive: bool -> transition -> transition wenzelm@5828: val print: transition -> transition wenzelm@16607: val print': string -> transition -> transition wenzelm@17363: val three_buffersN: string wenzelm@17363: val print3: transition -> transition wenzelm@9010: val no_timing: transition -> transition wenzelm@22056: val init_theory: (bool -> theory) -> (theory -> unit) -> (theory -> unit) -> wenzelm@22056: transition -> transition wenzelm@23911: val init_empty: (unit -> unit) -> transition -> transition wenzelm@6689: val exit: transition -> transition wenzelm@21958: val undo_exit: transition -> transition wenzelm@6689: val kill: transition -> transition wenzelm@20128: val history: (node History.T -> node History.T) -> 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@5828: val theory: (theory -> theory) -> transition -> transition wenzelm@7612: val theory': (bool -> 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@20963: val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition wenzelm@20963: val present_local_theory: xstring option -> (bool -> node -> unit) -> transition -> transition berghofe@24453: val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) -> berghofe@24453: transition -> transition wenzelm@21007: val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) -> wenzelm@21007: 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@21177: val present_proof: (bool -> node -> unit) -> transition -> transition wenzelm@21177: val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition wenzelm@17904: val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition wenzelm@21177: val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition wenzelm@21177: val proof: (Proof.state -> Proof.state) -> transition -> transition wenzelm@16815: val actual_proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition wenzelm@16815: val skip_proof: (int History.T -> int History.T) -> transition -> transition wenzelm@17904: val skip_proof_to_theory: (int -> bool) -> transition -> transition wenzelm@9512: val unknown_theory: transition -> transition wenzelm@9512: val unknown_proof: transition -> transition wenzelm@9512: val unknown_context: transition -> transition wenzelm@17076: val present_excursion: (transition * (state -> state -> 'a -> 'a)) list -> 'a -> 'a wenzelm@5828: val excursion: transition list -> unit wenzelm@5828: val set_state: state -> unit wenzelm@5828: val get_state: unit -> state wenzelm@5828: val exn: unit -> (exn * string) option wenzelm@5828: val >> : transition -> bool wenzelm@14985: val >>> : transition list -> unit wenzelm@21958: val init_state: unit -> unit nipkow@14091: type 'a isar nipkow@14091: val loop: 'a isar -> unit wenzelm@5828: end; wenzelm@5828: wenzelm@6965: structure Toplevel: TOPLEVEL = wenzelm@5828: struct wenzelm@5828: wenzelm@21958: wenzelm@5828: (** toplevel state **) wenzelm@5828: wenzelm@19063: exception UNDEF; wenzelm@19063: wenzelm@19063: wenzelm@21294: (* local theory wrappers *) wenzelm@5828: wenzelm@20963: type generic_theory = Context.generic; (*theory or local_theory*) wenzelm@20963: wenzelm@24936: val loc_init = TheoryTarget.init_cmd; wenzelm@21294: wenzelm@21294: val loc_exit = ProofContext.theory_of o LocalTheory.exit; wenzelm@21294: haftmann@25269: fun loc_begin NONE (Context.Theory thy) = loc_init "-" thy haftmann@25269: | loc_begin (SOME loc) (Context.Theory thy) = loc_init loc thy wenzelm@21294: | loc_begin NONE (Context.Proof lthy) = lthy haftmann@25269: | loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy); wenzelm@21294: wenzelm@21294: fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit wenzelm@21294: | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore wenzelm@21294: | loc_finish (SOME _) (Context.Proof lthy) = Context.Proof o LocalTheory.reinit lthy o loc_exit; wenzelm@21294: wenzelm@21294: wenzelm@21958: (* datatype node *) wenzelm@21294: wenzelm@5828: datatype node = wenzelm@20963: Theory of generic_theory * Proof.context option | (*theory with presentation context*) wenzelm@21007: Proof of ProofHistory.T * ((Proof.context -> generic_theory) * generic_theory) | wenzelm@21007: (*history of proof states, finish, original theory*) wenzelm@21007: SkipProof of int History.T * (generic_theory * generic_theory); wenzelm@18563: (*history of proof depths, resulting theory, original theory*) wenzelm@5828: wenzelm@22056: val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF; wenzelm@20963: val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE; wenzelm@18589: val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; wenzelm@18589: wenzelm@20963: fun cases_node f _ (Theory (gthy, _)) = f gthy wenzelm@19063: | cases_node _ g (Proof (prf, _)) = g (ProofHistory.current prf) wenzelm@21007: | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy; wenzelm@19063: wenzelm@20963: fun presentation_context (SOME (Theory (_, SOME ctxt))) NONE = ctxt wenzelm@20963: | presentation_context (SOME node) NONE = cases_node Context.proof_of Proof.context_of node wenzelm@20963: | presentation_context (SOME node) (SOME loc) = haftmann@25269: loc_init loc (cases_node Context.theory_of Proof.theory_of node) wenzelm@20963: | presentation_context NONE _ = raise UNDEF; wenzelm@19063: wenzelm@21958: wenzelm@21958: (* datatype state *) wenzelm@21958: wenzelm@22056: type state_info = node History.T * ((theory -> unit) * (theory -> unit)); wenzelm@5828: wenzelm@21958: datatype state = wenzelm@21958: Toplevel of state_info option | (*outer toplevel, leftover end state*) wenzelm@21958: State of state_info; wenzelm@5828: wenzelm@21958: val toplevel = Toplevel NONE; wenzelm@21958: wenzelm@21958: fun is_toplevel (Toplevel _) = true wenzelm@7732: | is_toplevel _ = false; wenzelm@7732: wenzelm@21958: fun level (Toplevel _) = 0 wenzelm@21958: | level (State (node, _)) = wenzelm@17076: (case History.current node of wenzelm@21310: Theory _ => 0 wenzelm@21310: | Proof (prf, _) => Proof.level (ProofHistory.current prf) wenzelm@21310: | SkipProof (h, _) => History.current h + 1); (*different notion of proof depth!*) wenzelm@17076: wenzelm@21958: fun str_of_state (Toplevel _) = "at top level" wenzelm@21958: | str_of_state (State (node, _)) = wenzelm@16815: (case History.current node of wenzelm@20963: Theory (Context.Theory _, _) => "in theory mode" wenzelm@20963: | Theory (Context.Proof _, _) => "in local theory mode" wenzelm@16815: | Proof _ => "in proof mode" wenzelm@16815: | SkipProof _ => "in skipped proof mode"); wenzelm@5946: wenzelm@5946: wenzelm@5828: (* top node *) wenzelm@5828: wenzelm@21958: fun node_history_of (Toplevel _) = raise UNDEF wenzelm@21958: | node_history_of (State (node, _)) = node; wenzelm@6689: wenzelm@6689: val node_of = History.current o node_history_of; 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@18589: wenzelm@19063: fun node_case f g state = cases_node f g (node_of state); wenzelm@5828: 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@18589: Proof (prf, _) => ProofHistory.position prf wenzelm@18589: | _ => raise UNDEF); wenzelm@6664: wenzelm@21007: val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward; wenzelm@5828: wenzelm@5828: wenzelm@16815: (* print state *) wenzelm@16815: haftmann@25269: val pretty_context = LocalTheory.pretty o Context.cases (loc_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@24795: | SOME (SkipProof (_, (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@23701: Proof.pretty_state (ProofHistory.position prf) (ProofHistory.current prf) wenzelm@23701: | SOME (SkipProof (h, _)) => wenzelm@23701: [Pretty.str ("skipped proof: depth " ^ string_of_int (History.current h))]) wenzelm@23701: |> Pretty.markup_chunks Markup.state |> Pretty.writeln; wenzelm@16815: wenzelm@16815: wenzelm@15668: wenzelm@5828: (** toplevel transitions **) wenzelm@5828: wenzelm@16682: val quiet = ref false; wenzelm@22135: val debug = Output.debugging; wenzelm@17321: val interact = ref false; wenzelm@16682: val timing = Output.timing; wenzelm@16682: val profiling = ref 0; wenzelm@16815: val skip_proofs = ref false; wenzelm@16682: wenzelm@5828: exception TERMINATE; wenzelm@5990: exception RESTART; wenzelm@7022: exception EXCURSION_FAIL of exn * string; wenzelm@6689: exception FAILURE of state * exn; wenzelm@24055: exception TOPLEVEL_ERROR; wenzelm@5828: wenzelm@20128: wenzelm@20128: (* print exceptions *) wenzelm@20128: wenzelm@20128: local wenzelm@20128: wenzelm@20128: fun with_context f xs = wenzelm@22095: (case ML_Context.get_context () of NONE => [] wenzelm@22089: | SOME context => map (f (Context.proof_of context)) xs); wenzelm@20128: wenzelm@20128: fun raised name [] = "exception " ^ name ^ " raised" wenzelm@20128: | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg wenzelm@20128: | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs); wenzelm@20128: wenzelm@20128: fun exn_msg _ TERMINATE = "Exit." wenzelm@20128: | exn_msg _ RESTART = "Restart." wenzelm@20128: | exn_msg _ Interrupt = "Interrupt." wenzelm@24055: | exn_msg _ TOPLEVEL_ERROR = "Error." wenzelm@20128: | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg wenzelm@20128: | exn_msg _ (ERROR msg) = msg wenzelm@23975: | exn_msg detailed (Exn.EXCEPTIONS (exns, "")) = cat_lines (map (exn_msg detailed) exns) wenzelm@23963: | exn_msg detailed (Exn.EXCEPTIONS (exns, msg)) = cat_lines (map (exn_msg detailed) exns @ [msg]) wenzelm@20128: | exn_msg detailed (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_msg detailed exn, msg] wenzelm@20128: | exn_msg false (THEORY (msg, _)) = msg wenzelm@20128: | exn_msg true (THEORY (msg, thys)) = raised "THEORY" (msg :: map Context.str_of_thy thys) wenzelm@20128: | exn_msg false (Syntax.AST (msg, _)) = raised "AST" [msg] wenzelm@20128: | exn_msg true (Syntax.AST (msg, asts)) = wenzelm@20128: raised "AST" (msg :: map (Pretty.string_of o Syntax.pretty_ast) asts) wenzelm@20128: | exn_msg false (TYPE (msg, _, _)) = raised "TYPE" [msg] wenzelm@20128: | exn_msg true (TYPE (msg, Ts, ts)) = raised "TYPE" (msg :: wenzelm@24920: with_context Syntax.string_of_typ Ts @ with_context Syntax.string_of_term ts) wenzelm@20128: | exn_msg false (TERM (msg, _)) = raised "TERM" [msg] wenzelm@22089: | exn_msg true (TERM (msg, ts)) = wenzelm@24920: raised "TERM" (msg :: with_context Syntax.string_of_term ts) wenzelm@20128: | exn_msg false (THM (msg, _, _)) = raised "THM" [msg] wenzelm@20128: | exn_msg true (THM (msg, i, thms)) = wenzelm@22089: raised ("THM " ^ string_of_int i) (msg :: with_context ProofContext.string_of_thm thms) wenzelm@20128: | exn_msg _ Option.Option = raised "Option" [] wenzelm@20128: | exn_msg _ Library.UnequalLengths = raised "UnequalLengths" [] wenzelm@20128: | exn_msg _ Empty = raised "Empty" [] wenzelm@20128: | exn_msg _ Subscript = raised "Subscript" [] wenzelm@20128: | exn_msg _ (Fail msg) = raised "Fail" [msg] wenzelm@23940: | exn_msg _ exn = General.exnMessage exn; wenzelm@20128: wenzelm@20128: in wenzelm@20128: wenzelm@20128: fun exn_message exn = exn_msg (! debug) exn; wenzelm@20128: wenzelm@22014: fun print_exn NONE = () wenzelm@22014: | print_exn (SOME (exn, s)) = Output.error_msg (cat_lines [exn_message exn, s]); wenzelm@20128: wenzelm@20128: end; wenzelm@20128: wenzelm@20128: wenzelm@20128: (* controlled execution *) wenzelm@20128: wenzelm@20128: local wenzelm@20128: wenzelm@18685: fun debugging f x = wenzelm@23940: if ! debug then exception_trace (fn () => f x) wenzelm@18685: else f x; wenzelm@18685: wenzelm@20128: fun interruptible f x = wenzelm@25219: let val y = ref NONE wenzelm@25219: in raise_interrupt (fn () => y := SOME (f x)) (); the (! y) end; wenzelm@20128: wenzelm@24055: fun toplevel_error f x = f x wenzelm@24055: handle exn => (Output.error_msg (exn_message exn); raise TOPLEVEL_ERROR); wenzelm@24055: wenzelm@20128: in wenzelm@20128: wenzelm@20128: fun controlled_execution f = wenzelm@20128: f wenzelm@20128: |> debugging wenzelm@24055: |> interruptible; wenzelm@20128: wenzelm@20128: fun program f = wenzelm@24055: (f wenzelm@24055: |> debugging wenzelm@24055: |> toplevel_error) (); wenzelm@20128: wenzelm@20128: end; wenzelm@20128: wenzelm@5828: wenzelm@16815: (* node transactions and recovery from stale theories *) wenzelm@6689: wenzelm@16815: (*NB: proof commands should be non-destructive!*) wenzelm@7022: wenzelm@6689: local wenzelm@6689: wenzelm@16452: fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false; wenzelm@6689: wenzelm@18685: val stale_theory = ERROR "Stale theory encountered after succesful execution!"; wenzelm@16815: wenzelm@23363: fun map_theory f = History.map_current wenzelm@21177: (fn Theory (gthy, _) => Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, NONE) wenzelm@21177: | node => node); wenzelm@6689: wenzelm@23363: fun context_position pos = History.map_current wenzelm@24795: (fn Theory (gthy, ctxt) => Theory (ContextPosition.put pos gthy, ctxt) wenzelm@23363: | Proof (prf, x) => wenzelm@24795: Proof (ProofHistory.map_current (Proof.map_context (ContextPosition.put_ctxt pos)) prf, x) wenzelm@23363: | node => node); wenzelm@23363: skalberg@15531: fun return (result, NONE) = result skalberg@15531: | return (result, SOME exn) = raise FAILURE (result, exn); wenzelm@7022: wenzelm@6689: in wenzelm@6689: wenzelm@23363: fun transaction hist pos f (node, term) = wenzelm@20128: let wenzelm@21177: val cont_node = map_theory Theory.checkpoint node; wenzelm@21177: val back_node = map_theory Theory.copy cont_node; wenzelm@21958: fun state nd = State (nd, term); wenzelm@20128: fun normal_state nd = (state nd, NONE); wenzelm@20128: fun error_state nd exn = (state nd, SOME exn); wenzelm@6689: wenzelm@20128: val (result, err) = wenzelm@20128: cont_node wenzelm@23363: |> context_position pos wenzelm@24795: |> map_theory Theory.checkpoint wenzelm@20128: |> (f wenzelm@23363: |> (if hist then History.apply' (History.current back_node) else History.map_current) wenzelm@20128: |> controlled_execution) wenzelm@23363: |> context_position Position.none wenzelm@20128: |> normal_state wenzelm@20128: handle exn => error_state cont_node exn; wenzelm@20128: in wenzelm@20128: if is_stale result wenzelm@20128: then return (error_state back_node (the_default stale_theory err)) wenzelm@20128: else return (result, err) wenzelm@20128: end; wenzelm@6689: wenzelm@6689: end; wenzelm@6689: wenzelm@6689: wenzelm@6689: (* primitive transitions *) wenzelm@6689: wenzelm@18563: (*Note: Recovery from stale theories is provided only for theory-level wenzelm@18589: operations via Transaction. Other node or state operations should wenzelm@18589: not touch theories at all. Interrupts are enabled only for Keep and wenzelm@18589: Transaction.*) wenzelm@5828: wenzelm@5828: datatype trans = wenzelm@22056: Init of (bool -> theory) * ((theory -> unit) * (theory -> unit)) | wenzelm@21958: (*init node; with exit/kill operation*) wenzelm@23911: InitEmpty of unit -> unit | (*init empty toplevel*) wenzelm@21958: Exit | (*conclude node -- deferred until init*) wenzelm@21958: UndoExit | (*continue after conclusion*) wenzelm@21958: Kill | (*abort node*) wenzelm@21958: History of node History.T -> node History.T | (*history operation (undo etc.)*) wenzelm@21958: Keep of bool -> state -> unit | (*peek at state*) wenzelm@21958: Transaction of bool * (bool -> node -> node); (*node transaction*) wenzelm@6689: skalberg@15531: fun undo_limit int = if int then NONE else SOME 0; wenzelm@6689: wenzelm@22056: fun safe_exit (Toplevel (SOME (node, (exit, _)))) = wenzelm@22056: (case try the_global_theory (History.current node) of wenzelm@25219: SOME thy => controlled_execution exit thy wenzelm@22056: | NONE => ()) wenzelm@22056: | safe_exit _ = (); wenzelm@21958: wenzelm@6689: local wenzelm@5828: wenzelm@21958: fun keep_state int f = controlled_execution (fn x => tap (f int) x); wenzelm@21958: wenzelm@23363: fun apply_tr int _ (Init (f, term)) (state as Toplevel _) = wenzelm@22056: let val node = Theory (Context.Theory (f int), NONE) wenzelm@22056: in safe_exit state; State (History.init (undo_limit int) node, term) end wenzelm@23911: | apply_tr int _ (InitEmpty f) (state as Toplevel _) = wenzelm@23911: (safe_exit state; keep_state int (fn _ => fn _ => f ()) toplevel) wenzelm@23363: | apply_tr _ _ Exit (State (node, term)) = wenzelm@22056: (the_global_theory (History.current node); Toplevel (SOME (node, term))) wenzelm@23363: | apply_tr _ _ UndoExit (Toplevel (SOME state_info)) = State state_info wenzelm@23363: | apply_tr _ _ Kill (State (node, (_, kill))) = wenzelm@22056: (kill (the_global_theory (History.current node)); toplevel) wenzelm@23363: | apply_tr _ _ (History f) (State (node, term)) = State (f node, term) wenzelm@23363: | apply_tr int _ (Keep f) state = keep_state int f state wenzelm@23363: | apply_tr int pos (Transaction (hist, f)) (State state) = wenzelm@23363: transaction hist pos (fn x => f int x) state wenzelm@23363: | apply_tr _ _ _ _ = raise UNDEF; wenzelm@5828: wenzelm@23363: fun apply_union _ _ [] state = raise FAILURE (state, UNDEF) wenzelm@23363: | apply_union int pos (tr :: trs) state = wenzelm@23363: apply_tr int pos tr state wenzelm@23363: handle UNDEF => apply_union int pos trs state wenzelm@23363: | FAILURE (alt_state, UNDEF) => apply_union int pos trs alt_state wenzelm@6689: | exn as FAILURE _ => raise exn wenzelm@6689: | exn => raise FAILURE (state, exn); wenzelm@6689: wenzelm@6689: in wenzelm@6689: wenzelm@23363: fun apply_trans int pos trs state = (apply_union int pos 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@16815: {name: string, (*command name*) wenzelm@16815: pos: Position.T, (*source position*) wenzelm@16815: source: OuterLex.token list option, (*source text*) wenzelm@16815: int_only: bool, (*interactive-only*) wenzelm@16815: print: string list, (*print modes (union)*) wenzelm@16815: no_timing: bool, (*suppress timing*) wenzelm@16815: trans: trans list}; (*primitive transitions (union)*) wenzelm@5828: wenzelm@14923: fun make_transition (name, pos, source, int_only, print, no_timing, trans) = wenzelm@14923: Transition {name = name, pos = pos, source = source, wenzelm@14923: int_only = int_only, print = print, no_timing = no_timing, trans = trans}; wenzelm@5828: wenzelm@14923: fun map_transition f (Transition {name, pos, source, int_only, print, no_timing, trans}) = wenzelm@14923: make_transition (f (name, pos, source, int_only, print, no_timing, trans)); wenzelm@5828: wenzelm@16607: val empty = make_transition ("", Position.none, NONE, false, [], false, []); wenzelm@14923: wenzelm@14923: fun name_of (Transition {name, ...}) = name; wenzelm@14923: fun source_of (Transition {source, ...}) = source; wenzelm@5828: wenzelm@5828: wenzelm@5828: (* diagnostics *) wenzelm@5828: wenzelm@5828: fun str_of_transition (Transition {name, pos, ...}) = quote name ^ Position.str_of pos; wenzelm@5828: wenzelm@5828: fun command_msg msg tr = msg ^ "command " ^ str_of_transition tr; wenzelm@5828: 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@14923: fun name nm = map_transition (fn (_, pos, source, int_only, print, no_timing, trans) => wenzelm@14923: (nm, pos, source, int_only, print, no_timing, trans)); wenzelm@5828: wenzelm@14923: fun position pos = map_transition (fn (name, _, source, int_only, print, no_timing, trans) => wenzelm@14923: (name, pos, source, int_only, print, no_timing, trans)); wenzelm@9010: wenzelm@14923: fun source src = map_transition (fn (name, pos, _, int_only, print, no_timing, trans) => skalberg@15531: (name, pos, SOME src, int_only, print, no_timing, trans)); wenzelm@5828: wenzelm@14923: fun interactive int_only = map_transition (fn (name, pos, source, _, print, no_timing, trans) => wenzelm@14923: (name, pos, source, int_only, print, no_timing, trans)); wenzelm@14923: wenzelm@17363: val no_timing = map_transition (fn (name, pos, source, int_only, print, _, trans) => wenzelm@17363: (name, pos, source, int_only, print, true, trans)); wenzelm@17363: wenzelm@17363: fun add_trans tr = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => wenzelm@17363: (name, pos, source, int_only, print, no_timing, trans @ [tr])); wenzelm@17363: wenzelm@16607: fun print' mode = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => wenzelm@16607: (name, pos, source, int_only, insert (op =) mode print, no_timing, trans)); wenzelm@16607: wenzelm@16607: val print = print' ""; wenzelm@5828: wenzelm@17363: val three_buffersN = "three_buffers"; wenzelm@17363: val print3 = print' three_buffersN; wenzelm@5828: wenzelm@5828: wenzelm@21007: (* basic transitions *) wenzelm@5828: wenzelm@22056: fun init_theory f exit kill = add_trans (Init (f, (exit, kill))); wenzelm@21958: val init_empty = add_trans o InitEmpty; wenzelm@6689: val exit = add_trans Exit; wenzelm@21958: val undo_exit = add_trans UndoExit; wenzelm@6689: val kill = add_trans Kill; wenzelm@20128: val history = add_trans o History; wenzelm@7612: val keep' = add_trans o Keep; wenzelm@18592: fun map_current f = add_trans (Transaction (false, f)); wenzelm@18592: fun app_current f = add_trans (Transaction (true, f)); wenzelm@5828: wenzelm@7612: fun keep f = add_trans (Keep (fn _ => f)); wenzelm@5828: fun imperative f = keep (fn _ => f ()); wenzelm@5828: 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@21007: (* theory transitions *) wenzelm@15668: wenzelm@20963: fun theory' f = app_current (fn int => wenzelm@20963: (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE) wenzelm@20963: | _ => raise UNDEF)); wenzelm@20963: wenzelm@20963: fun theory f = theory' (K f); wenzelm@20963: wenzelm@21294: fun begin_local_theory begin f = app_current (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@21294: val end_local_theory = app_current (fn _ => wenzelm@21294: (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy) wenzelm@21007: | _ => raise UNDEF)); wenzelm@21007: wenzelm@21007: local wenzelm@21007: wenzelm@20963: fun local_theory_presentation loc f g = app_current (fn int => wenzelm@21294: (fn Theory (gthy, _) => wenzelm@21294: let wenzelm@24780: val pos = ContextPosition.get (Context.proof_of gthy); wenzelm@21294: val finish = loc_finish loc gthy; wenzelm@24795: val lthy' = f (ContextPosition.put_ctxt pos (loc_begin loc gthy)); wenzelm@21294: in Theory (finish lthy', SOME lthy') end wenzelm@20963: | _ => raise UNDEF) #> tap (g int)); wenzelm@15668: wenzelm@21007: in wenzelm@21007: wenzelm@20963: fun local_theory loc f = local_theory_presentation loc f (K I); wenzelm@20963: fun present_local_theory loc g = local_theory_presentation loc I g; wenzelm@18955: wenzelm@21007: end; wenzelm@21007: wenzelm@21007: wenzelm@21007: (* proof transitions *) wenzelm@21007: wenzelm@21007: fun end_proof f = map_current (fn int => wenzelm@24795: (fn Proof (prf, (finish, _)) => wenzelm@21007: let val state = ProofHistory.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@21007: | SkipProof (h, (gthy, _)) => wenzelm@21007: if History.current h = 0 then Theory (gthy, NONE) else raise UNDEF wenzelm@21007: | _ => raise UNDEF)); wenzelm@21007: wenzelm@21294: local wenzelm@21294: wenzelm@21294: fun begin_proof init finish = app_current (fn int => wenzelm@21294: (fn Theory (gthy, _) => wenzelm@21294: let berghofe@24453: val prf = init int gthy; wenzelm@21294: val schematic = Proof.schematic_goal prf; wenzelm@21294: in wenzelm@21294: if ! skip_proofs andalso schematic then wenzelm@21294: warning "Cannot skip proof of schematic goal statement" wenzelm@21294: else (); wenzelm@21294: if ! skip_proofs andalso not schematic then wenzelm@21294: SkipProof wenzelm@21294: (History.init (undo_limit int) 0, (finish gthy (Proof.global_skip_proof int prf), gthy)) wenzelm@21294: else Proof (ProofHistory.init (undo_limit int) prf, (finish gthy, 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@24780: (fn int => fn gthy => wenzelm@24795: f int (ContextPosition.put_ctxt (ContextPosition.get (Context.proof_of gthy)) wenzelm@24795: (loc_begin loc gthy))) wenzelm@24780: (loc_finish loc); 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@24780: (K (fn Context.Theory thy => f thy | _ => raise UNDEF)) wenzelm@24780: (K (Context.Theory o ProofContext.theory_of)); wenzelm@21294: wenzelm@21294: end; wenzelm@21294: wenzelm@21007: val forget_proof = map_current (fn _ => wenzelm@21007: (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) wenzelm@21007: | SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) wenzelm@21007: | _ => raise UNDEF)); wenzelm@21007: wenzelm@21177: fun present_proof f = map_current (fn int => wenzelm@21177: (fn Proof (prf, x) => Proof (ProofHistory.apply I prf, x) wenzelm@21177: | SkipProof (h, x) => SkipProof (History.apply I h, x) wenzelm@21177: | _ => raise UNDEF) #> tap (f int)); wenzelm@21177: wenzelm@17904: fun proofs' f = map_current (fn int => wenzelm@21007: (fn Proof (prf, x) => Proof (ProofHistory.applys (f int) prf, x) wenzelm@21007: | SkipProof (h, x) => SkipProof (History.apply I h, x) wenzelm@16815: | _ => raise UNDEF)); wenzelm@15668: wenzelm@17904: fun proof' f = proofs' (Seq.single oo f); wenzelm@17904: val proofs = proofs' o K; wenzelm@6689: val proof = proof' o K; wenzelm@16815: wenzelm@16815: fun actual_proof f = map_current (fn _ => wenzelm@21007: (fn Proof (prf, x) => Proof (f prf, x) wenzelm@20963: | _ => raise UNDEF)); wenzelm@16815: wenzelm@16815: fun skip_proof f = map_current (fn _ => wenzelm@21007: (fn SkipProof (h, x) => SkipProof (f h, x) wenzelm@18563: | _ => raise UNDEF)); wenzelm@18563: wenzelm@16815: fun skip_proof_to_theory p = map_current (fn _ => wenzelm@21007: (fn SkipProof (h, (gthy, _)) => wenzelm@21007: if p (History.current h) then Theory (gthy, NONE) wenzelm@17904: else raise UNDEF wenzelm@17904: | _ => raise UNDEF)); wenzelm@5828: wenzelm@5828: wenzelm@5828: wenzelm@5828: (** toplevel transactions **) wenzelm@5828: wenzelm@5828: (* apply transitions *) wenzelm@5828: wenzelm@6664: local wenzelm@6664: wenzelm@23363: fun app int (tr as Transition {trans, pos, int_only, print, no_timing, ...}) state = wenzelm@5828: let wenzelm@21962: val _ = wenzelm@21962: if not int andalso int_only then warning (command_msg "Interactive-only " tr) wenzelm@21962: else (); wenzelm@16682: wenzelm@22588: fun do_timing f x = (warning (command_msg "" tr); timeap f x); wenzelm@16682: fun do_profiling f x = profile (! profiling) f x; wenzelm@16682: wenzelm@6689: val (result, opt_exn) = wenzelm@23363: state |> (apply_trans int pos trans wenzelm@23428: |> (if ! profiling > 0 andalso not no_timing then do_profiling else I) wenzelm@16729: |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I)); wenzelm@21962: val _ = wenzelm@24634: if int andalso not (! quiet) andalso exists (member (op =) print) ("" :: print_mode_value ()) wenzelm@21962: then print_state false result else (); skalberg@15570: in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_exn) end; wenzelm@6664: wenzelm@6664: in wenzelm@5828: wenzelm@6689: fun apply int tr st = wenzelm@6965: (case app int tr st of skalberg@15531: (_, SOME TERMINATE) => NONE skalberg@15531: | (_, SOME RESTART) => SOME (toplevel, NONE) skalberg@15531: | (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info) skalberg@15531: | (state', SOME exn) => SOME (state', SOME (exn, at_command tr)) skalberg@15531: | (state', NONE) => SOME (state', NONE)); wenzelm@6664: wenzelm@6664: end; wenzelm@5828: wenzelm@5828: wenzelm@17076: (* excursion: toplevel -- apply transformers/presentation -- toplevel *) wenzelm@5828: wenzelm@6664: local wenzelm@6664: wenzelm@5828: fun excur [] x = x wenzelm@17076: | excur ((tr, pr) :: trs) (st, res) = wenzelm@17321: (case apply (! interact) tr st of skalberg@15531: SOME (st', NONE) => wenzelm@18685: excur trs (st', pr st st' res handle exn => wenzelm@10324: raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr)) skalberg@15531: | SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info skalberg@15531: | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr)); wenzelm@5828: wenzelm@17076: fun no_pr _ _ _ = (); wenzelm@17076: wenzelm@6664: in wenzelm@6664: wenzelm@17076: fun present_excursion trs res = wenzelm@21958: (case excur trs (toplevel, res) of wenzelm@22056: (state as Toplevel _, res') => (safe_exit state; res') wenzelm@18685: | _ => error "Unfinished development at end of input") wenzelm@9134: handle exn => error (exn_message exn); wenzelm@9134: wenzelm@17076: fun excursion trs = present_excursion (map (rpair no_pr) trs) (); wenzelm@7062: wenzelm@6664: end; wenzelm@6664: wenzelm@5828: wenzelm@5828: wenzelm@5828: (** interactive transformations **) wenzelm@5828: wenzelm@5828: (* the global state reference *) wenzelm@5828: skalberg@15531: val global_state = ref (toplevel, NONE: (exn * string) option); wenzelm@5828: skalberg@15531: fun set_state state = global_state := (state, NONE); wenzelm@5828: fun get_state () = fst (! global_state); wenzelm@5828: fun exn () = snd (! global_state); wenzelm@5828: wenzelm@5828: wenzelm@24295: (* apply transformers to global state --- NOT THREAD-SAFE! *) wenzelm@5828: wenzelm@14985: nonfix >> >>>; wenzelm@5828: wenzelm@24295: fun >> tr = wenzelm@5828: (case apply true tr (get_state ()) of skalberg@15531: NONE => false skalberg@15531: | SOME (state', exn_info) => wenzelm@5828: (global_state := (state', exn_info); wenzelm@22014: print_exn exn_info; wenzelm@24295: true)); wenzelm@5828: wenzelm@14985: fun >>> [] = () wenzelm@14985: | >>> (tr :: trs) = if >> tr then >>> trs else (); wenzelm@14985: wenzelm@21958: fun init_state () = (>> (init_empty (K ()) empty); ()); wenzelm@21958: wenzelm@21958: wenzelm@21958: (* the Isar source of transitions *) wenzelm@21958: wenzelm@21958: type 'a isar = wenzelm@21958: (transition, (transition option, wenzelm@21958: (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token, wenzelm@21958: Position.T * (Symbol.symbol, (string, 'a) Source.source) Source.source) wenzelm@21958: Source.source) Source.source) Source.source) Source.source) Source.source) Source.source; wenzelm@21958: wenzelm@17513: (*Spurious interrupts ahead! Race condition?*) wenzelm@17513: fun get_interrupt src = SOME (Source.get_single src) handle Interrupt => NONE; wenzelm@7602: wenzelm@20928: fun warn_secure () = wenzelm@20928: let val secure = Secure.is_secure () wenzelm@20928: in if secure then warning "Cannot exit to ML in secure mode" else (); secure end; wenzelm@20928: wenzelm@5828: fun raw_loop src = wenzelm@23720: let val prompt = Output.escape (Markup.enclose Markup.prompt Source.default_prompt) in wenzelm@23640: (case get_interrupt (Source.set_prompt prompt src) of wenzelm@23640: NONE => (writeln "\nInterrupt."; raw_loop src) wenzelm@23640: | SOME NONE => if warn_secure () then quit () else () wenzelm@23640: | SOME (SOME (tr, src')) => wenzelm@23640: if >> tr orelse warn_secure () then raw_loop src' wenzelm@23640: else ()) wenzelm@23640: end; wenzelm@5828: wenzelm@12987: fun loop src = ignore_interrupt raw_loop src; wenzelm@5828: wenzelm@5828: end;