author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17513 | 0393718c2f1c |
child 17904 | 21c6894b5998 |
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 |
|
16815 | 10 |
datatype node = |
17076 | 11 |
Theory of theory * Proof.context option | |
12 |
Proof of ProofHistory.T | |
|
13 |
SkipProof of int History.T * theory |
|
5828 | 14 |
type state |
15 |
val toplevel: state |
|
7732 | 16 |
val is_toplevel: state -> bool |
17076 | 17 |
val level: state -> int |
6689 | 18 |
exception UNDEF |
17266 | 19 |
val assert: bool -> unit |
6689 | 20 |
val node_history_of: state -> node History.T |
5828 | 21 |
val node_of: state -> node |
6664 | 22 |
val node_case: (theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a |
5828 | 23 |
val theory_of: state -> theory |
16452 | 24 |
val sign_of: state -> theory (*obsolete*) |
17076 | 25 |
val body_context_of: state -> Proof.context |
17208 | 26 |
val no_body_context: state -> state |
5828 | 27 |
val proof_of: state -> Proof.state |
17076 | 28 |
val is_proof: state -> bool |
9453 | 29 |
val enter_forward_proof: state -> Proof.state |
16815 | 30 |
val prompt_state_default: state -> string |
31 |
val prompt_state_fn: (state -> string) ref |
|
32 |
val print_state_context: state -> unit |
|
33 |
val print_state_default: bool -> state -> unit |
|
34 |
val print_state_hook: (bool -> state -> unit) -> unit |
|
35 |
val print_state_fn: (bool -> state -> unit) ref |
|
36 |
val print_state: bool -> state -> unit |
|
37 |
val pretty_state: bool -> state -> Pretty.T list |
|
16682 | 38 |
val quiet: bool ref |
39 |
val debug: bool ref |
|
17321
227f1f5c3959
added interact flag to control mode of excursions;
wenzelm
parents:
17320
diff
changeset
|
40 |
val interact: bool ref |
16682 | 41 |
val timing: bool ref |
42 |
val profiling: int ref |
|
16815 | 43 |
val skip_proofs: bool ref |
5828 | 44 |
exception TERMINATE |
5990 | 45 |
exception RESTART |
16682 | 46 |
type transition |
6689 | 47 |
val undo_limit: bool -> int option |
5828 | 48 |
val empty: transition |
14923 | 49 |
val name_of: transition -> string |
50 |
val source_of: transition -> OuterLex.token list option |
|
5828 | 51 |
val name: string -> transition -> transition |
52 |
val position: Position.T -> transition -> transition |
|
14923 | 53 |
val source: OuterLex.token list -> transition -> transition |
5828 | 54 |
val interactive: bool -> transition -> transition |
55 |
val print: transition -> transition |
|
16607 | 56 |
val print': string -> transition -> transition |
17363 | 57 |
val three_buffersN: string |
58 |
val print3: transition -> transition |
|
9010 | 59 |
val no_timing: transition -> transition |
5828 | 60 |
val reset: transition -> transition |
8465 | 61 |
val init: (bool -> node) -> (node -> unit) -> (node -> unit) -> transition -> transition |
6689 | 62 |
val exit: transition -> transition |
63 |
val kill: transition -> transition |
|
5828 | 64 |
val keep: (state -> unit) -> transition -> transition |
7612 | 65 |
val keep': (bool -> state -> unit) -> transition -> transition |
6689 | 66 |
val history: (node History.T -> node History.T) -> transition -> transition |
5828 | 67 |
val imperative: (unit -> unit) -> transition -> transition |
7105 | 68 |
val init_theory: (bool -> theory) -> (theory -> unit) -> (theory -> unit) |
6689 | 69 |
-> transition -> transition |
5828 | 70 |
val theory: (theory -> theory) -> transition -> transition |
7612 | 71 |
val theory': (bool -> theory -> theory) -> transition -> transition |
17076 | 72 |
val theory_context: (theory -> theory * Proof.context) -> transition -> transition |
17363 | 73 |
val theory_to_proof: (theory -> Proof.state) -> transition -> transition |
5828 | 74 |
val proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition |
6689 | 75 |
val proof': (bool -> ProofHistory.T -> ProofHistory.T) -> transition -> transition |
16815 | 76 |
val actual_proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition |
77 |
val skip_proof: (int History.T -> int History.T) -> transition -> transition |
|
5828 | 78 |
val proof_to_theory: (ProofHistory.T -> theory) -> transition -> transition |
12315 | 79 |
val proof_to_theory': (bool -> ProofHistory.T -> theory) -> transition -> transition |
17115 | 80 |
val proof_to_theory_context: (bool -> ProofHistory.T -> theory * Proof.context) |
17076 | 81 |
-> transition -> transition |
15237
250e9be7a09d
Some changes to allow skipping of proof scripts.
berghofe
parents:
14985
diff
changeset
|
82 |
val skip_proof_to_theory: (int History.T -> bool) -> transition -> transition |
9512 | 83 |
val unknown_theory: transition -> transition |
84 |
val unknown_proof: transition -> transition |
|
85 |
val unknown_context: transition -> transition |
|
5922 | 86 |
val exn_message: exn -> string |
5828 | 87 |
val apply: bool -> transition -> state -> (state * (exn * string) option) option |
17266 | 88 |
val command: transition -> state -> state |
17076 | 89 |
val present_excursion: (transition * (state -> state -> 'a -> 'a)) list -> 'a -> 'a |
5828 | 90 |
val excursion: transition list -> unit |
91 |
val set_state: state -> unit |
|
92 |
val get_state: unit -> state |
|
93 |
val exn: unit -> (exn * string) option |
|
94 |
val >> : transition -> bool |
|
14985 | 95 |
val >>> : transition list -> unit |
14091 | 96 |
type 'a isar |
97 |
val loop: 'a isar -> unit |
|
5828 | 98 |
end; |
99 |
||
6965 | 100 |
structure Toplevel: TOPLEVEL = |
5828 | 101 |
struct |
102 |
||
103 |
||
104 |
(** toplevel state **) |
|
105 |
||
16815 | 106 |
(* datatype state *) |
5828 | 107 |
|
108 |
datatype node = |
|
17076 | 109 |
Theory of theory * Proof.context option | (*theory with optional body context*) |
110 |
Proof of ProofHistory.T | (*history of proof states*) |
|
111 |
SkipProof of int History.T * theory; (*history of proof depths*) |
|
5828 | 112 |
|
8465 | 113 |
datatype state = State of (node History.T * ((node -> unit) * (node -> unit))) option; |
5828 | 114 |
|
15531 | 115 |
val toplevel = State NONE; |
5828 | 116 |
|
15531 | 117 |
fun is_toplevel (State NONE) = true |
7732 | 118 |
| is_toplevel _ = false; |
119 |
||
17076 | 120 |
fun level (State NONE) = 0 |
121 |
| level (State (SOME (node, _))) = |
|
122 |
(case History.current node of |
|
123 |
Theory _ => 0 |
|
124 |
| Proof prf => Proof.level (ProofHistory.current prf) |
|
125 |
| SkipProof (h, _) => History.current h + 1); (*different notion of proof depth!*) |
|
126 |
||
15531 | 127 |
fun str_of_state (State NONE) = "at top level" |
16815 | 128 |
| str_of_state (State (SOME (node, _))) = |
129 |
(case History.current node of |
|
130 |
Theory _ => "in theory mode" |
|
131 |
| Proof _ => "in proof mode" |
|
132 |
| SkipProof _ => "in skipped proof mode"); |
|
5946
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
133 |
|
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
134 |
|
5828 | 135 |
(* top node *) |
136 |
||
6689 | 137 |
exception UNDEF; |
138 |
||
17266 | 139 |
fun assert true = () |
140 |
| assert false = raise UNDEF; |
|
141 |
||
15531 | 142 |
fun node_history_of (State NONE) = raise UNDEF |
143 |
| node_history_of (State (SOME (node, _))) = node; |
|
6689 | 144 |
|
145 |
val node_of = History.current o node_history_of; |
|
5828 | 146 |
|
6664 | 147 |
fun node_case f g state = |
5828 | 148 |
(case node_of state of |
17076 | 149 |
Theory (thy, _) => f thy |
16815 | 150 |
| Proof prf => g (ProofHistory.current prf) |
151 |
| SkipProof (_, thy) => f thy); |
|
5828 | 152 |
|
6664 | 153 |
val theory_of = node_case I Proof.theory_of; |
16452 | 154 |
val sign_of = theory_of; |
17076 | 155 |
|
156 |
fun body_context_of state = |
|
157 |
(case node_of state of |
|
158 |
Theory (_, SOME ctxt) => ctxt |
|
159 |
| _ => node_case ProofContext.init Proof.context_of state); |
|
160 |
||
17208 | 161 |
fun no_body_context (State NONE) = State NONE |
162 |
| no_body_context (State (SOME (node, x))) = |
|
163 |
State (SOME (History.apply (fn Theory (thy, _) => Theory (thy, NONE) | nd => nd) node, x)); |
|
164 |
||
6664 | 165 |
val proof_of = node_case (fn _ => raise UNDEF) I; |
17076 | 166 |
val is_proof = can proof_of; |
6664 | 167 |
|
17363 | 168 |
val enter_forward_proof = node_case (Proof.init o ProofContext.init) Proof.enter_forward; |
5828 | 169 |
|
170 |
||
16815 | 171 |
(* prompt state *) |
172 |
||
173 |
fun prompt_state_default (State _) = Source.default_prompt; |
|
174 |
||
175 |
val prompt_state_fn = ref prompt_state_default; |
|
176 |
fun prompt_state state = ! prompt_state_fn state; |
|
177 |
||
178 |
||
179 |
(* print state *) |
|
180 |
||
181 |
fun pretty_context thy = [Pretty.block |
|
182 |
[Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy), |
|
183 |
Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]]; |
|
184 |
||
185 |
fun pretty_state_context state = |
|
186 |
(case try theory_of state of NONE => [] |
|
187 |
| SOME thy => pretty_context thy); |
|
188 |
||
17076 | 189 |
fun pretty_node prf_only (Theory (thy, _)) = if prf_only then [] else pretty_context thy |
16815 | 190 |
| pretty_node _ (Proof prf) = |
191 |
Proof.pretty_state (ProofHistory.position prf) (ProofHistory.current prf) |
|
192 |
| pretty_node _ (SkipProof (h, _)) = |
|
193 |
[Pretty.str ("skipped proof: depth " ^ string_of_int (History.current h))]; |
|
194 |
||
195 |
fun pretty_state prf_only state = |
|
196 |
let val ref (begin_state, end_state, _) = Display.current_goals_markers in |
|
197 |
(case try node_of state of NONE => [] |
|
198 |
| SOME node => |
|
199 |
(if begin_state = "" then [] else [Pretty.str begin_state]) @ |
|
200 |
pretty_node prf_only node @ |
|
201 |
(if end_state = "" then [] else [Pretty.str end_state])) |
|
202 |
end; |
|
203 |
||
204 |
val print_state_context = Pretty.writelns o pretty_state_context; |
|
205 |
fun print_state_default prf_only state = Pretty.writelns (pretty_state prf_only state); |
|
206 |
||
207 |
val print_state_hooks = ref ([]: (bool -> state -> unit) list); |
|
208 |
fun print_state_hook f = change print_state_hooks (cons f); |
|
209 |
val print_state_fn = ref print_state_default; |
|
210 |
||
211 |
fun print_state prf_only state = |
|
212 |
(List.app (fn f => (try (fn () => f prf_only state) (); ())) (! print_state_hooks); |
|
213 |
! print_state_fn prf_only state); |
|
214 |
||
215 |
||
15668 | 216 |
|
5828 | 217 |
(** toplevel transitions **) |
218 |
||
16682 | 219 |
val quiet = ref false; |
220 |
val debug = ref false; |
|
17321
227f1f5c3959
added interact flag to control mode of excursions;
wenzelm
parents:
17320
diff
changeset
|
221 |
val interact = ref false; |
16682 | 222 |
val timing = Output.timing; |
223 |
val profiling = ref 0; |
|
16815 | 224 |
val skip_proofs = ref false; |
16682 | 225 |
|
5828 | 226 |
exception TERMINATE; |
5990 | 227 |
exception RESTART; |
7022 | 228 |
exception EXCURSION_FAIL of exn * string; |
6689 | 229 |
exception FAILURE of state * exn; |
5828 | 230 |
|
231 |
||
16815 | 232 |
(* node transactions and recovery from stale theories *) |
6689 | 233 |
|
16815 | 234 |
(*NB: proof commands should be non-destructive!*) |
7022 | 235 |
|
6689 | 236 |
local |
237 |
||
16452 | 238 |
fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false; |
6689 | 239 |
|
16815 | 240 |
val stale_theory = ERROR_MESSAGE "Stale theory encountered after succesful execution!"; |
241 |
||
17076 | 242 |
fun checkpoint_node true (Theory (thy, ctxt)) = Theory (Theory.checkpoint thy, ctxt) |
6689 | 243 |
| checkpoint_node _ node = node; |
244 |
||
17076 | 245 |
fun copy_node true (Theory (thy, ctxt)) = Theory (Theory.copy thy, ctxt) |
6689 | 246 |
| copy_node _ node = node; |
247 |
||
15531 | 248 |
fun return (result, NONE) = result |
249 |
| return (result, SOME exn) = raise FAILURE (result, exn); |
|
7022 | 250 |
|
16729 | 251 |
fun debug_trans f x = |
252 |
if ! debug then |
|
16815 | 253 |
setmp Output.transform_exceptions false |
254 |
exception_trace (fn () => f x) |
|
16729 | 255 |
else f x; |
256 |
||
257 |
fun interruptible f x = |
|
258 |
let val y = ref x |
|
259 |
in raise_interrupt (fn () => y := f x) (); ! y end; |
|
260 |
||
6689 | 261 |
in |
262 |
||
16815 | 263 |
fun transaction _ _ _ (State NONE) = raise UNDEF |
264 |
| transaction int hist f (State (SOME (node, term))) = |
|
6689 | 265 |
let |
16046
7d68cd1b1b8b
node_trans: revert to original transaction code (pre 1.54);
wenzelm
parents:
15973
diff
changeset
|
266 |
val cont_node = History.map (checkpoint_node int) node; |
7d68cd1b1b8b
node_trans: revert to original transaction code (pre 1.54);
wenzelm
parents:
15973
diff
changeset
|
267 |
val back_node = History.map (copy_node int) cont_node; |
16729 | 268 |
fun state nd = State (SOME (nd, term)); |
269 |
fun normal_state nd = (state nd, NONE); |
|
270 |
fun error_state nd exn = (state nd, SOME exn); |
|
6689 | 271 |
|
16729 | 272 |
val (result, err) = |
273 |
cont_node |
|
274 |
|> ((fn nd => f int nd) |
|
275 |
|> (if hist then History.apply_copy (copy_node int) else History.map) |
|
276 |
|> debug_trans |
|
277 |
|> interruptible |
|
278 |
|> transform_error) |
|
279 |
|> normal_state |
|
280 |
handle exn => error_state cont_node exn; |
|
6689 | 281 |
in |
16729 | 282 |
if is_stale result |
283 |
then return (error_state back_node (if_none err stale_theory)) |
|
284 |
else return (result, err) |
|
6689 | 285 |
end; |
286 |
||
287 |
end; |
|
288 |
||
289 |
||
290 |
(* primitive transitions *) |
|
291 |
||
16815 | 292 |
(*NB: recovery from stale theories is provided only for theory-level |
293 |
operations via MapCurrent and AppCurrent. Other node or state |
|
294 |
operations should not touch theories at all. |
|
6965 | 295 |
|
16815 | 296 |
Interrupts are enabled only for Keep, MapCurrent, and AppCurrent.*) |
5828 | 297 |
|
298 |
datatype trans = |
|
6689 | 299 |
Reset | (*empty toplevel*) |
8465 | 300 |
Init of (bool -> node) * ((node -> unit) * (node -> unit)) | |
301 |
(*init node; provide exit/kill operation*) |
|
302 |
Exit | (*conclude node*) |
|
6689 | 303 |
Kill | (*abort node*) |
7612 | 304 |
Keep of bool -> state -> unit | (*peek at state*) |
6689 | 305 |
History of node History.T -> node History.T | (*history operation (undo etc.)*) |
306 |
MapCurrent of bool -> node -> node | (*change node, bypassing history*) |
|
307 |
AppCurrent of bool -> node -> node; (*change node, recording history*) |
|
308 |
||
15531 | 309 |
fun undo_limit int = if int then NONE else SOME 0; |
6689 | 310 |
|
311 |
local |
|
5828 | 312 |
|
6689 | 313 |
fun apply_tr _ Reset _ = toplevel |
15531 | 314 |
| apply_tr int (Init (f, term)) (State NONE) = |
315 |
State (SOME (History.init (undo_limit int) (f int), term)) |
|
316 |
| apply_tr _ (Init _ ) (State (SOME _)) = raise UNDEF |
|
317 |
| apply_tr _ Exit (State NONE) = raise UNDEF |
|
318 |
| apply_tr _ Exit (State (SOME (node, (exit, _)))) = |
|
319 |
(exit (History.current node); State NONE) |
|
320 |
| apply_tr _ Kill (State NONE) = raise UNDEF |
|
321 |
| apply_tr _ Kill (State (SOME (node, (_, kill)))) = |
|
322 |
(kill (History.current node); State NONE) |
|
12987 | 323 |
| apply_tr int (Keep f) state = (raise_interrupt (f int) state; state) |
15531 | 324 |
| apply_tr _ (History _) (State NONE) = raise UNDEF |
325 |
| apply_tr _ (History f) (State (SOME (node, term))) = State (SOME (f node, term)) |
|
16815 | 326 |
| apply_tr int (MapCurrent f) state = transaction int false f state |
327 |
| apply_tr int (AppCurrent f) state = transaction int true f state; |
|
5828 | 328 |
|
6689 | 329 |
fun apply_union _ [] state = raise FAILURE (state, UNDEF) |
330 |
| apply_union int (tr :: trs) state = |
|
331 |
transform_error (apply_tr int tr) state |
|
332 |
handle UNDEF => apply_union int trs state |
|
333 |
| FAILURE (alt_state, UNDEF) => apply_union int trs alt_state |
|
334 |
| exn as FAILURE _ => raise exn |
|
335 |
| exn => raise FAILURE (state, exn); |
|
336 |
||
337 |
in |
|
338 |
||
15531 | 339 |
fun apply_trans int trs state = (apply_union int trs state, NONE) |
340 |
handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); |
|
6689 | 341 |
|
342 |
end; |
|
5828 | 343 |
|
344 |
||
345 |
(* datatype transition *) |
|
346 |
||
347 |
datatype transition = Transition of |
|
16815 | 348 |
{name: string, (*command name*) |
349 |
pos: Position.T, (*source position*) |
|
350 |
source: OuterLex.token list option, (*source text*) |
|
351 |
int_only: bool, (*interactive-only*) |
|
352 |
print: string list, (*print modes (union)*) |
|
353 |
no_timing: bool, (*suppress timing*) |
|
354 |
trans: trans list}; (*primitive transitions (union)*) |
|
5828 | 355 |
|
14923 | 356 |
fun make_transition (name, pos, source, int_only, print, no_timing, trans) = |
357 |
Transition {name = name, pos = pos, source = source, |
|
358 |
int_only = int_only, print = print, no_timing = no_timing, trans = trans}; |
|
5828 | 359 |
|
14923 | 360 |
fun map_transition f (Transition {name, pos, source, int_only, print, no_timing, trans}) = |
361 |
make_transition (f (name, pos, source, int_only, print, no_timing, trans)); |
|
5828 | 362 |
|
16607 | 363 |
val empty = make_transition ("<unknown>", Position.none, NONE, false, [], false, []); |
14923 | 364 |
|
365 |
fun name_of (Transition {name, ...}) = name; |
|
366 |
fun source_of (Transition {source, ...}) = source; |
|
5828 | 367 |
|
368 |
||
369 |
(* diagnostics *) |
|
370 |
||
371 |
fun str_of_transition (Transition {name, pos, ...}) = quote name ^ Position.str_of pos; |
|
372 |
||
373 |
fun command_msg msg tr = msg ^ "command " ^ str_of_transition tr; |
|
374 |
fun at_command tr = command_msg "At " tr ^ "."; |
|
375 |
||
376 |
fun type_error tr state = |
|
6689 | 377 |
ERROR_MESSAGE (command_msg "Illegal application of " tr ^ " " ^ str_of_state state); |
5828 | 378 |
|
379 |
||
380 |
(* modify transitions *) |
|
381 |
||
14923 | 382 |
fun name nm = map_transition (fn (_, pos, source, int_only, print, no_timing, trans) => |
383 |
(nm, pos, source, int_only, print, no_timing, trans)); |
|
5828 | 384 |
|
14923 | 385 |
fun position pos = map_transition (fn (name, _, source, int_only, print, no_timing, trans) => |
386 |
(name, pos, source, int_only, print, no_timing, trans)); |
|
9010 | 387 |
|
14923 | 388 |
fun source src = map_transition (fn (name, pos, _, int_only, print, no_timing, trans) => |
15531 | 389 |
(name, pos, SOME src, int_only, print, no_timing, trans)); |
5828 | 390 |
|
14923 | 391 |
fun interactive int_only = map_transition (fn (name, pos, source, _, print, no_timing, trans) => |
392 |
(name, pos, source, int_only, print, no_timing, trans)); |
|
393 |
||
17363 | 394 |
val no_timing = map_transition (fn (name, pos, source, int_only, print, _, trans) => |
395 |
(name, pos, source, int_only, print, true, trans)); |
|
396 |
||
397 |
fun add_trans tr = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => |
|
398 |
(name, pos, source, int_only, print, no_timing, trans @ [tr])); |
|
399 |
||
16607 | 400 |
fun print' mode = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => |
401 |
(name, pos, source, int_only, insert (op =) mode print, no_timing, trans)); |
|
402 |
||
403 |
val print = print' ""; |
|
5828 | 404 |
|
17363 | 405 |
val three_buffersN = "three_buffers"; |
406 |
val print3 = print' three_buffersN; |
|
5828 | 407 |
|
408 |
||
409 |
(* build transitions *) |
|
410 |
||
411 |
val reset = add_trans Reset; |
|
6689 | 412 |
fun init f exit kill = add_trans (Init (f, (exit, kill))); |
413 |
val exit = add_trans Exit; |
|
414 |
val kill = add_trans Kill; |
|
7612 | 415 |
val keep' = add_trans o Keep; |
6689 | 416 |
val history = add_trans o History; |
5828 | 417 |
val map_current = add_trans o MapCurrent; |
6689 | 418 |
val app_current = add_trans o AppCurrent; |
5828 | 419 |
|
7612 | 420 |
fun keep f = add_trans (Keep (fn _ => f)); |
5828 | 421 |
fun imperative f = keep (fn _ => f ()); |
422 |
||
6689 | 423 |
fun init_theory f exit kill = |
17076 | 424 |
init (fn int => Theory (f int, NONE)) |
425 |
(fn Theory (thy, _) => exit thy | _ => raise UNDEF) |
|
426 |
(fn Theory (thy, _) => kill thy | _ => raise UNDEF); |
|
5828 | 427 |
|
15668 | 428 |
|
429 |
(* typed transitions *) |
|
430 |
||
17076 | 431 |
fun theory f = app_current (K (fn Theory (thy, _) => Theory (f thy, NONE) | _ => raise UNDEF)); |
432 |
fun theory' f = app_current (fn int => |
|
433 |
(fn Theory (thy, _) => Theory (f int thy, NONE) | _ => raise UNDEF)); |
|
434 |
||
435 |
fun theory_context f = |
|
436 |
app_current (K (fn Theory (thy, _) => Theory (apsnd SOME (f thy)) | _ => raise UNDEF)); |
|
15668 | 437 |
|
15237
250e9be7a09d
Some changes to allow skipping of proof scripts.
berghofe
parents:
14985
diff
changeset
|
438 |
fun theory_to_proof f = app_current (fn int => |
17076 | 439 |
(fn Theory (thy, _) => |
17452
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
440 |
let val st = f thy in |
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
441 |
if Theory.eq_thy (thy, Proof.theory_of st) then () |
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
442 |
else error "Theory changed at start of proof"; |
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
443 |
if ! skip_proofs then |
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
444 |
SkipProof (History.init (undo_limit int) 0, #1 (Proof.global_skip_proof int st)) |
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
445 |
else Proof (ProofHistory.init (undo_limit int) st) |
178344c74562
theory_to_proof: check theory of initial proof state, which must not be changed;
wenzelm
parents:
17434
diff
changeset
|
446 |
end |
15237
250e9be7a09d
Some changes to allow skipping of proof scripts.
berghofe
parents:
14985
diff
changeset
|
447 |
| _ => raise UNDEF)); |
15668 | 448 |
|
16815 | 449 |
fun proof' f = map_current (fn int => |
450 |
(fn Proof prf => Proof (f int prf) |
|
451 |
| SkipProof (h, thy) => SkipProof (History.apply I h, thy) (*approximate f*) |
|
452 |
| _ => raise UNDEF)); |
|
15668 | 453 |
|
6689 | 454 |
val proof = proof' o K; |
16815 | 455 |
|
456 |
fun actual_proof f = map_current (fn _ => |
|
457 |
(fn Proof prf => Proof (f prf) | _ => raise UNDEF)); |
|
458 |
||
459 |
fun skip_proof f = map_current (fn _ => |
|
460 |
(fn SkipProof (h, thy) => SkipProof (f h, thy) | _ => raise UNDEF)); |
|
15668 | 461 |
|
17076 | 462 |
fun end_proof f = |
15237
250e9be7a09d
Some changes to allow skipping of proof scripts.
berghofe
parents:
14985
diff
changeset
|
463 |
map_current (fn int => (fn Proof prf => Theory (f int prf) |
17076 | 464 |
| SkipProof (h, thy) => if History.current h = 0 then Theory (thy, NONE) else raise UNDEF |
15237
250e9be7a09d
Some changes to allow skipping of proof scripts.
berghofe
parents:
14985
diff
changeset
|
465 |
| _ => raise UNDEF)); |
15668 | 466 |
|
17076 | 467 |
fun proof_to_theory' f = end_proof (rpair NONE oo f); |
468 |
fun proof_to_theory f = proof_to_theory' (K f); |
|
17115 | 469 |
fun proof_to_theory_context f = end_proof (apsnd SOME oo f); |
15668 | 470 |
|
16815 | 471 |
fun skip_proof_to_theory p = map_current (fn _ => |
17076 | 472 |
(fn SkipProof (h, thy) => if p h then Theory (thy, NONE) else raise UNDEF | _ => raise UNDEF)); |
5828 | 473 |
|
9512 | 474 |
val unknown_theory = imperative (fn () => warning "Unknown theory context"); |
475 |
val unknown_proof = imperative (fn () => warning "Unknown proof context"); |
|
476 |
val unknown_context = imperative (fn () => warning "Unknown context"); |
|
477 |
||
5828 | 478 |
|
479 |
||
480 |
(** toplevel transactions **) |
|
481 |
||
482 |
(* print exceptions *) |
|
483 |
||
15668 | 484 |
local |
485 |
||
486 |
fun with_context f xs = |
|
487 |
(case Context.get_context () of NONE => [] |
|
16452 | 488 |
| SOME thy => map (f thy) xs); |
15668 | 489 |
|
490 |
fun raised name [] = "exception " ^ name ^ " raised" |
|
491 |
| raised name [msg] = "exception " ^ name ^ " raised: " ^ msg |
|
492 |
| raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs); |
|
5828 | 493 |
|
15668 | 494 |
fun exn_msg _ TERMINATE = "Exit." |
495 |
| exn_msg _ RESTART = "Restart." |
|
496 |
| exn_msg _ Interrupt = "Interrupt." |
|
497 |
| exn_msg _ ERROR = "ERROR." |
|
498 |
| exn_msg _ (ERROR_MESSAGE msg) = msg |
|
17363 | 499 |
| exn_msg detailed (EXCEPTION (exn, msg)) = cat_lines [exn_msg detailed exn, msg] |
16452 | 500 |
| exn_msg detailed (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_msg detailed exn, msg] |
15668 | 501 |
| exn_msg false (THEORY (msg, _)) = msg |
16452 | 502 |
| exn_msg true (THEORY (msg, thys)) = raised "THEORY" (msg :: map Context.str_of_thy thys) |
15668 | 503 |
| exn_msg _ (ProofContext.CONTEXT (msg, _)) = msg |
504 |
| exn_msg _ (Proof.STATE (msg, _)) = msg |
|
505 |
| exn_msg _ (ProofHistory.FAIL msg) = msg |
|
506 |
| exn_msg detailed (MetaSimplifier.SIMPROC_FAIL (name, exn)) = |
|
507 |
fail_msg detailed "simproc" ((name, Position.none), exn) |
|
508 |
| exn_msg detailed (Attrib.ATTRIB_FAIL info) = fail_msg detailed "attribute" info |
|
509 |
| exn_msg detailed (Method.METHOD_FAIL info) = fail_msg detailed "method" info |
|
510 |
| exn_msg detailed (Antiquote.ANTIQUOTE_FAIL info) = fail_msg detailed "antiquotation" info |
|
511 |
| exn_msg false (Syntax.AST (msg, _)) = raised "AST" [msg] |
|
512 |
| exn_msg true (Syntax.AST (msg, asts)) = |
|
513 |
raised "AST" (msg :: map (Pretty.string_of o Syntax.pretty_ast) asts) |
|
514 |
| exn_msg false (TYPE (msg, _, _)) = raised "TYPE" [msg] |
|
515 |
| exn_msg true (TYPE (msg, Ts, ts)) = raised "TYPE" (msg :: |
|
516 |
with_context Sign.string_of_typ Ts @ with_context Sign.string_of_term ts) |
|
517 |
| exn_msg false (TERM (msg, _)) = raised "TERM" [msg] |
|
518 |
| exn_msg true (TERM (msg, ts)) = raised "TERM" (msg :: with_context Sign.string_of_term ts) |
|
519 |
| exn_msg false (THM (msg, _, _)) = raised "THM" [msg] |
|
520 |
| exn_msg true (THM (msg, i, thms)) = |
|
521 |
raised ("THM " ^ string_of_int i) (msg :: map Display.string_of_thm thms) |
|
522 |
| exn_msg _ Option = raised "Option" [] |
|
523 |
| exn_msg _ UnequalLengths = raised "UnequalLengths" [] |
|
524 |
| exn_msg _ Empty = raised "Empty" [] |
|
525 |
| exn_msg _ Subscript = raised "Subscript" [] |
|
16197 | 526 |
| exn_msg _ (Fail msg) = raised "Fail" [msg] |
15668 | 527 |
| exn_msg _ exn = General.exnMessage exn |
528 |
and fail_msg detailed kind ((name, pos), exn) = |
|
529 |
"Error in " ^ kind ^ " " ^ quote name ^ Position.str_of pos ^ ":\n" ^ exn_msg detailed exn; |
|
15519 | 530 |
|
15668 | 531 |
in |
532 |
||
533 |
fun exn_message exn = exn_msg (! debug) exn; |
|
5828 | 534 |
|
15531 | 535 |
fun print_exn NONE = () |
536 |
| print_exn (SOME (exn, s)) = error_msg (cat_lines [exn_message exn, s]); |
|
5828 | 537 |
|
15668 | 538 |
end; |
539 |
||
5828 | 540 |
|
541 |
(* apply transitions *) |
|
542 |
||
6664 | 543 |
local |
544 |
||
9010 | 545 |
fun app int (tr as Transition {trans, int_only, print, no_timing, ...}) state = |
5828 | 546 |
let |
16682 | 547 |
val _ = conditional (not int andalso int_only) (fn () => |
548 |
warning (command_msg "Interactive-only " tr)); |
|
549 |
||
550 |
fun do_timing f x = (info (command_msg "" tr); timeap f x); |
|
551 |
fun do_profiling f x = profile (! profiling) f x; |
|
552 |
||
6689 | 553 |
val (result, opt_exn) = |
16729 | 554 |
state |> (apply_trans int trans |
555 |
|> (if ! profiling > 0 then do_profiling else I) |
|
556 |
|> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I)); |
|
16607 | 557 |
val _ = conditional (int andalso not (! quiet) andalso |
558 |
exists (fn m => m mem_string print) ("" :: ! print_mode)) |
|
559 |
(fn () => print_state false result); |
|
15570 | 560 |
in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_exn) end; |
6664 | 561 |
|
562 |
in |
|
5828 | 563 |
|
6689 | 564 |
fun apply int tr st = |
6965 | 565 |
(case app int tr st of |
15531 | 566 |
(_, SOME TERMINATE) => NONE |
567 |
| (_, SOME RESTART) => SOME (toplevel, NONE) |
|
568 |
| (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info) |
|
569 |
| (state', SOME exn) => SOME (state', SOME (exn, at_command tr)) |
|
570 |
| (state', NONE) => SOME (state', NONE)); |
|
6664 | 571 |
|
17266 | 572 |
fun command tr st = |
573 |
(case apply false tr st of |
|
574 |
SOME (st', NONE) => st' |
|
575 |
| SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info |
|
576 |
| NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr)); |
|
577 |
||
6664 | 578 |
end; |
5828 | 579 |
|
580 |
||
17076 | 581 |
(* excursion: toplevel -- apply transformers/presentation -- toplevel *) |
5828 | 582 |
|
6664 | 583 |
local |
584 |
||
5828 | 585 |
fun excur [] x = x |
17076 | 586 |
| excur ((tr, pr) :: trs) (st, res) = |
17321
227f1f5c3959
added interact flag to control mode of excursions;
wenzelm
parents:
17320
diff
changeset
|
587 |
(case apply (! interact) tr st of |
15531 | 588 |
SOME (st', NONE) => |
17076 | 589 |
excur trs (st', transform_error (fn () => pr st st' res) () handle exn => |
10324 | 590 |
raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr)) |
15531 | 591 |
| SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info |
592 |
| NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr)); |
|
5828 | 593 |
|
17076 | 594 |
fun no_pr _ _ _ = (); |
595 |
||
6664 | 596 |
in |
597 |
||
17076 | 598 |
fun present_excursion trs res = |
15531 | 599 |
(case excur trs (State NONE, res) of |
600 |
(State NONE, res') => res' |
|
9134 | 601 |
| _ => raise ERROR_MESSAGE "Unfinished development at end of input") |
602 |
handle exn => error (exn_message exn); |
|
603 |
||
17076 | 604 |
fun excursion trs = present_excursion (map (rpair no_pr) trs) (); |
7062 | 605 |
|
6664 | 606 |
end; |
607 |
||
5828 | 608 |
|
609 |
||
610 |
(** interactive transformations **) |
|
611 |
||
612 |
(* the global state reference *) |
|
613 |
||
15531 | 614 |
val global_state = ref (toplevel, NONE: (exn * string) option); |
5828 | 615 |
|
15531 | 616 |
fun set_state state = global_state := (state, NONE); |
5828 | 617 |
fun get_state () = fst (! global_state); |
618 |
fun exn () = snd (! global_state); |
|
619 |
||
620 |
||
6965 | 621 |
(* the Isar source of transitions *) |
622 |
||
14091 | 623 |
type 'a isar = |
6965 | 624 |
(transition, (transition option, |
12881 | 625 |
(OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token, |
14091 | 626 |
Position.T * (Symbol.symbol, (string, 'a) Source.source) Source.source) |
12881 | 627 |
Source.source) Source.source) Source.source) Source.source) Source.source) Source.source; |
6965 | 628 |
|
629 |
||
5828 | 630 |
(* apply transformers to global state *) |
631 |
||
14985 | 632 |
nonfix >> >>>; |
5828 | 633 |
|
634 |
fun >> tr = |
|
635 |
(case apply true tr (get_state ()) of |
|
15531 | 636 |
NONE => false |
637 |
| SOME (state', exn_info) => |
|
5828 | 638 |
(global_state := (state', exn_info); |
16729 | 639 |
print_exn exn_info; |
5828 | 640 |
true)); |
641 |
||
14985 | 642 |
fun >>> [] = () |
643 |
| >>> (tr :: trs) = if >> tr then >>> trs else (); |
|
644 |
||
17513
0393718c2f1c
get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
wenzelm
parents:
17500
diff
changeset
|
645 |
(*Spurious interrupts ahead! Race condition?*) |
0393718c2f1c
get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
wenzelm
parents:
17500
diff
changeset
|
646 |
fun get_interrupt src = SOME (Source.get_single src) handle Interrupt => NONE; |
7602 | 647 |
|
5828 | 648 |
fun raw_loop src = |
7602 | 649 |
(case get_interrupt (Source.set_prompt (prompt_state (get_state ())) src) of |
15531 | 650 |
NONE => (writeln "\nInterrupt."; raw_loop src) |
651 |
| SOME NONE => () |
|
652 |
| SOME (SOME (tr, src')) => if >> tr then raw_loop src' else ()); |
|
5828 | 653 |
|
12987 | 654 |
fun loop src = ignore_interrupt raw_loop src; |
5828 | 655 |
|
656 |
end; |