wenzelm@5828
|
1 |
(* Title: Pure/Isar/toplevel.ML
|
wenzelm@5828
|
2 |
ID: $Id$
|
wenzelm@5828
|
3 |
Author: Markus Wenzel, TU Muenchen
|
wenzelm@5828
|
4 |
|
wenzelm@5828
|
5 |
The Isabelle/Isar toplevel.
|
wenzelm@5828
|
6 |
*)
|
wenzelm@5828
|
7 |
|
wenzelm@5828
|
8 |
signature TOPLEVEL =
|
wenzelm@5828
|
9 |
sig
|
wenzelm@5828
|
10 |
datatype node = Theory of theory | Proof of ProofHistory.T
|
wenzelm@5828
|
11 |
type state
|
wenzelm@5828
|
12 |
val toplevel: state
|
wenzelm@7732
|
13 |
val is_toplevel: state -> bool
|
wenzelm@6689
|
14 |
val prompt_state_default: state -> string
|
wenzelm@6689
|
15 |
val prompt_state_fn: (state -> string) ref
|
wenzelm@7308
|
16 |
val print_state_context: state -> unit
|
wenzelm@5946
|
17 |
val print_state_default: state -> unit
|
wenzelm@5946
|
18 |
val print_state_fn: (state -> unit) ref
|
wenzelm@5828
|
19 |
val print_state: state -> unit
|
wenzelm@6689
|
20 |
exception UNDEF
|
wenzelm@6689
|
21 |
val node_history_of: state -> node History.T
|
wenzelm@5828
|
22 |
val node_of: state -> node
|
wenzelm@6664
|
23 |
val node_case: (theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a
|
wenzelm@5828
|
24 |
val theory_of: state -> theory
|
wenzelm@5828
|
25 |
val sign_of: state -> Sign.sg
|
wenzelm@7501
|
26 |
val context_of: state -> Proof.context
|
wenzelm@5828
|
27 |
val proof_of: state -> Proof.state
|
wenzelm@5828
|
28 |
type transition
|
wenzelm@5828
|
29 |
exception TERMINATE
|
wenzelm@5990
|
30 |
exception RESTART
|
wenzelm@6689
|
31 |
val undo_limit: bool -> int option
|
wenzelm@5828
|
32 |
val empty: transition
|
wenzelm@5828
|
33 |
val name: string -> transition -> transition
|
wenzelm@5828
|
34 |
val position: Position.T -> transition -> transition
|
wenzelm@5828
|
35 |
val interactive: bool -> transition -> transition
|
wenzelm@5828
|
36 |
val print: transition -> transition
|
wenzelm@5828
|
37 |
val reset: transition -> transition
|
wenzelm@8465
|
38 |
val init: (bool -> node) -> (node -> unit) -> (node -> unit) -> transition -> transition
|
wenzelm@6689
|
39 |
val exit: transition -> transition
|
wenzelm@6689
|
40 |
val kill: transition -> transition
|
wenzelm@5828
|
41 |
val keep: (state -> unit) -> transition -> transition
|
wenzelm@7612
|
42 |
val keep': (bool -> state -> unit) -> transition -> transition
|
wenzelm@6689
|
43 |
val history: (node History.T -> node History.T) -> transition -> transition
|
wenzelm@5828
|
44 |
val imperative: (unit -> unit) -> transition -> transition
|
wenzelm@7105
|
45 |
val init_theory: (bool -> theory) -> (theory -> unit) -> (theory -> unit)
|
wenzelm@6689
|
46 |
-> transition -> transition
|
wenzelm@5828
|
47 |
val theory: (theory -> theory) -> transition -> transition
|
wenzelm@7612
|
48 |
val theory': (bool -> theory -> theory) -> transition -> transition
|
wenzelm@6689
|
49 |
val theory_to_proof: (bool -> theory -> ProofHistory.T) -> transition -> transition
|
wenzelm@5828
|
50 |
val proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition
|
wenzelm@6689
|
51 |
val proof': (bool -> ProofHistory.T -> ProofHistory.T) -> transition -> transition
|
wenzelm@5828
|
52 |
val proof_to_theory: (ProofHistory.T -> theory) -> transition -> transition
|
wenzelm@7198
|
53 |
val quiet: bool ref
|
wenzelm@5828
|
54 |
val trace: bool ref
|
wenzelm@5922
|
55 |
val exn_message: exn -> string
|
wenzelm@5828
|
56 |
val apply: bool -> transition -> state -> (state * (exn * string) option) option
|
wenzelm@5828
|
57 |
val excursion: transition list -> unit
|
wenzelm@7062
|
58 |
val excursion_error: transition list -> unit
|
wenzelm@5828
|
59 |
val set_state: state -> unit
|
wenzelm@5828
|
60 |
val get_state: unit -> state
|
wenzelm@5828
|
61 |
val exn: unit -> (exn * string) option
|
wenzelm@5828
|
62 |
val >> : transition -> bool
|
wenzelm@6965
|
63 |
type isar
|
wenzelm@5828
|
64 |
val loop: isar -> unit
|
wenzelm@5828
|
65 |
end;
|
wenzelm@5828
|
66 |
|
wenzelm@6965
|
67 |
structure Toplevel: TOPLEVEL =
|
wenzelm@5828
|
68 |
struct
|
wenzelm@5828
|
69 |
|
wenzelm@5828
|
70 |
|
wenzelm@5828
|
71 |
(** toplevel state **)
|
wenzelm@5828
|
72 |
|
wenzelm@5828
|
73 |
(* datatype node *)
|
wenzelm@5828
|
74 |
|
wenzelm@5828
|
75 |
datatype node =
|
wenzelm@5828
|
76 |
Theory of theory |
|
wenzelm@5828
|
77 |
Proof of ProofHistory.T;
|
wenzelm@5828
|
78 |
|
wenzelm@6689
|
79 |
fun str_of_node (Theory _) = "in theory mode"
|
wenzelm@6689
|
80 |
| str_of_node (Proof _) = "in proof mode";
|
wenzelm@6689
|
81 |
|
wenzelm@7308
|
82 |
fun print_ctxt thy = Pretty.writeln (Pretty.block
|
wenzelm@7308
|
83 |
[Pretty.str "Theory:", Pretty.brk 1, Pretty.str (PureThy.get_name thy),
|
wenzelm@7308
|
84 |
Pretty.str " =", Pretty.brk 1, Display.pretty_theory thy]);
|
wenzelm@7308
|
85 |
|
wenzelm@7308
|
86 |
fun print_node_ctxt (Theory thy) = print_ctxt thy
|
wenzelm@7308
|
87 |
| print_node_ctxt (Proof prf) = print_ctxt (Proof.theory_of (ProofHistory.current prf));
|
wenzelm@7308
|
88 |
|
wenzelm@7308
|
89 |
fun print_node (Theory thy) = print_ctxt thy
|
wenzelm@7308
|
90 |
| print_node (Proof prf) = Proof.print_state (ProofHistory.position prf)
|
wenzelm@7308
|
91 |
(ProofHistory.current prf);
|
wenzelm@5828
|
92 |
|
wenzelm@5828
|
93 |
|
wenzelm@5828
|
94 |
(* datatype state *)
|
wenzelm@5828
|
95 |
|
wenzelm@8465
|
96 |
datatype state = State of (node History.T * ((node -> unit) * (node -> unit))) option;
|
wenzelm@5828
|
97 |
|
wenzelm@8465
|
98 |
val toplevel = State None;
|
wenzelm@5828
|
99 |
|
wenzelm@8465
|
100 |
fun is_toplevel (State None) = true
|
wenzelm@7732
|
101 |
| is_toplevel _ = false;
|
wenzelm@7732
|
102 |
|
wenzelm@8465
|
103 |
fun str_of_state (State None) = "at top level"
|
wenzelm@8465
|
104 |
| str_of_state (State (Some (node, _))) = str_of_node (History.current node);
|
wenzelm@6689
|
105 |
|
wenzelm@6689
|
106 |
|
wenzelm@6689
|
107 |
(* prompt_state hook *)
|
wenzelm@6689
|
108 |
|
wenzelm@8465
|
109 |
fun prompt_state_default (State _) = Source.default_prompt;
|
wenzelm@6689
|
110 |
|
wenzelm@6689
|
111 |
val prompt_state_fn = ref prompt_state_default;
|
wenzelm@6689
|
112 |
fun prompt_state state = ! prompt_state_fn state;
|
wenzelm@5828
|
113 |
|
wenzelm@5828
|
114 |
|
wenzelm@7308
|
115 |
(* print state *)
|
wenzelm@5946
|
116 |
|
wenzelm@8465
|
117 |
fun print_current_node _ (State None) = ()
|
wenzelm@8465
|
118 |
| print_current_node prt (State (Some (node, _))) = prt (History.current node);
|
wenzelm@7308
|
119 |
|
wenzelm@8465
|
120 |
val print_state_context = print_current_node print_node_ctxt;
|
wenzelm@5946
|
121 |
|
wenzelm@5946
|
122 |
fun print_state_default state =
|
wenzelm@5946
|
123 |
let val ref (begin_state, end_state, _) = Goals.current_goals_markers in
|
wenzelm@5946
|
124 |
if begin_state = "" then () else writeln begin_state;
|
wenzelm@8465
|
125 |
print_current_node print_node state;
|
wenzelm@5946
|
126 |
if end_state = "" then () else writeln end_state
|
wenzelm@5946
|
127 |
end;
|
wenzelm@5946
|
128 |
|
wenzelm@5946
|
129 |
val print_state_fn = ref print_state_default;
|
wenzelm@5946
|
130 |
fun print_state state = ! print_state_fn state;
|
wenzelm@5946
|
131 |
|
wenzelm@5946
|
132 |
|
wenzelm@5828
|
133 |
(* top node *)
|
wenzelm@5828
|
134 |
|
wenzelm@6689
|
135 |
exception UNDEF;
|
wenzelm@6689
|
136 |
|
wenzelm@8465
|
137 |
fun node_history_of (State None) = raise UNDEF
|
wenzelm@8465
|
138 |
| node_history_of (State (Some (node, _))) = node;
|
wenzelm@6689
|
139 |
|
wenzelm@6689
|
140 |
val node_of = History.current o node_history_of;
|
wenzelm@5828
|
141 |
|
wenzelm@6664
|
142 |
fun node_case f g state =
|
wenzelm@5828
|
143 |
(case node_of state of
|
wenzelm@5828
|
144 |
Theory thy => f thy
|
wenzelm@5828
|
145 |
| Proof prf => g (ProofHistory.current prf));
|
wenzelm@5828
|
146 |
|
wenzelm@6664
|
147 |
val theory_of = node_case I Proof.theory_of;
|
wenzelm@7501
|
148 |
val context_of = node_case ProofContext.init Proof.context_of;
|
wenzelm@5828
|
149 |
val sign_of = Theory.sign_of o theory_of;
|
wenzelm@6664
|
150 |
val proof_of = node_case (fn _ => raise UNDEF) I;
|
wenzelm@6664
|
151 |
|
wenzelm@5828
|
152 |
|
wenzelm@5828
|
153 |
|
wenzelm@5828
|
154 |
(** toplevel transitions **)
|
wenzelm@5828
|
155 |
|
wenzelm@5828
|
156 |
exception TERMINATE;
|
wenzelm@5990
|
157 |
exception RESTART;
|
wenzelm@7022
|
158 |
exception EXCURSION_FAIL of exn * string;
|
wenzelm@6689
|
159 |
exception FAILURE of state * exn;
|
wenzelm@5828
|
160 |
|
wenzelm@5828
|
161 |
|
wenzelm@6689
|
162 |
(* recovery from stale signatures *)
|
wenzelm@6689
|
163 |
|
wenzelm@7022
|
164 |
(*note: proof commands should be non-destructive!*)
|
wenzelm@7022
|
165 |
|
wenzelm@6689
|
166 |
local
|
wenzelm@6689
|
167 |
|
wenzelm@6689
|
168 |
fun is_stale state = Sign.is_stale (sign_of state) handle UNDEF => false;
|
wenzelm@6689
|
169 |
|
wenzelm@6689
|
170 |
fun checkpoint_node true (Theory thy) = Theory (PureThy.checkpoint thy)
|
wenzelm@6689
|
171 |
| checkpoint_node _ node = node;
|
wenzelm@6689
|
172 |
|
wenzelm@6689
|
173 |
fun copy_node true (Theory thy) = Theory (Theory.copy thy)
|
wenzelm@6689
|
174 |
| copy_node _ node = node;
|
wenzelm@6689
|
175 |
|
wenzelm@6965
|
176 |
fun interruptible f x =
|
wenzelm@6965
|
177 |
let val y = ref (None: node History.T option);
|
wenzelm@6965
|
178 |
in exhibit_interrupt (fn () => y := Some (f x)) (); the (! y) end;
|
wenzelm@6965
|
179 |
|
wenzelm@7022
|
180 |
val rollback = ERROR_MESSAGE "Stale signature encountered after succesful execution!";
|
wenzelm@7022
|
181 |
|
wenzelm@7022
|
182 |
fun return (result, None) = result
|
wenzelm@7022
|
183 |
| return (result, Some exn) = raise FAILURE (result, exn);
|
wenzelm@7022
|
184 |
|
wenzelm@6689
|
185 |
in
|
wenzelm@6689
|
186 |
|
wenzelm@8465
|
187 |
fun node_trans _ _ _ (State None) = raise UNDEF
|
wenzelm@8465
|
188 |
| node_trans int hist f (State (Some (node, term))) =
|
wenzelm@6689
|
189 |
let
|
wenzelm@8465
|
190 |
fun mk_state nd = State (Some (nd, term));
|
wenzelm@6689
|
191 |
|
wenzelm@6689
|
192 |
val cont_node = History.map (checkpoint_node int) node;
|
wenzelm@6689
|
193 |
val back_node = History.map (copy_node int) cont_node;
|
wenzelm@6689
|
194 |
|
wenzelm@6689
|
195 |
val trans = if hist then History.apply_copy (copy_node int) else History.map;
|
wenzelm@6965
|
196 |
val (result, opt_exn) =
|
wenzelm@6965
|
197 |
(mk_state (transform_error (interruptible (trans (f int))) cont_node), None)
|
wenzelm@6965
|
198 |
handle exn => (mk_state cont_node, Some exn);
|
wenzelm@6689
|
199 |
in
|
wenzelm@7022
|
200 |
if is_stale result then return (mk_state back_node, Some (if_none opt_exn rollback))
|
wenzelm@7022
|
201 |
else return (result, opt_exn)
|
wenzelm@6689
|
202 |
end;
|
wenzelm@6689
|
203 |
|
wenzelm@6689
|
204 |
fun check_stale state =
|
wenzelm@6689
|
205 |
if not (is_stale state) then ()
|
wenzelm@7022
|
206 |
else warning "Stale signature encountered! Should restart current theory.";
|
wenzelm@6689
|
207 |
|
wenzelm@6689
|
208 |
end;
|
wenzelm@6689
|
209 |
|
wenzelm@6689
|
210 |
|
wenzelm@6689
|
211 |
(* primitive transitions *)
|
wenzelm@6689
|
212 |
|
wenzelm@6689
|
213 |
(*Important note: recovery from stale signatures is provided only for
|
wenzelm@6689
|
214 |
theory-level operations via MapCurrent and AppCurrent. Other node
|
wenzelm@6965
|
215 |
or state operations should not touch signatures at all.
|
wenzelm@6965
|
216 |
|
wenzelm@6965
|
217 |
Also note that interrupts are enabled for Keep, MapCurrent, and
|
wenzelm@6965
|
218 |
AppCurrent only.*)
|
wenzelm@5828
|
219 |
|
wenzelm@5828
|
220 |
datatype trans =
|
wenzelm@6689
|
221 |
Reset | (*empty toplevel*)
|
wenzelm@8465
|
222 |
Init of (bool -> node) * ((node -> unit) * (node -> unit)) |
|
wenzelm@8465
|
223 |
(*init node; provide exit/kill operation*)
|
wenzelm@8465
|
224 |
Exit | (*conclude node*)
|
wenzelm@6689
|
225 |
Kill | (*abort node*)
|
wenzelm@7612
|
226 |
Keep of bool -> state -> unit | (*peek at state*)
|
wenzelm@6689
|
227 |
History of node History.T -> node History.T | (*history operation (undo etc.)*)
|
wenzelm@6689
|
228 |
MapCurrent of bool -> node -> node | (*change node, bypassing history*)
|
wenzelm@6689
|
229 |
AppCurrent of bool -> node -> node; (*change node, recording history*)
|
wenzelm@6689
|
230 |
|
wenzelm@6689
|
231 |
fun undo_limit int = if int then None else Some 0;
|
wenzelm@6689
|
232 |
|
wenzelm@6689
|
233 |
local
|
wenzelm@5828
|
234 |
|
wenzelm@6689
|
235 |
fun apply_tr _ Reset _ = toplevel
|
wenzelm@8465
|
236 |
| apply_tr int (Init (f, term)) (State None) =
|
wenzelm@8465
|
237 |
State (Some (History.init (undo_limit int) (f int), term))
|
wenzelm@8465
|
238 |
| apply_tr _ (Init _ ) (State (Some _)) = raise UNDEF
|
wenzelm@8465
|
239 |
| apply_tr _ Exit (State None) = raise UNDEF
|
wenzelm@8465
|
240 |
| apply_tr _ Exit (State (Some (node, (exit, _)))) =
|
wenzelm@8465
|
241 |
(exit (History.current node); State None)
|
wenzelm@8465
|
242 |
| apply_tr _ Kill (State None) = raise UNDEF
|
wenzelm@8465
|
243 |
| apply_tr _ Kill (State (Some (node, (_, kill)))) =
|
wenzelm@8465
|
244 |
(kill (History.current node); State None)
|
wenzelm@7612
|
245 |
| apply_tr int (Keep f) state = (exhibit_interrupt (f int) state; state)
|
wenzelm@8465
|
246 |
| apply_tr _ (History _) (State None) = raise UNDEF
|
wenzelm@8465
|
247 |
| apply_tr _ (History f) (State (Some (node, term))) = State (Some (f node, term))
|
wenzelm@6689
|
248 |
| apply_tr int (MapCurrent f) state = node_trans int false f state
|
wenzelm@6689
|
249 |
| apply_tr int (AppCurrent f) state = node_trans int true f state;
|
wenzelm@5828
|
250 |
|
wenzelm@6689
|
251 |
fun apply_union _ [] state = raise FAILURE (state, UNDEF)
|
wenzelm@6689
|
252 |
| apply_union int (tr :: trs) state =
|
wenzelm@6689
|
253 |
transform_error (apply_tr int tr) state
|
wenzelm@6689
|
254 |
handle UNDEF => apply_union int trs state
|
wenzelm@6689
|
255 |
| FAILURE (alt_state, UNDEF) => apply_union int trs alt_state
|
wenzelm@6689
|
256 |
| exn as FAILURE _ => raise exn
|
wenzelm@6689
|
257 |
| exn => raise FAILURE (state, exn);
|
wenzelm@6689
|
258 |
|
wenzelm@6689
|
259 |
in
|
wenzelm@6689
|
260 |
|
wenzelm@6689
|
261 |
fun apply_trans int trs state = (apply_union int trs state, None)
|
wenzelm@6689
|
262 |
handle FAILURE (alt_state, exn) => (alt_state, Some exn) | exn => (state, Some exn);
|
wenzelm@6689
|
263 |
|
wenzelm@6689
|
264 |
end;
|
wenzelm@5828
|
265 |
|
wenzelm@5828
|
266 |
|
wenzelm@5828
|
267 |
(* datatype transition *)
|
wenzelm@5828
|
268 |
|
wenzelm@5828
|
269 |
datatype transition = Transition of
|
wenzelm@5828
|
270 |
{name: string,
|
wenzelm@5828
|
271 |
pos: Position.T,
|
wenzelm@5828
|
272 |
int_only: bool,
|
wenzelm@5828
|
273 |
print: bool,
|
wenzelm@5828
|
274 |
trans: trans list};
|
wenzelm@5828
|
275 |
|
wenzelm@5828
|
276 |
fun make_transition (name, pos, int_only, print, trans) =
|
wenzelm@5828
|
277 |
Transition {name = name, pos = pos, int_only = int_only, print = print, trans = trans};
|
wenzelm@5828
|
278 |
|
wenzelm@5828
|
279 |
fun map_transition f (Transition {name, pos, int_only, print, trans}) =
|
wenzelm@5828
|
280 |
make_transition (f (name, pos, int_only, print, trans));
|
wenzelm@5828
|
281 |
|
wenzelm@5828
|
282 |
val empty = make_transition ("<unknown>", Position.none, false, false, []);
|
wenzelm@5828
|
283 |
|
wenzelm@5828
|
284 |
|
wenzelm@5828
|
285 |
(* diagnostics *)
|
wenzelm@5828
|
286 |
|
wenzelm@5828
|
287 |
fun str_of_transition (Transition {name, pos, ...}) = quote name ^ Position.str_of pos;
|
wenzelm@5828
|
288 |
|
wenzelm@5828
|
289 |
fun command_msg msg tr = msg ^ "command " ^ str_of_transition tr;
|
wenzelm@5828
|
290 |
fun at_command tr = command_msg "At " tr ^ ".";
|
wenzelm@5828
|
291 |
|
wenzelm@5828
|
292 |
fun type_error tr state =
|
wenzelm@6689
|
293 |
ERROR_MESSAGE (command_msg "Illegal application of " tr ^ " " ^ str_of_state state);
|
wenzelm@5828
|
294 |
|
wenzelm@5828
|
295 |
|
wenzelm@5828
|
296 |
(* modify transitions *)
|
wenzelm@5828
|
297 |
|
wenzelm@5828
|
298 |
fun name nm = map_transition
|
wenzelm@5828
|
299 |
(fn (_, pos, int_only, print, trans) => (nm, pos, int_only, print, trans));
|
wenzelm@5828
|
300 |
|
wenzelm@5828
|
301 |
fun position pos = map_transition
|
wenzelm@5828
|
302 |
(fn (name, _, int_only, print, trans) => (name, pos, int_only, print, trans));
|
wenzelm@5828
|
303 |
|
wenzelm@5828
|
304 |
fun interactive int_only = map_transition
|
wenzelm@5828
|
305 |
(fn (name, pos, _, print, trans) => (name, pos, int_only, print, trans));
|
wenzelm@5828
|
306 |
|
wenzelm@5828
|
307 |
val print = map_transition
|
wenzelm@5828
|
308 |
(fn (name, pos, int_only, _, trans) => (name, pos, int_only, true, trans));
|
wenzelm@5828
|
309 |
|
wenzelm@5828
|
310 |
fun add_trans tr = map_transition
|
wenzelm@5828
|
311 |
(fn (name, pos, int_only, print, trans) => (name, pos, int_only, print, trans @ [tr]));
|
wenzelm@5828
|
312 |
|
wenzelm@5828
|
313 |
|
wenzelm@5828
|
314 |
(* build transitions *)
|
wenzelm@5828
|
315 |
|
wenzelm@5828
|
316 |
val reset = add_trans Reset;
|
wenzelm@6689
|
317 |
fun init f exit kill = add_trans (Init (f, (exit, kill)));
|
wenzelm@6689
|
318 |
val exit = add_trans Exit;
|
wenzelm@6689
|
319 |
val kill = add_trans Kill;
|
wenzelm@7612
|
320 |
val keep' = add_trans o Keep;
|
wenzelm@6689
|
321 |
val history = add_trans o History;
|
wenzelm@5828
|
322 |
val map_current = add_trans o MapCurrent;
|
wenzelm@6689
|
323 |
val app_current = add_trans o AppCurrent;
|
wenzelm@5828
|
324 |
|
wenzelm@7612
|
325 |
fun keep f = add_trans (Keep (fn _ => f));
|
wenzelm@5828
|
326 |
fun imperative f = keep (fn _ => f ());
|
wenzelm@5828
|
327 |
|
wenzelm@6689
|
328 |
fun init_theory f exit kill =
|
wenzelm@8465
|
329 |
init (fn int => Theory (f int))
|
wenzelm@6689
|
330 |
(fn Theory thy => exit thy | _ => raise UNDEF)
|
wenzelm@6689
|
331 |
(fn Theory thy => kill thy | _ => raise UNDEF);
|
wenzelm@5828
|
332 |
|
wenzelm@6689
|
333 |
fun theory f = app_current (fn _ => (fn Theory thy => Theory (f thy) | _ => raise UNDEF));
|
wenzelm@7612
|
334 |
fun theory' f = app_current (fn int => (fn Theory thy => Theory (f int thy) | _ => raise UNDEF));
|
wenzelm@6689
|
335 |
fun theory_to_proof f =
|
wenzelm@6689
|
336 |
app_current (fn int => (fn Theory thy => Proof (f int thy) | _ => raise UNDEF));
|
wenzelm@6689
|
337 |
fun proof' f = map_current (fn int => (fn Proof prf => Proof (f int prf) | _ => raise UNDEF));
|
wenzelm@6689
|
338 |
val proof = proof' o K;
|
wenzelm@6689
|
339 |
fun proof_to_theory f = map_current (fn _ => (fn Proof prf => Theory (f prf) | _ => raise UNDEF));
|
wenzelm@5828
|
340 |
|
wenzelm@5828
|
341 |
|
wenzelm@5828
|
342 |
|
wenzelm@5828
|
343 |
(** toplevel transactions **)
|
wenzelm@5828
|
344 |
|
wenzelm@7198
|
345 |
val quiet = ref false;
|
wenzelm@5828
|
346 |
val trace = ref false;
|
wenzelm@5828
|
347 |
|
wenzelm@5828
|
348 |
|
wenzelm@5828
|
349 |
(* print exceptions *)
|
wenzelm@5828
|
350 |
|
wenzelm@5828
|
351 |
fun raised name = "exception " ^ name ^ " raised";
|
wenzelm@5828
|
352 |
fun raised_msg name msg = raised name ^ ": " ^ msg;
|
wenzelm@5828
|
353 |
|
wenzelm@5828
|
354 |
fun exn_message TERMINATE = "Exit."
|
wenzelm@5990
|
355 |
| exn_message RESTART = "Restart."
|
wenzelm@7022
|
356 |
| exn_message (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_message exn, msg]
|
wenzelm@6971
|
357 |
| exn_message Interrupt = "Interrupt."
|
wenzelm@6002
|
358 |
| exn_message ERROR = "ERROR."
|
wenzelm@5828
|
359 |
| exn_message (ERROR_MESSAGE msg) = msg
|
wenzelm@5828
|
360 |
| exn_message (THEORY (msg, _)) = msg
|
wenzelm@5828
|
361 |
| exn_message (ProofContext.CONTEXT (msg, _)) = msg
|
wenzelm@5828
|
362 |
| exn_message (Proof.STATE (msg, _)) = msg
|
wenzelm@5828
|
363 |
| exn_message (ProofHistory.FAIL msg) = msg
|
wenzelm@5920
|
364 |
| exn_message (Attrib.ATTRIB_FAIL info) = fail_message "attribute" info
|
wenzelm@5920
|
365 |
| exn_message (Method.METHOD_FAIL info) = fail_message "method" info
|
wenzelm@5828
|
366 |
| exn_message (Syntax.AST (msg, _)) = raised_msg "AST" msg
|
wenzelm@5828
|
367 |
| exn_message (TYPE (msg, _, _)) = raised_msg "TYPE" msg
|
wenzelm@5828
|
368 |
| exn_message (TERM (msg, _)) = raised_msg "TERM" msg
|
wenzelm@5828
|
369 |
| exn_message (THM (msg, _, _)) = raised_msg "THM" msg
|
wenzelm@5828
|
370 |
| exn_message Library.OPTION = raised "Library.OPTION"
|
wenzelm@5828
|
371 |
| exn_message (Library.LIST msg) = raised_msg "Library.LIST" msg
|
wenzelm@5920
|
372 |
| exn_message exn = General.exnMessage exn
|
wenzelm@5920
|
373 |
and fail_message kind ((name, pos), exn) =
|
wenzelm@5920
|
374 |
"Error in " ^ kind ^ " " ^ name ^ Position.str_of pos ^ ":\n" ^ exn_message exn;
|
wenzelm@5828
|
375 |
|
wenzelm@5828
|
376 |
fun print_exn None = ()
|
wenzelm@5828
|
377 |
| print_exn (Some (exn, s)) = error_msg (cat_lines [exn_message exn, s]);
|
wenzelm@5828
|
378 |
|
wenzelm@5828
|
379 |
|
wenzelm@5828
|
380 |
(* apply transitions *)
|
wenzelm@5828
|
381 |
|
wenzelm@6664
|
382 |
local
|
wenzelm@6664
|
383 |
|
wenzelm@6689
|
384 |
fun app int (tr as Transition {trans, int_only, print, ...}) state =
|
wenzelm@5828
|
385 |
let
|
wenzelm@5828
|
386 |
val _ =
|
wenzelm@5828
|
387 |
if int orelse not int_only then ()
|
wenzelm@5828
|
388 |
else warning (command_msg "Executing interactive-only " tr);
|
wenzelm@6689
|
389 |
val (result, opt_exn) =
|
wenzelm@6689
|
390 |
(if ! trace then (writeln (command_msg "" tr); timeap) else I) (apply_trans int trans) state;
|
wenzelm@7198
|
391 |
val _ = if int andalso print andalso not (! quiet) then print_state result else ();
|
wenzelm@6689
|
392 |
in (result, apsome (fn UNDEF => type_error tr state | exn => exn) opt_exn) end;
|
wenzelm@6664
|
393 |
|
wenzelm@6664
|
394 |
in
|
wenzelm@5828
|
395 |
|
wenzelm@6689
|
396 |
fun apply int tr st =
|
wenzelm@6965
|
397 |
(case app int tr st of
|
wenzelm@6689
|
398 |
(_, Some TERMINATE) => None
|
wenzelm@6664
|
399 |
| (_, Some RESTART) => Some (toplevel, None)
|
wenzelm@7022
|
400 |
| (state', Some (EXCURSION_FAIL exn_info)) => Some (state', Some exn_info)
|
wenzelm@6689
|
401 |
| (state', Some exn) => Some (state', Some (exn, at_command tr))
|
wenzelm@6689
|
402 |
| (state', None) => Some (state', None));
|
wenzelm@6664
|
403 |
|
wenzelm@6664
|
404 |
end;
|
wenzelm@5828
|
405 |
|
wenzelm@5828
|
406 |
|
wenzelm@5828
|
407 |
(* excursion: toplevel -- apply transformers -- toplevel *)
|
wenzelm@5828
|
408 |
|
wenzelm@6664
|
409 |
local
|
wenzelm@6664
|
410 |
|
wenzelm@5828
|
411 |
fun excur [] x = x
|
wenzelm@5828
|
412 |
| excur (tr :: trs) x =
|
wenzelm@5828
|
413 |
(case apply false tr x of
|
wenzelm@5828
|
414 |
Some (x', None) => excur trs x'
|
wenzelm@7022
|
415 |
| Some (x', Some exn_info) => raise EXCURSION_FAIL exn_info
|
wenzelm@7022
|
416 |
| None => raise EXCURSION_FAIL (TERMINATE, at_command tr));
|
wenzelm@5828
|
417 |
|
wenzelm@6664
|
418 |
in
|
wenzelm@6664
|
419 |
|
wenzelm@5828
|
420 |
fun excursion trs =
|
wenzelm@8465
|
421 |
(case excur trs (State None) of
|
wenzelm@8465
|
422 |
State None => ()
|
wenzelm@8465
|
423 |
| _ => raise ERROR_MESSAGE "Unfinished development at end of input");
|
wenzelm@5828
|
424 |
|
wenzelm@7062
|
425 |
fun excursion_error trs =
|
wenzelm@7062
|
426 |
excursion trs handle exn => error (exn_message exn);
|
wenzelm@7062
|
427 |
|
wenzelm@6664
|
428 |
end;
|
wenzelm@6664
|
429 |
|
wenzelm@5828
|
430 |
|
wenzelm@5828
|
431 |
|
wenzelm@5828
|
432 |
(** interactive transformations **)
|
wenzelm@5828
|
433 |
|
wenzelm@5828
|
434 |
(* the global state reference *)
|
wenzelm@5828
|
435 |
|
wenzelm@6689
|
436 |
val global_state = ref (toplevel, None: (exn * string) option);
|
wenzelm@5828
|
437 |
|
wenzelm@5828
|
438 |
fun set_state state = global_state := (state, None);
|
wenzelm@5828
|
439 |
fun get_state () = fst (! global_state);
|
wenzelm@5828
|
440 |
fun exn () = snd (! global_state);
|
wenzelm@5828
|
441 |
|
wenzelm@5828
|
442 |
|
wenzelm@6965
|
443 |
(* the Isar source of transitions *)
|
wenzelm@6965
|
444 |
|
wenzelm@6965
|
445 |
type isar =
|
wenzelm@6965
|
446 |
(transition, (transition option,
|
wenzelm@6965
|
447 |
(OuterLex.token, (OuterLex.token,
|
wenzelm@6965
|
448 |
Position.T * (Symbol.symbol, (string, unit) Source.source) Source.source)
|
wenzelm@6965
|
449 |
Source.source) Source.source) Source.source) Source.source;
|
wenzelm@6965
|
450 |
|
wenzelm@6965
|
451 |
|
wenzelm@5828
|
452 |
(* apply transformers to global state *)
|
wenzelm@5828
|
453 |
|
wenzelm@5828
|
454 |
nonfix >>;
|
wenzelm@5828
|
455 |
|
wenzelm@5828
|
456 |
fun >> tr =
|
wenzelm@5828
|
457 |
(case apply true tr (get_state ()) of
|
wenzelm@5828
|
458 |
None => false
|
wenzelm@5828
|
459 |
| Some (state', exn_info) =>
|
wenzelm@5828
|
460 |
(global_state := (state', exn_info);
|
wenzelm@5828
|
461 |
check_stale state'; print_exn exn_info;
|
wenzelm@5828
|
462 |
true));
|
wenzelm@5828
|
463 |
|
wenzelm@7602
|
464 |
(*Note: this is for Poly/ML only, we really do not intend to exhibit
|
wenzelm@7602
|
465 |
interrupts here*)
|
wenzelm@7602
|
466 |
fun get_interrupt src = Some (Source.get_single src) handle Interrupt => None;
|
wenzelm@7602
|
467 |
|
wenzelm@5828
|
468 |
fun raw_loop src =
|
wenzelm@7602
|
469 |
(case get_interrupt (Source.set_prompt (prompt_state (get_state ())) src) of
|
wenzelm@6971
|
470 |
None => (writeln "\nInterrupt."; raw_loop src)
|
wenzelm@5828
|
471 |
| Some None => ()
|
wenzelm@5828
|
472 |
| Some (Some (tr, src')) => if >> tr then raw_loop src' else ());
|
wenzelm@5828
|
473 |
|
wenzelm@5828
|
474 |
|
wenzelm@5828
|
475 |
fun loop src = mask_interrupt raw_loop src;
|
wenzelm@5828
|
476 |
|
wenzelm@5828
|
477 |
|
wenzelm@5828
|
478 |
end;
|