| author | wenzelm |
| Tue, 15 Jul 2008 15:59:49 +0200 | |
| changeset 27608 | 8fd5662ccd97 |
| parent 27606 | 82399f3a8ca8 |
| child 27840 | e9483ef084cc |
| permissions | -rw-r--r-- |
| 5828 | 1 |
(* Title: Pure/Isar/toplevel.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
5 |
Isabelle/Isar toplevel transactions. |
| 5828 | 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 |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
14 |
val proof_node: node -> ProofNode.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 |
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
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 |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
23 |
val previous_node_of: state -> node option |
| 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 |
| 22089 | 27 |
val generic_theory_of: state -> generic_theory |
| 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 print_state_context: state -> unit |
33 |
val print_state: bool -> state -> unit |
|
| 16682 | 34 |
val quiet: bool ref |
35 |
val debug: bool ref |
|
|
17321
227f1f5c3959
added interact flag to control mode of excursions;
wenzelm
parents:
17320
diff
changeset
|
36 |
val interact: bool ref |
| 16682 | 37 |
val timing: bool ref |
38 |
val profiling: int ref |
|
| 16815 | 39 |
val skip_proofs: bool ref |
| 5828 | 40 |
exception TERMINATE |
| 27583 | 41 |
exception TOPLEVEL_ERROR |
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
42 |
exception CONTEXT of Proof.context * exn |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
43 |
val exn_message: exn -> string |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
44 |
val program: (unit -> 'a) -> 'a |
| 16682 | 45 |
type transition |
| 5828 | 46 |
val empty: transition |
| 27441 | 47 |
val init_of: transition -> string option |
| 27427 | 48 |
val name_of: transition -> string |
| 27500 | 49 |
val str_of: transition -> string |
| 5828 | 50 |
val name: string -> transition -> transition |
51 |
val position: Position.T -> transition -> transition |
|
52 |
val interactive: bool -> transition -> transition |
|
53 |
val print: transition -> transition |
|
| 9010 | 54 |
val no_timing: transition -> transition |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
55 |
val init_theory: string -> (bool -> theory) -> (theory -> unit) -> transition -> transition |
| 6689 | 56 |
val exit: transition -> transition |
| 5828 | 57 |
val keep: (state -> unit) -> transition -> transition |
| 7612 | 58 |
val keep': (bool -> state -> unit) -> transition -> transition |
| 5828 | 59 |
val imperative: (unit -> unit) -> transition -> transition |
60 |
val theory: (theory -> theory) -> transition -> transition |
|
| 26491 | 61 |
val generic_theory: (generic_theory -> generic_theory) -> transition -> transition |
| 7612 | 62 |
val theory': (bool -> theory -> theory) -> transition -> transition |
| 20985 | 63 |
val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition |
| 21007 | 64 |
val end_local_theory: transition -> transition |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
65 |
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
|
66 |
val present_local_theory: xstring option -> (bool -> node -> unit) -> transition -> transition |
| 24453 | 67 |
val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) -> |
68 |
transition -> transition |
|
| 21007 | 69 |
val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) -> |
70 |
transition -> transition |
|
| 17363 | 71 |
val theory_to_proof: (theory -> Proof.state) -> transition -> transition |
| 21007 | 72 |
val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition |
73 |
val forget_proof: transition -> transition |
|
| 21177 | 74 |
val present_proof: (bool -> node -> unit) -> transition -> transition |
75 |
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
|
76 |
val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition |
| 21177 | 77 |
val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition |
78 |
val proof: (Proof.state -> Proof.state) -> transition -> transition |
|
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
79 |
val actual_proof: (ProofNode.T -> ProofNode.T) -> transition -> transition |
|
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
80 |
val skip_proof: (int -> int) -> transition -> transition |
|
17904
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents:
17513
diff
changeset
|
81 |
val skip_proof_to_theory: (int -> bool) -> transition -> transition |
| 27427 | 82 |
val get_id: transition -> string option |
83 |
val put_id: string -> transition -> transition |
|
| 9512 | 84 |
val unknown_theory: transition -> transition |
85 |
val unknown_proof: transition -> transition |
|
86 |
val unknown_context: transition -> transition |
|
| 27606 | 87 |
val status: transition -> Markup.T -> unit |
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
88 |
val error_msg: transition -> exn * string -> unit |
|
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
89 |
val transition: bool -> transition -> state -> (state * (exn * string) option) option |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
90 |
val commit_exit: transition |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
91 |
val excursion: transition list -> unit |
| 17076 | 92 |
val present_excursion: (transition * (state -> state -> 'a -> 'a)) list -> 'a -> 'a |
| 5828 | 93 |
end; |
94 |
||
| 6965 | 95 |
structure Toplevel: TOPLEVEL = |
| 5828 | 96 |
struct |
97 |
||
98 |
(** toplevel state **) |
|
99 |
||
| 19063 | 100 |
exception UNDEF; |
101 |
||
102 |
||
| 21294 | 103 |
(* local theory wrappers *) |
| 5828 | 104 |
|
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
105 |
type generic_theory = Context.generic; (*theory or local_theory*) |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
106 |
|
| 25292 | 107 |
val loc_init = TheoryTarget.context; |
| 21294 | 108 |
val loc_exit = ProofContext.theory_of o LocalTheory.exit; |
109 |
||
| 25292 | 110 |
fun loc_begin loc (Context.Theory thy) = loc_init (the_default "-" loc) thy |
| 21294 | 111 |
| loc_begin NONE (Context.Proof lthy) = lthy |
| 25269 | 112 |
| loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy); |
| 21294 | 113 |
|
114 |
fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit |
|
115 |
| loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore |
|
| 25292 | 116 |
| loc_finish (SOME _) (Context.Proof lthy) = fn lthy' => |
117 |
Context.Proof (LocalTheory.reinit (LocalTheory.raw_theory (K (loc_exit lthy')) lthy)); |
|
| 21294 | 118 |
|
119 |
||
|
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
120 |
(* datatype node *) |
| 21294 | 121 |
|
| 5828 | 122 |
datatype node = |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
123 |
Theory of generic_theory * Proof.context option |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
124 |
(*theory with presentation context*) | |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
125 |
Proof of ProofNode.T * ((Proof.context -> generic_theory) * generic_theory) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
126 |
(*proof node, finish, original theory*) | |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
127 |
SkipProof of int * (generic_theory * generic_theory); |
|
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
128 |
(*proof depth, resulting theory, original theory*) |
| 5828 | 129 |
|
| 22056 | 130 |
val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF; |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
131 |
val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE; |
| 18589 | 132 |
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; |
133 |
||
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
134 |
fun cases_node f _ (Theory (gthy, _)) = f gthy |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
135 |
| cases_node _ g (Proof (prf, _)) = g (ProofNode.current prf) |
| 21007 | 136 |
| cases_node f _ (SkipProof (_, (gthy, _))) = f gthy; |
| 19063 | 137 |
|
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
138 |
fun presentation_context (SOME (Theory (_, SOME ctxt))) NONE = ctxt |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
139 |
| 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
|
140 |
| presentation_context (SOME node) (SOME loc) = |
| 25269 | 141 |
loc_init 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
|
142 |
| presentation_context NONE _ = raise UNDEF; |
| 19063 | 143 |
|
|
26624
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
144 |
fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE) |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
145 |
| reset_presentation node = node; |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
146 |
|
|
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
147 |
|
|
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
148 |
(* datatype state *) |
|
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
149 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
150 |
type node_info = node * (theory -> unit); (*node with exit operation*) |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
151 |
datatype state = State of node_info option * node_info option; (*current, previous*) |
| 5828 | 152 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
153 |
val toplevel = State (NONE, NONE); |
| 5828 | 154 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
155 |
fun is_toplevel (State (NONE, _)) = true |
| 7732 | 156 |
| is_toplevel _ = false; |
157 |
||
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
158 |
fun level (State (NONE, _)) = 0 |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
159 |
| level (State (SOME (Theory _, _), _)) = 0 |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
160 |
| level (State (SOME (Proof (prf, _), _), _)) = Proof.level (ProofNode.current prf) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
161 |
| level (State (SOME (SkipProof (d, _), _), _)) = d + 1; (*different notion of proof depth!*) |
| 17076 | 162 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
163 |
fun str_of_state (State (NONE, _)) = "at top level" |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
164 |
| str_of_state (State (SOME (Theory (Context.Theory _, _), _), _)) = "in theory mode" |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
165 |
| str_of_state (State (SOME (Theory (Context.Proof _, _), _), _)) = "in local theory mode" |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
166 |
| str_of_state (State (SOME (Proof _, _), _)) = "in proof mode" |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
167 |
| str_of_state (State (SOME (SkipProof _, _), _)) = "in skipped proof mode"; |
|
5946
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
168 |
|
|
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
169 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
170 |
(* current node *) |
| 5828 | 171 |
|
| 27603 | 172 |
fun previous_node_of (State (_, prev)) = Option.map #1 prev; |
| 6689 | 173 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
174 |
fun node_of (State (NONE, _)) = raise UNDEF |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
175 |
| node_of (State (SOME (node, _), _)) = node; |
| 5828 | 176 |
|
| 18589 | 177 |
fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state)); |
178 |
fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state)); |
|
179 |
||
| 19063 | 180 |
fun node_case f g state = cases_node f g (node_of state); |
| 5828 | 181 |
|
| 21506 | 182 |
val context_of = node_case Context.proof_of Proof.context_of; |
| 22089 | 183 |
val generic_theory_of = node_case I (Context.Proof o Proof.context_of); |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
184 |
val theory_of = node_case Context.theory_of Proof.theory_of; |
| 18589 | 185 |
val proof_of = node_case (fn _ => raise UNDEF) I; |
| 17208 | 186 |
|
| 18589 | 187 |
fun proof_position_of state = |
188 |
(case node_of state of |
|
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
189 |
Proof (prf, _) => ProofNode.position prf |
| 18589 | 190 |
| _ => raise UNDEF); |
| 6664 | 191 |
|
| 21007 | 192 |
val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward; |
| 5828 | 193 |
|
194 |
||
| 16815 | 195 |
(* print state *) |
196 |
||
| 25292 | 197 |
val pretty_context = LocalTheory.pretty o Context.cases (TheoryTarget.init NONE) I; |
| 16815 | 198 |
|
|
23640
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
wenzelm
parents:
23619
diff
changeset
|
199 |
fun print_state_context state = |
|
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
200 |
(case try node_of state of |
| 21506 | 201 |
NONE => [] |
|
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
202 |
| SOME (Theory (gthy, _)) => pretty_context gthy |
|
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
203 |
| SOME (Proof (_, (_, gthy))) => pretty_context gthy |
|
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
204 |
| SOME (SkipProof (_, (gthy, _))) => pretty_context gthy) |
|
23640
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
wenzelm
parents:
23619
diff
changeset
|
205 |
|> Pretty.chunks |> Pretty.writeln; |
| 16815 | 206 |
|
|
23640
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
wenzelm
parents:
23619
diff
changeset
|
207 |
fun print_state prf_only state = |
| 23701 | 208 |
(case try node_of state of |
209 |
NONE => [] |
|
210 |
| SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy |
|
211 |
| SOME (Proof (prf, _)) => |
|
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
212 |
Proof.pretty_state (ProofNode.position prf) (ProofNode.current prf) |
|
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
213 |
| SOME (SkipProof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)])
|
| 23701 | 214 |
|> Pretty.markup_chunks Markup.state |> Pretty.writeln; |
| 16815 | 215 |
|
216 |
||
| 15668 | 217 |
|
| 5828 | 218 |
(** toplevel transitions **) |
219 |
||
| 16682 | 220 |
val quiet = ref false; |
| 22135 | 221 |
val debug = Output.debugging; |
|
17321
227f1f5c3959
added interact flag to control mode of excursions;
wenzelm
parents:
17320
diff
changeset
|
222 |
val interact = ref false; |
| 16682 | 223 |
val timing = Output.timing; |
224 |
val profiling = ref 0; |
|
| 16815 | 225 |
val skip_proofs = ref false; |
| 16682 | 226 |
|
| 5828 | 227 |
exception TERMINATE; |
| 7022 | 228 |
exception EXCURSION_FAIL of exn * string; |
| 6689 | 229 |
exception FAILURE of state * exn; |
|
24055
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
wenzelm
parents:
23975
diff
changeset
|
230 |
exception TOPLEVEL_ERROR; |
| 5828 | 231 |
|
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
232 |
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
233 |
(* print exceptions *) |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
234 |
|
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
235 |
exception CONTEXT of Proof.context * exn; |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
236 |
|
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
237 |
fun exn_context NONE exn = exn |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
238 |
| exn_context (SOME ctxt) exn = CONTEXT (ctxt, exn); |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
239 |
|
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
240 |
local |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
241 |
|
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
242 |
fun if_context NONE _ _ = [] |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
243 |
| if_context (SOME ctxt) f xs = map (f ctxt) xs; |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
244 |
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
245 |
fun raised name [] = "exception " ^ name ^ " raised" |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
246 |
| raised name [msg] = "exception " ^ name ^ " raised: " ^ msg |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
247 |
| raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
248 |
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
249 |
in |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
250 |
|
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
251 |
fun exn_message e = |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
252 |
let |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
253 |
val detailed = ! debug; |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
254 |
|
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
255 |
fun exn_msg _ (CONTEXT (ctxt, exn)) = exn_msg (SOME ctxt) exn |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
256 |
| exn_msg ctxt (Exn.EXCEPTIONS (exns, "")) = cat_lines (map (exn_msg ctxt) exns) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
257 |
| exn_msg ctxt (Exn.EXCEPTIONS (exns, msg)) = cat_lines (map (exn_msg ctxt) exns @ [msg]) |
| 26293 | 258 |
| exn_msg ctxt (EXCURSION_FAIL (exn, loc)) = |
259 |
exn_msg ctxt exn ^ Markup.markup Markup.location ("\n" ^ loc)
|
|
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
260 |
| exn_msg _ TERMINATE = "Exit." |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
261 |
| exn_msg _ Interrupt = "Interrupt." |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
262 |
| exn_msg _ TimeLimit.TimeOut = "Timeout." |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
263 |
| exn_msg _ TOPLEVEL_ERROR = "Error." |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
264 |
| exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
265 |
| exn_msg _ (ERROR msg) = msg |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
266 |
| exn_msg _ (Fail msg) = raised "Fail" [msg] |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
267 |
| exn_msg _ (THEORY (msg, thys)) = |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
268 |
raised "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else [])) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
269 |
| exn_msg _ (Syntax.AST (msg, asts)) = raised "AST" (msg :: |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
270 |
(if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else [])) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
271 |
| exn_msg ctxt (TYPE (msg, Ts, ts)) = raised "TYPE" (msg :: |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
272 |
(if detailed then |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
273 |
if_context ctxt Syntax.string_of_typ Ts @ if_context ctxt Syntax.string_of_term ts |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
274 |
else [])) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
275 |
| exn_msg ctxt (TERM (msg, ts)) = raised "TERM" (msg :: |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
276 |
(if detailed then if_context ctxt Syntax.string_of_term ts else [])) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
277 |
| exn_msg ctxt (THM (msg, i, thms)) = raised ("THM " ^ string_of_int i) (msg ::
|
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
278 |
(if detailed then if_context ctxt ProofContext.string_of_thm thms else [])) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
279 |
| exn_msg _ exn = raised (General.exnMessage exn) [] |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
280 |
in exn_msg NONE e end; |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
281 |
|
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
282 |
end; |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
283 |
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
284 |
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
285 |
(* controlled execution *) |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
286 |
|
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
287 |
local |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
288 |
|
| 18685 | 289 |
fun debugging f x = |
| 23940 | 290 |
if ! debug then exception_trace (fn () => f x) |
| 18685 | 291 |
else f x; |
292 |
||
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
293 |
fun toplevel_error f x = |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
294 |
let val ctxt = try ML_Context.the_local_context () in |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
295 |
f x handle exn => |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
296 |
(Output.error_msg (exn_message (exn_context ctxt exn)); raise TOPLEVEL_ERROR) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
297 |
end; |
|
24055
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
wenzelm
parents:
23975
diff
changeset
|
298 |
|
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
299 |
in |
|
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 |
fun controlled_execution f = |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
302 |
f |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
303 |
|> debugging |
|
24055
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
wenzelm
parents:
23975
diff
changeset
|
304 |
|> interruptible; |
|
20128
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 program f = |
|
24055
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
wenzelm
parents:
23975
diff
changeset
|
307 |
(f |
|
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
wenzelm
parents:
23975
diff
changeset
|
308 |
|> debugging |
|
f7483532537b
added TOPLEVEL_ERROR (simplified version from output.ML);
wenzelm
parents:
23975
diff
changeset
|
309 |
|> toplevel_error) (); |
|
20128
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 |
end; |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
312 |
|
| 5828 | 313 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
314 |
(* node transactions -- maintaining stable checkpoints *) |
| 7022 | 315 |
|
| 6689 | 316 |
local |
317 |
||
|
26624
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
318 |
fun is_draft_theory (Theory (gthy, _)) = Context.is_draft (Context.theory_of gthy) |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
319 |
| is_draft_theory _ = false; |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
320 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
321 |
fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false; |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
322 |
|
|
26624
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
323 |
fun stale_error NONE = SOME (ERROR "Stale theory encountered after successful execution!") |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
324 |
| stale_error some = some; |
| 16815 | 325 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
326 |
fun map_theory f (Theory (gthy, ctxt)) = |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
327 |
Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, ctxt) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
328 |
| map_theory _ node = node; |
| 6689 | 329 |
|
330 |
in |
|
331 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
332 |
fun apply_transaction pos f (node, exit) = |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
333 |
let |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
334 |
val _ = is_draft_theory node andalso error "Illegal draft theory in toplevel state"; |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
335 |
val cont_node = reset_presentation node; |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
336 |
val back_node = map_theory (Theory.checkpoint o Theory.copy) cont_node; |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
337 |
fun state_error e nd = (State (SOME (nd, exit), SOME (node, exit)), e); |
| 6689 | 338 |
|
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
339 |
val (result, err) = |
|
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
340 |
cont_node |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
341 |
|> controlled_execution f |
|
26624
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
342 |
|> map_theory Theory.checkpoint |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
343 |
|> state_error NONE |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
344 |
handle exn => state_error (SOME exn) cont_node; |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
345 |
|
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
346 |
val (result', err') = |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
347 |
if is_stale result then state_error (stale_error err) back_node |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
348 |
else (result, err); |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
349 |
in |
|
26624
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
350 |
(case err' of |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
351 |
NONE => result' |
|
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
352 |
| SOME exn => raise FAILURE (result', exn)) |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
353 |
end; |
| 6689 | 354 |
|
355 |
end; |
|
356 |
||
357 |
||
358 |
(* primitive transitions *) |
|
359 |
||
| 5828 | 360 |
datatype trans = |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
361 |
Init of string * (bool -> theory) * (theory -> unit) | (*theory name, init, exit*) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
362 |
Exit | (*formal exit of theory -- without committing*) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
363 |
CommitExit | (*exit and commit final theory*) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
364 |
Keep of bool -> state -> unit | (*peek at state*) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
365 |
Transaction of bool -> node -> node; (*node transaction*) |
|
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
366 |
|
| 6689 | 367 |
local |
| 5828 | 368 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
369 |
fun apply_tr int _ (Init (_, f, exit)) (State (NONE, _)) = |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
370 |
State (SOME (Theory (Context.Theory (Theory.checkpoint (f int)), NONE), exit), NONE) |
| 27603 | 371 |
| apply_tr _ _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) = |
372 |
State (NONE, prev) |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
373 |
| apply_tr _ _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) = |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
374 |
(controlled_execution exit thy; toplevel) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
375 |
| apply_tr int _ (Keep f) state = |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
376 |
controlled_execution (fn x => tap (f int) x) state |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
377 |
| apply_tr int pos (Transaction f) (State (SOME state, _)) = |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
378 |
apply_transaction pos (fn x => f int x) state |
|
23363
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
379 |
| apply_tr _ _ _ _ = raise UNDEF; |
| 5828 | 380 |
|
|
23363
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
381 |
fun apply_union _ _ [] state = raise FAILURE (state, UNDEF) |
|
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
382 |
| apply_union int pos (tr :: trs) state = |
|
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
383 |
apply_tr int pos tr state |
|
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
384 |
handle UNDEF => apply_union int pos trs state |
|
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
385 |
| FAILURE (alt_state, UNDEF) => apply_union int pos trs alt_state |
| 6689 | 386 |
| exn as FAILURE _ => raise exn |
387 |
| exn => raise FAILURE (state, exn); |
|
388 |
||
389 |
in |
|
390 |
||
|
23363
9981199bf865
transaction: context_position (non-destructive only);
wenzelm
parents:
22588
diff
changeset
|
391 |
fun apply_trans int pos trs state = (apply_union int pos trs state, NONE) |
| 15531 | 392 |
handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); |
| 6689 | 393 |
|
394 |
end; |
|
| 5828 | 395 |
|
396 |
||
397 |
(* datatype transition *) |
|
398 |
||
399 |
datatype transition = Transition of |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
400 |
{name: string, (*command name*)
|
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
401 |
pos: Position.T, (*source position*) |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
402 |
int_only: bool, (*interactive-only*) |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
403 |
print: bool, (*print result state*) |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
404 |
no_timing: bool, (*suppress timing*) |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
405 |
trans: trans list}; (*primitive transitions (union)*) |
| 5828 | 406 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
407 |
fun make_transition (name, pos, int_only, print, no_timing, trans) = |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
408 |
Transition {name = name, pos = pos, int_only = int_only, print = print, no_timing = no_timing,
|
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
409 |
trans = trans}; |
| 5828 | 410 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
411 |
fun map_transition f (Transition {name, pos, int_only, print, no_timing, trans}) =
|
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
412 |
make_transition (f (name, pos, int_only, print, no_timing, trans)); |
| 5828 | 413 |
|
| 27441 | 414 |
val empty = make_transition ("", Position.none, false, false, false, []);
|
| 5828 | 415 |
|
416 |
||
417 |
(* diagnostics *) |
|
418 |
||
| 27441 | 419 |
fun init_of (Transition {trans = [Init (name, _, _)], ...}) = SOME name
|
420 |
| init_of _ = NONE; |
|
421 |
||
| 27427 | 422 |
fun name_of (Transition {name, ...}) = name;
|
423 |
fun str_of (Transition {name, pos, ...}) = quote name ^ Position.str_of pos;
|
|
| 5828 | 424 |
|
| 27427 | 425 |
fun command_msg msg tr = msg ^ "command " ^ str_of tr; |
| 5828 | 426 |
fun at_command tr = command_msg "At " tr ^ "."; |
427 |
||
428 |
fun type_error tr state = |
|
| 18685 | 429 |
ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state); |
| 5828 | 430 |
|
431 |
||
432 |
(* modify transitions *) |
|
433 |
||
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
434 |
fun name nm = map_transition (fn (_, pos, int_only, print, no_timing, trans) => |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
435 |
(nm, pos, int_only, print, no_timing, trans)); |
| 9010 | 436 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
437 |
fun position pos = map_transition (fn (name, _, int_only, print, no_timing, trans) => |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
438 |
(name, pos, int_only, print, no_timing, trans)); |
| 5828 | 439 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
440 |
fun interactive int_only = map_transition (fn (name, pos, _, print, no_timing, trans) => |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
441 |
(name, pos, int_only, print, no_timing, trans)); |
| 14923 | 442 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
443 |
val no_timing = map_transition (fn (name, pos, int_only, print, _, trans) => |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
444 |
(name, pos, int_only, print, true, trans)); |
| 17363 | 445 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
446 |
fun add_trans tr = map_transition (fn (name, pos, int_only, print, no_timing, trans) => |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
447 |
(name, pos, int_only, print, no_timing, trans @ [tr])); |
| 16607 | 448 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
449 |
val print = map_transition (fn (name, pos, int_only, _, no_timing, trans) => |
|
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
450 |
(name, pos, int_only, true, no_timing, trans)); |
| 5828 | 451 |
|
452 |
||
| 21007 | 453 |
(* basic transitions *) |
| 5828 | 454 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
455 |
fun init_theory name f exit = add_trans (Init (name, f, exit)); |
| 6689 | 456 |
val exit = add_trans Exit; |
| 7612 | 457 |
val keep' = add_trans o Keep; |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
458 |
fun transaction f = add_trans (Transaction f); |
| 5828 | 459 |
|
| 7612 | 460 |
fun keep f = add_trans (Keep (fn _ => f)); |
| 5828 | 461 |
fun imperative f = keep (fn _ => f ()); |
462 |
||
| 21007 | 463 |
val unknown_theory = imperative (fn () => warning "Unknown theory context"); |
464 |
val unknown_proof = imperative (fn () => warning "Unknown proof context"); |
|
465 |
val unknown_context = imperative (fn () => warning "Unknown context"); |
|
| 15668 | 466 |
|
| 21007 | 467 |
|
468 |
(* theory transitions *) |
|
| 15668 | 469 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
470 |
fun generic_theory f = transaction (fn _ => |
| 26491 | 471 |
(fn Theory (gthy, _) => Theory (f gthy, NONE) |
472 |
| _ => raise UNDEF)); |
|
473 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
474 |
fun theory' f = transaction (fn int => |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
475 |
(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
|
476 |
| _ => raise UNDEF)); |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
477 |
|
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
478 |
fun theory f = theory' (K f); |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
479 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
480 |
fun begin_local_theory begin f = transaction (fn _ => |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
481 |
(fn Theory (Context.Theory thy, _) => |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
482 |
let |
| 20985 | 483 |
val lthy = f thy; |
| 21294 | 484 |
val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy); |
485 |
in Theory (gthy, SOME lthy) end |
|
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
486 |
| _ => raise UNDEF)); |
| 17076 | 487 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
488 |
val end_local_theory = transaction (fn _ => |
| 21294 | 489 |
(fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy) |
| 21007 | 490 |
| _ => raise UNDEF)); |
491 |
||
492 |
local |
|
493 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
494 |
fun local_theory_presentation loc f g = transaction (fn int => |
| 21294 | 495 |
(fn Theory (gthy, _) => |
496 |
let |
|
497 |
val finish = loc_finish loc gthy; |
|
| 25960 | 498 |
val lthy' = f (loc_begin loc gthy); |
| 21294 | 499 |
in Theory (finish lthy', SOME lthy') end |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
500 |
| _ => raise UNDEF) #> tap (g int)); |
| 15668 | 501 |
|
| 21007 | 502 |
in |
503 |
||
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
504 |
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
|
505 |
fun present_local_theory loc g = local_theory_presentation loc I g; |
| 18955 | 506 |
|
| 21007 | 507 |
end; |
508 |
||
509 |
||
510 |
(* proof transitions *) |
|
511 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
512 |
fun end_proof f = transaction (fn int => |
|
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
513 |
(fn Proof (prf, (finish, _)) => |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
514 |
let val state = ProofNode.current prf in |
| 21007 | 515 |
if can (Proof.assert_bottom true) state then |
516 |
let |
|
517 |
val ctxt' = f int state; |
|
518 |
val gthy' = finish ctxt'; |
|
519 |
in Theory (gthy', SOME ctxt') end |
|
520 |
else raise UNDEF |
|
521 |
end |
|
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
522 |
| SkipProof (0, (gthy, _)) => Theory (gthy, NONE) |
| 21007 | 523 |
| _ => raise UNDEF)); |
524 |
||
| 21294 | 525 |
local |
526 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
527 |
fun begin_proof init finish = transaction (fn int => |
| 21294 | 528 |
(fn Theory (gthy, _) => |
529 |
let |
|
| 24453 | 530 |
val prf = init int gthy; |
| 21294 | 531 |
val schematic = Proof.schematic_goal prf; |
532 |
in |
|
533 |
if ! skip_proofs andalso schematic then |
|
534 |
warning "Cannot skip proof of schematic goal statement" |
|
535 |
else (); |
|
536 |
if ! skip_proofs andalso not schematic then |
|
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
537 |
SkipProof (0, (finish gthy (Proof.global_skip_proof int prf), gthy)) |
|
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
538 |
else Proof (ProofNode.init prf, (finish gthy, gthy)) |
| 21294 | 539 |
end |
540 |
| _ => raise UNDEF)); |
|
541 |
||
542 |
in |
|
543 |
||
|
24780
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
544 |
fun local_theory_to_proof' loc f = begin_proof |
| 25960 | 545 |
(fn int => fn gthy => f int (loc_begin loc gthy)) |
|
24780
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
546 |
(loc_finish loc); |
|
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
547 |
|
| 24453 | 548 |
fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f); |
| 21294 | 549 |
|
550 |
fun theory_to_proof f = begin_proof |
|
|
24780
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
551 |
(K (fn Context.Theory thy => f thy | _ => raise UNDEF)) |
|
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
552 |
(K (Context.Theory o ProofContext.theory_of)); |
| 21294 | 553 |
|
554 |
end; |
|
555 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
556 |
val forget_proof = transaction (fn _ => |
| 21007 | 557 |
(fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) |
558 |
| SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) |
|
559 |
| _ => raise UNDEF)); |
|
560 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
561 |
fun present_proof f = transaction (fn int => |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
562 |
(fn Proof (prf, x) => Proof (ProofNode.apply I prf, x) |
|
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
563 |
| skip as SkipProof _ => skip |
| 21177 | 564 |
| _ => raise UNDEF) #> tap (f int)); |
565 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
566 |
fun proofs' f = transaction (fn int => |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
567 |
(fn Proof (prf, x) => Proof (ProofNode.applys (f int) prf, x) |
|
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
568 |
| skip as SkipProof _ => skip |
| 16815 | 569 |
| _ => raise UNDEF)); |
| 15668 | 570 |
|
|
17904
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents:
17513
diff
changeset
|
571 |
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
|
572 |
val proofs = proofs' o K; |
| 6689 | 573 |
val proof = proof' o K; |
| 16815 | 574 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
575 |
fun actual_proof f = transaction (fn _ => |
| 21007 | 576 |
(fn Proof (prf, x) => Proof (f prf, x) |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
577 |
| _ => raise UNDEF)); |
| 16815 | 578 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
579 |
fun skip_proof f = transaction (fn _ => |
| 21007 | 580 |
(fn SkipProof (h, x) => SkipProof (f h, x) |
| 18563 | 581 |
| _ => raise UNDEF)); |
582 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
583 |
fun skip_proof_to_theory pred = transaction (fn _ => |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
584 |
(fn SkipProof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF |
|
17904
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents:
17513
diff
changeset
|
585 |
| _ => raise UNDEF)); |
| 5828 | 586 |
|
587 |
||
588 |
||
589 |
(** toplevel transactions **) |
|
590 |
||
| 27427 | 591 |
(* identification *) |
592 |
||
593 |
fun get_id (Transition {pos, ...}) = Position.get_id pos;
|
|
594 |
fun put_id id (tr as Transition {pos, ...}) = position (Position.put_id id pos) tr;
|
|
595 |
||
596 |
||
| 25960 | 597 |
(* thread position *) |
| 25799 | 598 |
|
| 25960 | 599 |
fun setmp_thread_position (Transition {pos, ...}) f x =
|
|
25819
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
wenzelm
parents:
25809
diff
changeset
|
600 |
Position.setmp_thread_data pos f x; |
| 25799 | 601 |
|
| 27606 | 602 |
fun status tr m = |
603 |
setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) (); |
|
604 |
||
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
605 |
fun error_msg tr exn_info = |
|
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
606 |
setmp_thread_position tr (fn () => Output.error_msg (exn_message (EXCURSION_FAIL exn_info))) (); |
|
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
607 |
|
| 25799 | 608 |
|
| 5828 | 609 |
(* apply transitions *) |
610 |
||
| 6664 | 611 |
local |
612 |
||
| 25799 | 613 |
fun app int (tr as Transition {trans, pos, int_only, print, no_timing, ...}) =
|
|
25819
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
wenzelm
parents:
25809
diff
changeset
|
614 |
setmp_thread_position tr (fn state => |
| 25799 | 615 |
let |
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
616 |
val _ = if not int andalso int_only then warning (command_msg "Interactive-only " tr) else (); |
| 16682 | 617 |
|
| 25799 | 618 |
fun do_timing f x = (warning (command_msg "" tr); timeap f x); |
619 |
fun do_profiling f x = profile (! profiling) f x; |
|
620 |
||
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
621 |
val (result, status) = |
| 25799 | 622 |
state |> (apply_trans int pos trans |
623 |
|> (if ! profiling > 0 andalso not no_timing then do_profiling else I) |
|
624 |
|> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I)); |
|
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
625 |
|
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
626 |
val _ = if int andalso not (! quiet) andalso print then print_state false result else (); |
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
627 |
in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) status) end); |
| 6664 | 628 |
|
629 |
in |
|
| 5828 | 630 |
|
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
631 |
fun transition int tr st = |
|
26256
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
632 |
let val ctxt = try context_of st in |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
633 |
(case app int tr st of |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
634 |
(_, SOME TERMINATE) => NONE |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
635 |
| (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
636 |
| (state', SOME exn) => SOME (state', SOME (exn_context ctxt exn, at_command tr)) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
637 |
| (state', NONE) => SOME (state', NONE)) |
|
3e7939e978c6
added exception CONTEXT, indicating context of another exception;
wenzelm
parents:
26081
diff
changeset
|
638 |
end; |
| 6664 | 639 |
|
640 |
end; |
|
| 5828 | 641 |
|
642 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
643 |
(* commit final exit *) |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
644 |
|
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
645 |
val commit_exit = |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
646 |
name "end" empty |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
647 |
|> add_trans CommitExit |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
648 |
|> imperative (fn () => warning "Expected to find finished theory"); |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
649 |
|
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
650 |
|
| 17076 | 651 |
(* excursion: toplevel -- apply transformers/presentation -- toplevel *) |
| 5828 | 652 |
|
| 6664 | 653 |
local |
654 |
||
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
655 |
fun no_pr _ _ _ = (); |
| 5828 | 656 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
657 |
fun excur (tr, pr) (st, res) = |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
658 |
(case transition (! interact) tr st of |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
659 |
SOME (st', NONE) => |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
660 |
(st', setmp_thread_position tr (fn () => pr st st' res) () handle exn => |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
661 |
raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr)) |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
662 |
| SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
663 |
| NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr)); |
|
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
664 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
665 |
fun excurs trs st_res = |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
666 |
let val (st', res') = fold excur trs st_res in |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
667 |
if is_toplevel st' then (excur (commit_exit, no_pr) (st', ()); res') |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
668 |
else error "Unfinished development at end of input" |
|
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
669 |
end; |
| 17076 | 670 |
|
| 6664 | 671 |
in |
672 |
||
| 17076 | 673 |
fun present_excursion trs res = |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
674 |
excurs trs (toplevel, res) handle exn => error (exn_message exn); |
| 9134 | 675 |
|
| 17076 | 676 |
fun excursion trs = present_excursion (map (rpair no_pr) trs) (); |
| 7062 | 677 |
|
| 6664 | 678 |
end; |
679 |
||
| 5828 | 680 |
end; |