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