| author | chaieb | 
| Wed, 27 Feb 2008 14:39:49 +0100 | |
| changeset 26155 | 7c265e3da23c | 
| parent 26081 | fbdb1161b4b0 | 
| child 26256 | 3e7939e978c6 | 
| permissions | -rw-r--r-- | 
| 5828 | 1 | (* Title: Pure/Isar/toplevel.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | The Isabelle/Isar toplevel. | |
| 6 | *) | |
| 7 | ||
| 8 | signature TOPLEVEL = | |
| 9 | sig | |
| 19063 | 10 | exception UNDEF | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 11 | type generic_theory | 
| 18589 | 12 | type node | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 13 | val theory_node: node -> generic_theory option | 
| 18589 | 14 | val proof_node: node -> ProofHistory.T option | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 15 | val cases_node: (generic_theory -> 'a) -> (Proof.state -> 'a) -> node -> 'a | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 16 | val presentation_context: node option -> xstring option -> Proof.context | 
| 5828 | 17 | type state | 
| 7732 | 18 | val is_toplevel: state -> bool | 
| 18589 | 19 | val is_theory: state -> bool | 
| 20 | val is_proof: state -> bool | |
| 17076 | 21 | val level: state -> int | 
| 6689 | 22 | val node_history_of: state -> node History.T | 
| 5828 | 23 | val node_of: state -> node | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 24 | val node_case: (generic_theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a | 
| 21506 | 25 | val context_of: state -> Proof.context | 
| 22089 | 26 | val generic_theory_of: state -> generic_theory | 
| 5828 | 27 | val theory_of: state -> theory | 
| 28 | val proof_of: state -> Proof.state | |
| 18589 | 29 | val proof_position_of: state -> int | 
| 21007 | 30 | val enter_proof_body: state -> Proof.state | 
| 16815 | 31 | val print_state_context: state -> unit | 
| 32 | val print_state: bool -> state -> unit | |
| 16682 | 33 | val quiet: bool ref | 
| 34 | val debug: bool ref | |
| 17321 
227f1f5c3959
added interact flag to control mode of excursions;
 wenzelm parents: 
17320diff
changeset | 35 | val interact: bool ref | 
| 16682 | 36 | val timing: bool ref | 
| 37 | val profiling: int ref | |
| 16815 | 38 | val skip_proofs: bool ref | 
| 25584 | 39 | val crashes: exn list ref | 
| 5828 | 40 | exception TERMINATE | 
| 5990 | 41 | exception RESTART | 
| 24055 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 42 | exception TOPLEVEL_ERROR | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 43 | val exn_message: exn -> string | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 44 | val program: (unit -> 'a) -> 'a | 
| 16682 | 45 | type transition | 
| 6689 | 46 | val undo_limit: bool -> int option | 
| 5828 | 47 | val empty: transition | 
| 14923 | 48 | val name_of: transition -> string | 
| 49 | val source_of: transition -> OuterLex.token list option | |
| 5828 | 50 | val name: string -> transition -> transition | 
| 51 | val position: Position.T -> transition -> transition | |
| 14923 | 52 | val source: OuterLex.token list -> transition -> transition | 
| 5828 | 53 | val interactive: bool -> transition -> transition | 
| 54 | val print: transition -> transition | |
| 16607 | 55 | val print': string -> transition -> transition | 
| 17363 | 56 | val three_buffersN: string | 
| 57 | val print3: transition -> transition | |
| 9010 | 58 | val no_timing: transition -> transition | 
| 22056 | 59 | val init_theory: (bool -> theory) -> (theory -> unit) -> (theory -> unit) -> | 
| 60 | transition -> transition | |
| 25441 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 61 | val init_empty: (state -> bool) -> (unit -> unit) -> transition -> transition | 
| 6689 | 62 | val exit: transition -> transition | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 63 | val undo_exit: transition -> transition | 
| 6689 | 64 | val kill: transition -> transition | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 65 | val history: (node History.T -> node History.T) -> transition -> transition | 
| 5828 | 66 | val keep: (state -> unit) -> transition -> transition | 
| 7612 | 67 | val keep': (bool -> state -> unit) -> transition -> transition | 
| 5828 | 68 | val imperative: (unit -> unit) -> transition -> transition | 
| 69 | val theory: (theory -> theory) -> transition -> transition | |
| 7612 | 70 | val theory': (bool -> theory -> theory) -> transition -> transition | 
| 20985 | 71 | val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition | 
| 21007 | 72 | val end_local_theory: transition -> transition | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 73 | val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 74 | val present_local_theory: xstring option -> (bool -> node -> unit) -> transition -> transition | 
| 24453 | 75 | val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) -> | 
| 76 | transition -> transition | |
| 21007 | 77 | val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) -> | 
| 78 | transition -> transition | |
| 17363 | 79 | val theory_to_proof: (theory -> Proof.state) -> transition -> transition | 
| 21007 | 80 | val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition | 
| 81 | val forget_proof: transition -> transition | |
| 21177 | 82 | val present_proof: (bool -> node -> unit) -> transition -> transition | 
| 83 | val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 84 | val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition | 
| 21177 | 85 | val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition | 
| 86 | val proof: (Proof.state -> Proof.state) -> transition -> transition | |
| 16815 | 87 | val actual_proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition | 
| 88 | val skip_proof: (int History.T -> int History.T) -> transition -> transition | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 89 | val skip_proof_to_theory: (int -> bool) -> transition -> transition | 
| 9512 | 90 | val unknown_theory: transition -> transition | 
| 91 | val unknown_proof: transition -> transition | |
| 92 | val unknown_context: transition -> transition | |
| 17076 | 93 | val present_excursion: (transition * (state -> state -> 'a -> 'a)) list -> 'a -> 'a | 
| 5828 | 94 | val excursion: transition list -> unit | 
| 95 | val set_state: state -> unit | |
| 96 | val get_state: unit -> state | |
| 97 | val exn: unit -> (exn * string) option | |
| 98 | val >> : transition -> bool | |
| 14985 | 99 | val >>> : transition list -> unit | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 100 | val init_state: unit -> unit | 
| 14091 | 101 | type 'a isar | 
| 25527 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 102 | val loop: bool -> 'a isar -> unit | 
| 5828 | 103 | end; | 
| 104 | ||
| 6965 | 105 | structure Toplevel: TOPLEVEL = | 
| 5828 | 106 | struct | 
| 107 | ||
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 108 | |
| 5828 | 109 | (** toplevel state **) | 
| 110 | ||
| 19063 | 111 | exception UNDEF; | 
| 112 | ||
| 113 | ||
| 21294 | 114 | (* local theory wrappers *) | 
| 5828 | 115 | |
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 116 | type generic_theory = Context.generic; (*theory or local_theory*) | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 117 | |
| 25292 | 118 | val loc_init = TheoryTarget.context; | 
| 21294 | 119 | val loc_exit = ProofContext.theory_of o LocalTheory.exit; | 
| 120 | ||
| 25292 | 121 | fun loc_begin loc (Context.Theory thy) = loc_init (the_default "-" loc) thy | 
| 21294 | 122 | | loc_begin NONE (Context.Proof lthy) = lthy | 
| 25269 | 123 | | loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy); | 
| 21294 | 124 | |
| 125 | fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit | |
| 126 | | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore | |
| 25292 | 127 | | loc_finish (SOME _) (Context.Proof lthy) = fn lthy' => | 
| 128 | Context.Proof (LocalTheory.reinit (LocalTheory.raw_theory (K (loc_exit lthy')) lthy)); | |
| 21294 | 129 | |
| 130 | ||
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 131 | (* datatype node *) | 
| 21294 | 132 | |
| 5828 | 133 | datatype node = | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 134 | Theory of generic_theory * Proof.context option | (*theory with presentation context*) | 
| 21007 | 135 | Proof of ProofHistory.T * ((Proof.context -> generic_theory) * generic_theory) | | 
| 136 | (*history of proof states, finish, original theory*) | |
| 137 | SkipProof of int History.T * (generic_theory * generic_theory); | |
| 18563 | 138 | (*history of proof depths, resulting theory, original theory*) | 
| 5828 | 139 | |
| 22056 | 140 | val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF; | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 141 | val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE; | 
| 18589 | 142 | val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; | 
| 143 | ||
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 144 | fun cases_node f _ (Theory (gthy, _)) = f gthy | 
| 19063 | 145 | | cases_node _ g (Proof (prf, _)) = g (ProofHistory.current prf) | 
| 21007 | 146 | | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy; | 
| 19063 | 147 | |
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 148 | fun presentation_context (SOME (Theory (_, SOME ctxt))) NONE = ctxt | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 149 | | presentation_context (SOME node) NONE = cases_node Context.proof_of Proof.context_of node | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 150 | | presentation_context (SOME node) (SOME loc) = | 
| 25269 | 151 | loc_init loc (cases_node Context.theory_of Proof.theory_of node) | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 152 | | presentation_context NONE _ = raise UNDEF; | 
| 19063 | 153 | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 154 | |
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 155 | (* datatype state *) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 156 | |
| 22056 | 157 | type state_info = node History.T * ((theory -> unit) * (theory -> unit)); | 
| 5828 | 158 | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 159 | datatype state = | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 160 | Toplevel of state_info option | (*outer toplevel, leftover end state*) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 161 | State of state_info; | 
| 5828 | 162 | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 163 | val toplevel = Toplevel NONE; | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 164 | |
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 165 | fun is_toplevel (Toplevel _) = true | 
| 7732 | 166 | | is_toplevel _ = false; | 
| 167 | ||
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 168 | fun level (Toplevel _) = 0 | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 169 | | level (State (node, _)) = | 
| 17076 | 170 | (case History.current node of | 
| 21310 
bfcc24fc7c46
level: do not account for local theory blocks (relevant for document preparation);
 wenzelm parents: 
21294diff
changeset | 171 | Theory _ => 0 | 
| 
bfcc24fc7c46
level: do not account for local theory blocks (relevant for document preparation);
 wenzelm parents: 
21294diff
changeset | 172 | | Proof (prf, _) => Proof.level (ProofHistory.current prf) | 
| 
bfcc24fc7c46
level: do not account for local theory blocks (relevant for document preparation);
 wenzelm parents: 
21294diff
changeset | 173 | | SkipProof (h, _) => History.current h + 1); (*different notion of proof depth!*) | 
| 17076 | 174 | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 175 | fun str_of_state (Toplevel _) = "at top level" | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 176 | | str_of_state (State (node, _)) = | 
| 16815 | 177 | (case History.current node of | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 178 | Theory (Context.Theory _, _) => "in theory mode" | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 179 | | Theory (Context.Proof _, _) => "in local theory mode" | 
| 16815 | 180 | | Proof _ => "in proof mode" | 
| 181 | | SkipProof _ => "in skipped proof mode"); | |
| 5946 
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
 wenzelm parents: 
5939diff
changeset | 182 | |
| 
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
 wenzelm parents: 
5939diff
changeset | 183 | |
| 5828 | 184 | (* top node *) | 
| 185 | ||
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 186 | fun node_history_of (Toplevel _) = raise UNDEF | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 187 | | node_history_of (State (node, _)) = node; | 
| 6689 | 188 | |
| 189 | val node_of = History.current o node_history_of; | |
| 5828 | 190 | |
| 18589 | 191 | fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state)); | 
| 192 | fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state)); | |
| 193 | ||
| 19063 | 194 | fun node_case f g state = cases_node f g (node_of state); | 
| 5828 | 195 | |
| 21506 | 196 | val context_of = node_case Context.proof_of Proof.context_of; | 
| 22089 | 197 | val generic_theory_of = node_case I (Context.Proof o Proof.context_of); | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 198 | val theory_of = node_case Context.theory_of Proof.theory_of; | 
| 18589 | 199 | val proof_of = node_case (fn _ => raise UNDEF) I; | 
| 17208 | 200 | |
| 18589 | 201 | fun proof_position_of state = | 
| 202 | (case node_of state of | |
| 203 | Proof (prf, _) => ProofHistory.position prf | |
| 204 | | _ => raise UNDEF); | |
| 6664 | 205 | |
| 21007 | 206 | val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward; | 
| 5828 | 207 | |
| 208 | ||
| 16815 | 209 | (* print state *) | 
| 210 | ||
| 25292 | 211 | val pretty_context = LocalTheory.pretty o Context.cases (TheoryTarget.init NONE) I; | 
| 16815 | 212 | |
| 23640 
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
 wenzelm parents: 
23619diff
changeset | 213 | fun print_state_context state = | 
| 24795 
6f5cb7885fd7
print_state_context: local theory context, not proof context;
 wenzelm parents: 
24780diff
changeset | 214 | (case try node_of state of | 
| 21506 | 215 | NONE => [] | 
| 24795 
6f5cb7885fd7
print_state_context: local theory context, not proof context;
 wenzelm parents: 
24780diff
changeset | 216 | | SOME (Theory (gthy, _)) => pretty_context gthy | 
| 
6f5cb7885fd7
print_state_context: local theory context, not proof context;
 wenzelm parents: 
24780diff
changeset | 217 | | SOME (Proof (_, (_, gthy))) => pretty_context gthy | 
| 
6f5cb7885fd7
print_state_context: local theory context, not proof context;
 wenzelm parents: 
24780diff
changeset | 218 | | SOME (SkipProof (_, (gthy, _))) => pretty_context gthy) | 
| 23640 
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
 wenzelm parents: 
23619diff
changeset | 219 | |> Pretty.chunks |> Pretty.writeln; | 
| 16815 | 220 | |
| 23640 
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
 wenzelm parents: 
23619diff
changeset | 221 | fun print_state prf_only state = | 
| 23701 | 222 | (case try node_of state of | 
| 223 | NONE => [] | |
| 224 | | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy | |
| 225 | | SOME (Proof (prf, _)) => | |
| 226 | Proof.pretty_state (ProofHistory.position prf) (ProofHistory.current prf) | |
| 227 | | SOME (SkipProof (h, _)) => | |
| 228 |       [Pretty.str ("skipped proof: depth " ^ string_of_int (History.current h))])
 | |
| 229 | |> Pretty.markup_chunks Markup.state |> Pretty.writeln; | |
| 16815 | 230 | |
| 231 | ||
| 15668 | 232 | |
| 5828 | 233 | (** toplevel transitions **) | 
| 234 | ||
| 16682 | 235 | val quiet = ref false; | 
| 22135 | 236 | val debug = Output.debugging; | 
| 17321 
227f1f5c3959
added interact flag to control mode of excursions;
 wenzelm parents: 
17320diff
changeset | 237 | val interact = ref false; | 
| 16682 | 238 | val timing = Output.timing; | 
| 239 | val profiling = ref 0; | |
| 16815 | 240 | val skip_proofs = ref false; | 
| 25584 | 241 | val crashes = ref ([]: exn list); | 
| 16682 | 242 | |
| 5828 | 243 | exception TERMINATE; | 
| 5990 | 244 | exception RESTART; | 
| 7022 | 245 | exception EXCURSION_FAIL of exn * string; | 
| 6689 | 246 | exception FAILURE of state * exn; | 
| 24055 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 247 | exception TOPLEVEL_ERROR; | 
| 5828 | 248 | |
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 249 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 250 | (* print exceptions *) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 251 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 252 | local | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 253 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 254 | fun with_context f xs = | 
| 22095 
07875394618e
moved ML context stuff to from Context to ML_Context;
 wenzelm parents: 
22089diff
changeset | 255 | (case ML_Context.get_context () of NONE => [] | 
| 22089 | 256 | | SOME context => map (f (Context.proof_of context)) xs); | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 257 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 258 | fun raised name [] = "exception " ^ name ^ " raised" | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 259 | | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 260 |   | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);
 | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 261 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 262 | fun exn_msg _ TERMINATE = "Exit." | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 263 | | exn_msg _ RESTART = "Restart." | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 264 | | exn_msg _ Interrupt = "Interrupt." | 
| 26081 | 265 | | exn_msg _ TimeLimit.TimeOut = "Timeout." | 
| 24055 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 266 | | exn_msg _ TOPLEVEL_ERROR = "Error." | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 267 | | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 268 | | exn_msg _ (ERROR msg) = msg | 
| 23975 | 269 | | exn_msg detailed (Exn.EXCEPTIONS (exns, "")) = cat_lines (map (exn_msg detailed) exns) | 
| 23963 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23940diff
changeset | 270 | | exn_msg detailed (Exn.EXCEPTIONS (exns, msg)) = cat_lines (map (exn_msg detailed) exns @ [msg]) | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 271 | | exn_msg detailed (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_msg detailed exn, msg] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 272 | | exn_msg false (THEORY (msg, _)) = msg | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 273 | | exn_msg true (THEORY (msg, thys)) = raised "THEORY" (msg :: map Context.str_of_thy thys) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 274 | | exn_msg false (Syntax.AST (msg, _)) = raised "AST" [msg] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 275 | | exn_msg true (Syntax.AST (msg, asts)) = | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 276 | raised "AST" (msg :: map (Pretty.string_of o Syntax.pretty_ast) asts) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 277 | | exn_msg false (TYPE (msg, _, _)) = raised "TYPE" [msg] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 278 | | exn_msg true (TYPE (msg, Ts, ts)) = raised "TYPE" (msg :: | 
| 24920 | 279 | with_context Syntax.string_of_typ Ts @ with_context Syntax.string_of_term ts) | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 280 | | exn_msg false (TERM (msg, _)) = raised "TERM" [msg] | 
| 22089 | 281 | | exn_msg true (TERM (msg, ts)) = | 
| 24920 | 282 | raised "TERM" (msg :: with_context Syntax.string_of_term ts) | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 283 | | exn_msg false (THM (msg, _, _)) = raised "THM" [msg] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 284 | | exn_msg true (THM (msg, i, thms)) = | 
| 22089 | 285 |       raised ("THM " ^ string_of_int i) (msg :: with_context ProofContext.string_of_thm thms)
 | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 286 | | exn_msg _ Option.Option = raised "Option" [] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 287 | | exn_msg _ Library.UnequalLengths = raised "UnequalLengths" [] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 288 | | exn_msg _ Empty = raised "Empty" [] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 289 | | exn_msg _ Subscript = raised "Subscript" [] | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 290 | | exn_msg _ (Fail msg) = raised "Fail" [msg] | 
| 23940 | 291 | | exn_msg _ exn = General.exnMessage exn; | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 292 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 293 | in | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 294 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 295 | fun exn_message exn = exn_msg (! debug) exn; | 
| 25809 
decb98ff92dd
toplevel print_exn: proper setmp_thread_properties;
 wenzelm parents: 
25799diff
changeset | 296 | fun print_exn (exn, s) = Output.error_msg (cat_lines [exn_message exn, s]); | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 297 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 298 | end; | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 299 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 300 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 301 | (* controlled execution *) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 302 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 303 | local | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 304 | |
| 18685 | 305 | fun debugging f x = | 
| 23940 | 306 | if ! debug then exception_trace (fn () => f x) | 
| 18685 | 307 | else f x; | 
| 308 | ||
| 24055 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 309 | fun toplevel_error f x = f x | 
| 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 310 | handle exn => (Output.error_msg (exn_message exn); raise TOPLEVEL_ERROR); | 
| 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 311 | |
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 312 | in | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 313 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 314 | fun controlled_execution f = | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 315 | f | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 316 | |> debugging | 
| 24055 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 317 | |> interruptible; | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 318 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 319 | fun program f = | 
| 24055 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 320 | (f | 
| 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 321 | |> debugging | 
| 
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
 wenzelm parents: 
23975diff
changeset | 322 | |> toplevel_error) (); | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 323 | |
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 324 | end; | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 325 | |
| 5828 | 326 | |
| 16815 | 327 | (* node transactions and recovery from stale theories *) | 
| 6689 | 328 | |
| 16815 | 329 | (*NB: proof commands should be non-destructive!*) | 
| 7022 | 330 | |
| 6689 | 331 | local | 
| 332 | ||
| 16452 | 333 | fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false; | 
| 6689 | 334 | |
| 18685 | 335 | val stale_theory = ERROR "Stale theory encountered after succesful execution!"; | 
| 16815 | 336 | |
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 337 | fun map_theory f = History.map_current | 
| 21177 | 338 | (fn Theory (gthy, _) => Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, NONE) | 
| 339 | | node => node); | |
| 6689 | 340 | |
| 15531 | 341 | fun return (result, NONE) = result | 
| 342 | | return (result, SOME exn) = raise FAILURE (result, exn); | |
| 7022 | 343 | |
| 6689 | 344 | in | 
| 345 | ||
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 346 | fun transaction hist pos f (node, term) = | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 347 | let | 
| 21177 | 348 | val cont_node = map_theory Theory.checkpoint node; | 
| 349 | val back_node = map_theory Theory.copy cont_node; | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 350 | fun state nd = State (nd, term); | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 351 | fun normal_state nd = (state nd, NONE); | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 352 | fun error_state nd exn = (state nd, SOME exn); | 
| 6689 | 353 | |
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 354 | val (result, err) = | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 355 | cont_node | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 356 | |> (f | 
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 357 | |> (if hist then History.apply' (History.current back_node) else History.map_current) | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 358 | |> controlled_execution) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 359 | |> normal_state | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 360 | handle exn => error_state cont_node exn; | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 361 | in | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 362 | if is_stale result | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 363 | then return (error_state back_node (the_default stale_theory err)) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 364 | else return (result, err) | 
| 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 365 | end; | 
| 6689 | 366 | |
| 367 | end; | |
| 368 | ||
| 369 | ||
| 370 | (* primitive transitions *) | |
| 371 | ||
| 18563 | 372 | (*Note: Recovery from stale theories is provided only for theory-level | 
| 18589 | 373 | operations via Transaction. Other node or state operations should | 
| 374 | not touch theories at all. Interrupts are enabled only for Keep and | |
| 375 | Transaction.*) | |
| 5828 | 376 | |
| 377 | datatype trans = | |
| 22056 | 378 | Init of (bool -> theory) * ((theory -> unit) * (theory -> unit)) | | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 379 | (*init node; with exit/kill operation*) | 
| 25441 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 380 | InitEmpty of (state -> bool) * (unit -> unit) | (*init empty toplevel*) | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 381 | Exit | (*conclude node -- deferred until init*) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 382 | UndoExit | (*continue after conclusion*) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 383 | Kill | (*abort node*) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 384 | History of node History.T -> node History.T | (*history operation (undo etc.)*) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 385 | Keep of bool -> state -> unit | (*peek at state*) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 386 | Transaction of bool * (bool -> node -> node); (*node transaction*) | 
| 6689 | 387 | |
| 15531 | 388 | fun undo_limit int = if int then NONE else SOME 0; | 
| 6689 | 389 | |
| 22056 | 390 | fun safe_exit (Toplevel (SOME (node, (exit, _)))) = | 
| 391 | (case try the_global_theory (History.current node) of | |
| 25219 | 392 | SOME thy => controlled_execution exit thy | 
| 22056 | 393 | | NONE => ()) | 
| 394 | | safe_exit _ = (); | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 395 | |
| 6689 | 396 | local | 
| 5828 | 397 | |
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 398 | fun keep_state int f = controlled_execution (fn x => tap (f int) x); | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 399 | |
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 400 | fun apply_tr int _ (Init (f, term)) (state as Toplevel _) = | 
| 22056 | 401 | let val node = Theory (Context.Theory (f int), NONE) | 
| 402 | in safe_exit state; State (History.init (undo_limit int) node, term) end | |
| 25441 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 403 | | apply_tr int _ (InitEmpty (check, f)) (state as Toplevel _) = | 
| 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 404 | if check state then (safe_exit state; keep_state int (fn _ => fn _ => f ()) toplevel) | 
| 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 405 | else raise UNDEF | 
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 406 | | apply_tr _ _ Exit (State (node, term)) = | 
| 22056 | 407 | (the_global_theory (History.current node); Toplevel (SOME (node, term))) | 
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 408 | | apply_tr _ _ UndoExit (Toplevel (SOME state_info)) = State state_info | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 409 | | apply_tr _ _ Kill (State (node, (_, kill))) = | 
| 22056 | 410 | (kill (the_global_theory (History.current node)); toplevel) | 
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 411 | | apply_tr _ _ (History f) (State (node, term)) = State (f node, term) | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 412 | | apply_tr int _ (Keep f) state = keep_state int f state | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 413 | | apply_tr int pos (Transaction (hist, f)) (State state) = | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 414 | transaction hist pos (fn x => f int x) state | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 415 | | apply_tr _ _ _ _ = raise UNDEF; | 
| 5828 | 416 | |
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 417 | fun apply_union _ _ [] state = raise FAILURE (state, UNDEF) | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 418 | | apply_union int pos (tr :: trs) state = | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 419 | apply_tr int pos tr state | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 420 | handle UNDEF => apply_union int pos trs state | 
| 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 421 | | FAILURE (alt_state, UNDEF) => apply_union int pos trs alt_state | 
| 6689 | 422 | | exn as FAILURE _ => raise exn | 
| 423 | | exn => raise FAILURE (state, exn); | |
| 424 | ||
| 425 | in | |
| 426 | ||
| 23363 
9981199bf865
transaction: context_position (non-destructive only);
 wenzelm parents: 
22588diff
changeset | 427 | fun apply_trans int pos trs state = (apply_union int pos trs state, NONE) | 
| 15531 | 428 | handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); | 
| 6689 | 429 | |
| 430 | end; | |
| 5828 | 431 | |
| 432 | ||
| 433 | (* datatype transition *) | |
| 434 | ||
| 435 | datatype transition = Transition of | |
| 16815 | 436 |  {name: string,                        (*command name*)
 | 
| 437 | pos: Position.T, (*source position*) | |
| 438 | source: OuterLex.token list option, (*source text*) | |
| 439 | int_only: bool, (*interactive-only*) | |
| 440 | print: string list, (*print modes (union)*) | |
| 441 | no_timing: bool, (*suppress timing*) | |
| 442 | trans: trans list}; (*primitive transitions (union)*) | |
| 5828 | 443 | |
| 25960 | 444 | fun make_transition (name, pos, source, int_only, print, no_timing, trans) = | 
| 445 |   Transition {name = name, pos = pos, source = source,
 | |
| 14923 | 446 | int_only = int_only, print = print, no_timing = no_timing, trans = trans}; | 
| 5828 | 447 | |
| 25960 | 448 | fun map_transition f (Transition {name, pos, source, int_only, print, no_timing, trans}) =
 | 
| 449 | make_transition (f (name, pos, source, int_only, print, no_timing, trans)); | |
| 5828 | 450 | |
| 25960 | 451 | val empty = make_transition ("<unknown>", Position.none, NONE, false, [], false, []);
 | 
| 14923 | 452 | |
| 453 | fun name_of (Transition {name, ...}) = name;
 | |
| 454 | fun source_of (Transition {source, ...}) = source;
 | |
| 5828 | 455 | |
| 456 | ||
| 457 | (* diagnostics *) | |
| 458 | ||
| 459 | fun str_of_transition (Transition {name, pos, ...}) = quote name ^ Position.str_of pos;
 | |
| 460 | ||
| 461 | fun command_msg msg tr = msg ^ "command " ^ str_of_transition tr; | |
| 462 | fun at_command tr = command_msg "At " tr ^ "."; | |
| 463 | ||
| 464 | fun type_error tr state = | |
| 18685 | 465 | ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state); | 
| 5828 | 466 | |
| 467 | ||
| 468 | (* modify transitions *) | |
| 469 | ||
| 25960 | 470 | fun name nm = map_transition (fn (_, pos, source, int_only, print, no_timing, trans) => | 
| 471 | (nm, pos, source, int_only, print, no_timing, trans)); | |
| 5828 | 472 | |
| 25960 | 473 | fun position pos = map_transition (fn (name, _, source, int_only, print, no_timing, trans) => | 
| 474 | (name, pos, source, int_only, print, no_timing, trans)); | |
| 9010 | 475 | |
| 25960 | 476 | fun source src = map_transition (fn (name, pos, _, int_only, print, no_timing, trans) => | 
| 477 | (name, pos, SOME src, int_only, print, no_timing, trans)); | |
| 5828 | 478 | |
| 25960 | 479 | fun interactive int_only = map_transition (fn (name, pos, source, _, print, no_timing, trans) => | 
| 480 | (name, pos, source, int_only, print, no_timing, trans)); | |
| 14923 | 481 | |
| 25960 | 482 | val no_timing = map_transition (fn (name, pos, source, int_only, print, _, trans) => | 
| 483 | (name, pos, source, int_only, print, true, trans)); | |
| 17363 | 484 | |
| 25960 | 485 | fun add_trans tr = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => | 
| 486 | (name, pos, source, int_only, print, no_timing, trans @ [tr])); | |
| 17363 | 487 | |
| 25960 | 488 | fun print' mode = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => | 
| 489 | (name, pos, source, int_only, insert (op =) mode print, no_timing, trans)); | |
| 16607 | 490 | |
| 491 | val print = print' ""; | |
| 5828 | 492 | |
| 17363 | 493 | val three_buffersN = "three_buffers"; | 
| 494 | val print3 = print' three_buffersN; | |
| 5828 | 495 | |
| 496 | ||
| 21007 | 497 | (* basic transitions *) | 
| 5828 | 498 | |
| 22056 | 499 | fun init_theory f exit kill = add_trans (Init (f, (exit, kill))); | 
| 25441 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 500 | fun init_empty check f = add_trans (InitEmpty (check, f)); | 
| 6689 | 501 | val exit = add_trans Exit; | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 502 | val undo_exit = add_trans UndoExit; | 
| 6689 | 503 | val kill = add_trans Kill; | 
| 20128 
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
 wenzelm parents: 
19996diff
changeset | 504 | val history = add_trans o History; | 
| 7612 | 505 | val keep' = add_trans o Keep; | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 506 | fun map_current f = add_trans (Transaction (false, f)); | 
| 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 507 | fun app_current f = add_trans (Transaction (true, f)); | 
| 5828 | 508 | |
| 7612 | 509 | fun keep f = add_trans (Keep (fn _ => f)); | 
| 5828 | 510 | fun imperative f = keep (fn _ => f ()); | 
| 511 | ||
| 21007 | 512 | val unknown_theory = imperative (fn () => warning "Unknown theory context"); | 
| 513 | val unknown_proof = imperative (fn () => warning "Unknown proof context"); | |
| 514 | val unknown_context = imperative (fn () => warning "Unknown context"); | |
| 15668 | 515 | |
| 21007 | 516 | |
| 517 | (* theory transitions *) | |
| 15668 | 518 | |
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 519 | fun theory' f = app_current (fn int => | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 520 | (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE) | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 521 | | _ => raise UNDEF)); | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 522 | |
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 523 | fun theory f = theory' (K f); | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 524 | |
| 21294 | 525 | fun begin_local_theory begin f = app_current (fn _ => | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 526 | (fn Theory (Context.Theory thy, _) => | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 527 | let | 
| 20985 | 528 | val lthy = f thy; | 
| 21294 | 529 | val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy); | 
| 530 | in Theory (gthy, SOME lthy) end | |
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 531 | | _ => raise UNDEF)); | 
| 17076 | 532 | |
| 21294 | 533 | val end_local_theory = app_current (fn _ => | 
| 534 | (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy) | |
| 21007 | 535 | | _ => raise UNDEF)); | 
| 536 | ||
| 537 | local | |
| 538 | ||
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 539 | fun local_theory_presentation loc f g = app_current (fn int => | 
| 21294 | 540 | (fn Theory (gthy, _) => | 
| 541 | let | |
| 542 | val finish = loc_finish loc gthy; | |
| 25960 | 543 | val lthy' = f (loc_begin loc gthy); | 
| 21294 | 544 | in Theory (finish lthy', SOME lthy') end | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 545 | | _ => raise UNDEF) #> tap (g int)); | 
| 15668 | 546 | |
| 21007 | 547 | in | 
| 548 | ||
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 549 | fun local_theory loc f = local_theory_presentation loc f (K I); | 
| 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 550 | fun present_local_theory loc g = local_theory_presentation loc I g; | 
| 18955 | 551 | |
| 21007 | 552 | end; | 
| 553 | ||
| 554 | ||
| 555 | (* proof transitions *) | |
| 556 | ||
| 557 | fun end_proof f = map_current (fn int => | |
| 24795 
6f5cb7885fd7
print_state_context: local theory context, not proof context;
 wenzelm parents: 
24780diff
changeset | 558 | (fn Proof (prf, (finish, _)) => | 
| 21007 | 559 | let val state = ProofHistory.current prf in | 
| 560 | if can (Proof.assert_bottom true) state then | |
| 561 | let | |
| 562 | val ctxt' = f int state; | |
| 563 | val gthy' = finish ctxt'; | |
| 564 | in Theory (gthy', SOME ctxt') end | |
| 565 | else raise UNDEF | |
| 566 | end | |
| 567 | | SkipProof (h, (gthy, _)) => | |
| 568 | if History.current h = 0 then Theory (gthy, NONE) else raise UNDEF | |
| 569 | | _ => raise UNDEF)); | |
| 570 | ||
| 21294 | 571 | local | 
| 572 | ||
| 573 | fun begin_proof init finish = app_current (fn int => | |
| 574 | (fn Theory (gthy, _) => | |
| 575 | let | |
| 24453 | 576 | val prf = init int gthy; | 
| 21294 | 577 | val schematic = Proof.schematic_goal prf; | 
| 578 | in | |
| 579 | if ! skip_proofs andalso schematic then | |
| 580 | warning "Cannot skip proof of schematic goal statement" | |
| 581 | else (); | |
| 582 | if ! skip_proofs andalso not schematic then | |
| 583 | SkipProof | |
| 584 | (History.init (undo_limit int) 0, (finish gthy (Proof.global_skip_proof int prf), gthy)) | |
| 585 | else Proof (ProofHistory.init (undo_limit int) prf, (finish gthy, gthy)) | |
| 586 | end | |
| 587 | | _ => raise UNDEF)); | |
| 588 | ||
| 589 | in | |
| 590 | ||
| 24780 
47bb1e380d83
local_theory transactions: more careful treatment of context position;
 wenzelm parents: 
24634diff
changeset | 591 | fun local_theory_to_proof' loc f = begin_proof | 
| 25960 | 592 | (fn int => fn gthy => f int (loc_begin loc gthy)) | 
| 24780 
47bb1e380d83
local_theory transactions: more careful treatment of context position;
 wenzelm parents: 
24634diff
changeset | 593 | (loc_finish loc); | 
| 
47bb1e380d83
local_theory transactions: more careful treatment of context position;
 wenzelm parents: 
24634diff
changeset | 594 | |
| 24453 | 595 | fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f); | 
| 21294 | 596 | |
| 597 | fun theory_to_proof f = begin_proof | |
| 24780 
47bb1e380d83
local_theory transactions: more careful treatment of context position;
 wenzelm parents: 
24634diff
changeset | 598 | (K (fn Context.Theory thy => f thy | _ => raise UNDEF)) | 
| 
47bb1e380d83
local_theory transactions: more careful treatment of context position;
 wenzelm parents: 
24634diff
changeset | 599 | (K (Context.Theory o ProofContext.theory_of)); | 
| 21294 | 600 | |
| 601 | end; | |
| 602 | ||
| 21007 | 603 | val forget_proof = map_current (fn _ => | 
| 604 | (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) | |
| 605 | | SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) | |
| 606 | | _ => raise UNDEF)); | |
| 607 | ||
| 21177 | 608 | fun present_proof f = map_current (fn int => | 
| 609 | (fn Proof (prf, x) => Proof (ProofHistory.apply I prf, x) | |
| 610 | | SkipProof (h, x) => SkipProof (History.apply I h, x) | |
| 611 | | _ => raise UNDEF) #> tap (f int)); | |
| 612 | ||
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 613 | fun proofs' f = map_current (fn int => | 
| 21007 | 614 | (fn Proof (prf, x) => Proof (ProofHistory.applys (f int) prf, x) | 
| 615 | | SkipProof (h, x) => SkipProof (History.apply I h, x) | |
| 16815 | 616 | | _ => raise UNDEF)); | 
| 15668 | 617 | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 618 | fun proof' f = proofs' (Seq.single oo f); | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 619 | val proofs = proofs' o K; | 
| 6689 | 620 | val proof = proof' o K; | 
| 16815 | 621 | |
| 622 | fun actual_proof f = map_current (fn _ => | |
| 21007 | 623 | (fn Proof (prf, x) => Proof (f prf, x) | 
| 20963 
a7fd8f05a2be
added type global_theory -- theory or local_theory;
 wenzelm parents: 
20928diff
changeset | 624 | | _ => raise UNDEF)); | 
| 16815 | 625 | |
| 626 | fun skip_proof f = map_current (fn _ => | |
| 21007 | 627 | (fn SkipProof (h, x) => SkipProof (f h, x) | 
| 18563 | 628 | | _ => raise UNDEF)); | 
| 629 | ||
| 16815 | 630 | fun skip_proof_to_theory p = map_current (fn _ => | 
| 21007 | 631 | (fn SkipProof (h, (gthy, _)) => | 
| 632 | if p (History.current h) then Theory (gthy, NONE) | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 633 | else raise UNDEF | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 634 | | _ => raise UNDEF)); | 
| 5828 | 635 | |
| 636 | ||
| 637 | ||
| 638 | (** toplevel transactions **) | |
| 639 | ||
| 25960 | 640 | (* thread position *) | 
| 25799 | 641 | |
| 25960 | 642 | fun setmp_thread_position (Transition {pos, ...}) f x =
 | 
| 25819 
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
 wenzelm parents: 
25809diff
changeset | 643 | Position.setmp_thread_data pos f x; | 
| 25799 | 644 | |
| 645 | ||
| 5828 | 646 | (* apply transitions *) | 
| 647 | ||
| 6664 | 648 | local | 
| 649 | ||
| 25799 | 650 | fun app int (tr as Transition {trans, pos, int_only, print, no_timing, ...}) =
 | 
| 25819 
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
 wenzelm parents: 
25809diff
changeset | 651 | setmp_thread_position tr (fn state => | 
| 25799 | 652 | let | 
| 653 | val _ = | |
| 654 | if not int andalso int_only then warning (command_msg "Interactive-only " tr) | |
| 655 | else (); | |
| 16682 | 656 | |
| 25799 | 657 | fun do_timing f x = (warning (command_msg "" tr); timeap f x); | 
| 658 | fun do_profiling f x = profile (! profiling) f x; | |
| 659 | ||
| 660 | val (result, opt_exn) = | |
| 661 | state |> (apply_trans int pos trans | |
| 662 | |> (if ! profiling > 0 andalso not no_timing then do_profiling else I) | |
| 663 | |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I)); | |
| 664 | val _ = | |
| 665 |         if int andalso not (! quiet) andalso exists (member (op =) print) ("" :: print_mode_value ())
 | |
| 666 | then print_state false result else (); | |
| 667 | in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_exn) end); | |
| 6664 | 668 | |
| 669 | in | |
| 5828 | 670 | |
| 6689 | 671 | fun apply int tr st = | 
| 6965 | 672 | (case app int tr st of | 
| 15531 | 673 | (_, SOME TERMINATE) => NONE | 
| 674 | | (_, SOME RESTART) => SOME (toplevel, NONE) | |
| 675 | | (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info) | |
| 676 | | (state', SOME exn) => SOME (state', SOME (exn, at_command tr)) | |
| 677 | | (state', NONE) => SOME (state', NONE)); | |
| 6664 | 678 | |
| 679 | end; | |
| 5828 | 680 | |
| 681 | ||
| 17076 | 682 | (* excursion: toplevel -- apply transformers/presentation -- toplevel *) | 
| 5828 | 683 | |
| 6664 | 684 | local | 
| 685 | ||
| 5828 | 686 | fun excur [] x = x | 
| 17076 | 687 | | excur ((tr, pr) :: trs) (st, res) = | 
| 17321 
227f1f5c3959
added interact flag to control mode of excursions;
 wenzelm parents: 
17320diff
changeset | 688 | (case apply (! interact) tr st of | 
| 15531 | 689 | SOME (st', NONE) => | 
| 18685 | 690 | excur trs (st', pr st st' res handle exn => | 
| 10324 | 691 | raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr)) | 
| 15531 | 692 | | SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info | 
| 693 | | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr)); | |
| 5828 | 694 | |
| 17076 | 695 | fun no_pr _ _ _ = (); | 
| 696 | ||
| 6664 | 697 | in | 
| 698 | ||
| 17076 | 699 | fun present_excursion trs res = | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 700 | (case excur trs (toplevel, res) of | 
| 22056 | 701 | (state as Toplevel _, res') => (safe_exit state; res') | 
| 18685 | 702 | | _ => error "Unfinished development at end of input") | 
| 9134 | 703 | handle exn => error (exn_message exn); | 
| 704 | ||
| 17076 | 705 | fun excursion trs = present_excursion (map (rpair no_pr) trs) (); | 
| 7062 | 706 | |
| 6664 | 707 | end; | 
| 708 | ||
| 5828 | 709 | |
| 710 | ||
| 711 | (** interactive transformations **) | |
| 712 | ||
| 713 | (* the global state reference *) | |
| 714 | ||
| 15531 | 715 | val global_state = ref (toplevel, NONE: (exn * string) option); | 
| 5828 | 716 | |
| 15531 | 717 | fun set_state state = global_state := (state, NONE); | 
| 5828 | 718 | fun get_state () = fst (! global_state); | 
| 719 | fun exn () = snd (! global_state); | |
| 720 | ||
| 721 | ||
| 24295 
af8dbc7a2305
global state transformation: non-critical, but also non-thread-safe;
 wenzelm parents: 
24058diff
changeset | 722 | (* apply transformers to global state --- NOT THREAD-SAFE! *) | 
| 5828 | 723 | |
| 14985 | 724 | nonfix >> >>>; | 
| 5828 | 725 | |
| 24295 
af8dbc7a2305
global state transformation: non-critical, but also non-thread-safe;
 wenzelm parents: 
24058diff
changeset | 726 | fun >> tr = | 
| 5828 | 727 | (case apply true tr (get_state ()) of | 
| 15531 | 728 | NONE => false | 
| 729 | | SOME (state', exn_info) => | |
| 5828 | 730 | (global_state := (state', exn_info); | 
| 25809 
decb98ff92dd
toplevel print_exn: proper setmp_thread_properties;
 wenzelm parents: 
25799diff
changeset | 731 | (case exn_info of | 
| 
decb98ff92dd
toplevel print_exn: proper setmp_thread_properties;
 wenzelm parents: 
25799diff
changeset | 732 | NONE => () | 
| 25819 
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
 wenzelm parents: 
25809diff
changeset | 733 | | SOME err => setmp_thread_position tr print_exn err); | 
| 24295 
af8dbc7a2305
global state transformation: non-critical, but also non-thread-safe;
 wenzelm parents: 
24058diff
changeset | 734 | true)); | 
| 5828 | 735 | |
| 14985 | 736 | fun >>> [] = () | 
| 737 | | >>> (tr :: trs) = if >> tr then >>> trs else (); | |
| 738 | ||
| 25441 
4028958d19ff
init_empty: check before change (avoids non-linear update);
 wenzelm parents: 
25292diff
changeset | 739 | fun init_state () = (>> (init_empty (K true) (K ()) empty); ()); | 
| 21958 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 740 | |
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 741 | |
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 742 | (* the Isar source of transitions *) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 743 | |
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 744 | type 'a isar = | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 745 | (transition, (transition option, | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 746 | (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token, | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 747 | Position.T * (Symbol.symbol, (string, 'a) Source.source) Source.source) | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 748 | Source.source) Source.source) Source.source) Source.source) Source.source) Source.source; | 
| 
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
 wenzelm parents: 
21861diff
changeset | 749 | |
| 25527 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 750 | local | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 751 | |
| 17513 
0393718c2f1c
get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
 wenzelm parents: 
17500diff
changeset | 752 | (*Spurious interrupts ahead! Race condition?*) | 
| 
0393718c2f1c
get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
 wenzelm parents: 
17500diff
changeset | 753 | fun get_interrupt src = SOME (Source.get_single src) handle Interrupt => NONE; | 
| 7602 | 754 | |
| 25527 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 755 | fun raw_loop secure src = | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 756 | let | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 757 | fun check_secure () = | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 758 | (if secure then warning "Secure loop -- cannot exit to ML" else (); secure); | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 759 | in | 
| 25847 | 760 | (case get_interrupt (Source.set_prompt Source.default_prompt src) of | 
| 25527 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 761 | NONE => (writeln "\nInterrupt."; raw_loop secure src) | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 762 | | SOME NONE => if secure then quit () else () | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 763 | | SOME (SOME (tr, src')) => if >> tr orelse check_secure () then raw_loop secure src' else ()) | 
| 25584 | 764 | handle exn => (CRITICAL (fn () => change crashes (cons exn)); | 
| 765 | warning "Recovering after Isar toplevel crash -- see also Toplevel.crashes"; | |
| 766 | raw_loop secure src) | |
| 23640 
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
 wenzelm parents: 
23619diff
changeset | 767 | end; | 
| 5828 | 768 | |
| 25527 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 769 | in | 
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 770 | |
| 26081 | 771 | fun loop secure src = uninterruptible (fn _ => raw_loop secure) src; | 
| 5828 | 772 | |
| 773 | end; | |
| 25527 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 774 | |
| 
330ca6e1dca8
Toplevel.loop: explicit argument for secure loop, no warning on quit;
 wenzelm parents: 
25441diff
changeset | 775 | end; |