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