| author | wenzelm |
| Thu, 29 Dec 2022 12:08:58 +0100 | |
| changeset 76810 | 3f211d126ddc |
| parent 76809 | f05327293f07 |
| child 76811 | 56d76e8cecf4 |
| permissions | -rw-r--r-- |
| 5828 | 1 |
(* Title: Pure/Isar/toplevel.ML |
2 |
Author: Markus Wenzel, TU Muenchen |
|
3 |
||
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
4 |
Isabelle/Isar toplevel transactions. |
| 5828 | 5 |
*) |
6 |
||
7 |
signature TOPLEVEL = |
|
8 |
sig |
|
| 19063 | 9 |
exception UNDEF |
| 5828 | 10 |
type state |
| 76415 | 11 |
val make_state: theory option -> state |
| 7732 | 12 |
val is_toplevel: state -> bool |
| 18589 | 13 |
val is_theory: state -> bool |
14 |
val is_proof: state -> bool |
|
| 51555 | 15 |
val is_skipped_proof: state -> bool |
| 17076 | 16 |
val level: state -> int |
| 67391 | 17 |
val previous_theory_of: state -> theory option |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
18 |
val output_of: state -> Latex.text option |
| 21506 | 19 |
val context_of: state -> Proof.context |
| 22089 | 20 |
val generic_theory_of: state -> generic_theory |
| 5828 | 21 |
val theory_of: state -> theory |
22 |
val proof_of: state -> Proof.state |
|
| 18589 | 23 |
val proof_position_of: state -> int |
| 68505 | 24 |
val is_end_theory: state -> bool |
|
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
25 |
val end_theory: Position.T -> state -> theory |
| 67381 | 26 |
val presentation_context: state -> Proof.context |
27 |
val presentation_state: Proof.context -> state |
|
| 56893 | 28 |
val pretty_context: state -> Pretty.T list |
|
56887
1ca814da47ae
clarified print_state, which goes back to TTY loop before Proof General, and before separate print_context;
wenzelm
parents:
56867
diff
changeset
|
29 |
val pretty_state: state -> Pretty.T list |
| 61208 | 30 |
val string_of_state: state -> string |
| 37858 | 31 |
val pretty_abstract: state -> Pretty.T |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
32 |
type presentation = state -> Latex.text option |
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
33 |
val presentation: (state -> Latex.text) -> presentation |
|
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
34 |
val no_presentation: presentation |
| 16682 | 35 |
type transition |
| 5828 | 36 |
val empty: transition |
| 27427 | 37 |
val name_of: transition -> string |
| 28105 | 38 |
val pos_of: transition -> Position.T |
| 69877 | 39 |
val timing_of: transition -> Time.time |
| 60076 | 40 |
val type_error: transition -> string |
| 5828 | 41 |
val name: string -> transition -> transition |
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
42 |
val position: Position.T -> transition -> transition |
| 69887 | 43 |
val markers: Input.source list -> transition -> transition |
| 69877 | 44 |
val timing: Time.time -> transition -> transition |
|
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
45 |
val init_theory: (unit -> theory) -> transition -> transition |
|
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
46 |
val is_init: transition -> bool |
|
44186
806f0ec1a43d
simplified Toplevel.init_theory: discontinued special master argument;
wenzelm
parents:
44185
diff
changeset
|
47 |
val modify_init: (unit -> theory) -> transition -> transition |
| 6689 | 48 |
val exit: transition -> transition |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
49 |
val present: (state -> Latex.text) -> transition -> transition |
| 5828 | 50 |
val keep: (state -> unit) -> transition -> transition |
| 7612 | 51 |
val keep': (bool -> state -> unit) -> transition -> transition |
|
60190
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
52 |
val keep_proof: (state -> unit) -> transition -> transition |
|
51268
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents:
51241
diff
changeset
|
53 |
val is_ignored: transition -> bool |
| 74672 | 54 |
val is_malformed: transition -> bool |
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
55 |
val ignored: Position.T -> transition |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
56 |
val malformed: Position.T -> string -> transition |
| 26491 | 57 |
val generic_theory: (generic_theory -> generic_theory) -> transition -> transition |
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
58 |
val theory': (bool -> theory -> theory) -> presentation -> transition -> transition |
|
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
59 |
val theory: (theory -> theory) -> transition -> transition |
| 72434 | 60 |
val begin_main_target: bool -> (theory -> local_theory) -> transition -> transition |
61 |
val end_main_target: transition -> transition |
|
| 72453 | 62 |
val begin_nested_target: (Context.generic -> local_theory) -> transition -> transition |
| 72434 | 63 |
val end_nested_target: transition -> transition |
|
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59923
diff
changeset
|
64 |
val local_theory': (bool * Position.T) option -> (xstring * Position.T) option -> |
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
65 |
(bool -> local_theory -> local_theory) -> presentation -> transition -> transition |
|
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59923
diff
changeset
|
66 |
val local_theory: (bool * Position.T) option -> (xstring * Position.T) option -> |
|
59923
b21c82422d65
support private scope for individual local theory commands;
wenzelm
parents:
59472
diff
changeset
|
67 |
(local_theory -> local_theory) -> transition -> transition |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
68 |
val present_local_theory: (xstring * Position.T) option -> (state -> Latex.text) -> |
| 24453 | 69 |
transition -> transition |
|
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59923
diff
changeset
|
70 |
val local_theory_to_proof': (bool * Position.T) option -> (xstring * Position.T) option -> |
|
45488
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
71 |
(bool -> local_theory -> Proof.state) -> transition -> transition |
|
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59923
diff
changeset
|
72 |
val local_theory_to_proof: (bool * Position.T) option -> (xstring * Position.T) option -> |
|
45488
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
73 |
(local_theory -> Proof.state) -> transition -> transition |
| 17363 | 74 |
val theory_to_proof: (theory -> Proof.state) -> transition -> transition |
| 21007 | 75 |
val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition |
|
68876
cefaac3d24ff
no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68875
diff
changeset
|
76 |
val forget_proof: transition -> transition |
|
49863
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
wenzelm
parents:
49062
diff
changeset
|
77 |
val proofs': (bool -> Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition |
|
17904
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents:
17513
diff
changeset
|
78 |
val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition |
|
49863
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
wenzelm
parents:
49062
diff
changeset
|
79 |
val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition |
| 21177 | 80 |
val proof: (Proof.state -> Proof.state) -> transition -> transition |
| 33390 | 81 |
val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition |
|
60693
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
82 |
val skip_proof: (unit -> unit) -> transition -> transition |
|
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
83 |
val skip_proof_open: transition -> transition |
|
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
84 |
val skip_proof_close: transition -> transition |
| 52536 | 85 |
val exec_id: Document_ID.exec -> transition -> transition |
| 28425 | 86 |
val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
|
|
59055
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
87 |
val transition: bool -> transition -> state -> state * (exn * string) option |
| 51323 | 88 |
val command_errors: bool -> transition -> state -> Runtime.error list * state option |
|
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
89 |
val command_exception: bool -> transition -> state -> state |
| 56937 | 90 |
val reset_theory: state -> state option |
91 |
val reset_proof: state -> state option |
|
|
68877
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
92 |
val reset_notepad: state -> state option |
| 69887 | 93 |
val fork_presentation: transition -> transition * transition |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
94 |
type result |
|
73687
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
95 |
val join_results: result -> (state * transition * state) list |
| 68839 | 96 |
val element_result: Keyword.keywords -> transition Thy_Element.element -> state -> result * state |
| 5828 | 97 |
end; |
98 |
||
| 6965 | 99 |
structure Toplevel: TOPLEVEL = |
| 5828 | 100 |
struct |
101 |
||
102 |
(** toplevel state **) |
|
103 |
||
|
31476
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
31431
diff
changeset
|
104 |
exception UNDEF = Runtime.UNDEF; |
| 19063 | 105 |
|
106 |
||
|
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
107 |
(* datatype node *) |
| 21294 | 108 |
|
| 5828 | 109 |
datatype node = |
| 69883 | 110 |
Toplevel |
111 |
(*toplevel outside of theory body*) | |
|
| 69878 | 112 |
Theory of generic_theory |
113 |
(*global or local theory*) | |
|
| 33390 | 114 |
Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory) |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
115 |
(*proof node, finish, original theory*) | |
| 51555 | 116 |
Skipped_Proof of int * (generic_theory * generic_theory); |
|
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
117 |
(*proof depth, resulting theory, original theory*) |
| 5828 | 118 |
|
| 69878 | 119 |
val theory_node = fn Theory gthy => SOME gthy | _ => NONE; |
| 18589 | 120 |
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; |
| 51555 | 121 |
val skipped_proof_node = fn Skipped_Proof _ => true | _ => false; |
| 18589 | 122 |
|
| 69883 | 123 |
fun cases_node f _ _ Toplevel = f () |
124 |
| cases_node _ g _ (Theory gthy) = g gthy |
|
125 |
| cases_node _ _ h (Proof (prf, _)) = h (Proof_Node.current prf) |
|
126 |
| cases_node _ g _ (Skipped_Proof (_, (gthy, _))) = g gthy; |
|
127 |
||
128 |
fun cases_proper_node g h = cases_node (fn () => raise UNDEF) g h; |
|
129 |
||
130 |
val get_theory = cases_node (K NONE) (SOME o Context.theory_of) (SOME o Proof.theory_of); |
|
| 19063 | 131 |
|
|
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
132 |
|
|
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
133 |
(* datatype state *) |
|
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
134 |
|
| 69883 | 135 |
type node_presentation = node * Proof.context; |
| 5828 | 136 |
|
| 69886 | 137 |
fun init_presentation () = |
138 |
Proof_Context.init_global (Theory.get_pure_bootstrap ()); |
|
139 |
||
140 |
fun node_presentation node = |
|
141 |
(node, cases_node init_presentation Context.proof_of Proof.context_of node); |
|
142 |
||
| 69883 | 143 |
datatype state = |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
144 |
State of node_presentation * (theory option * Latex.text option); |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
145 |
(*current node with presentation context, previous theory, document output*) |
|
62895
54c2abe7e9a4
treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
wenzelm
parents:
62889
diff
changeset
|
146 |
|
| 69883 | 147 |
fun node_of (State ((node, _), _)) = node; |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
148 |
fun previous_theory_of (State (_, (prev_thy, _))) = prev_thy; |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
149 |
fun output_of (State (_, (_, output))) = output; |
| 5828 | 150 |
|
| 76415 | 151 |
fun make_state opt_thy = |
152 |
let val node = (case opt_thy of NONE => Toplevel | SOME thy => Theory (Context.Theory thy)) |
|
153 |
in State (node_presentation node, (NONE, NONE)) end; |
|
| 69883 | 154 |
|
155 |
fun level state = |
|
156 |
(case node_of state of |
|
157 |
Toplevel => 0 |
|
158 |
| Theory _ => 0 |
|
159 |
| Proof (prf, _) => Proof.level (Proof_Node.current prf) |
|
160 |
| Skipped_Proof (d, _) => d + 1); (*different notion of proof depth!*) |
|
| 17076 | 161 |
|
| 69883 | 162 |
fun str_of_state state = |
163 |
(case node_of state of |
|
164 |
Toplevel => |
|
165 |
(case previous_theory_of state of |
|
166 |
NONE => "at top level" |
|
167 |
| SOME thy => "at top level, result theory " ^ quote (Context.theory_name thy)) |
|
168 |
| Theory (Context.Theory _) => "in theory mode" |
|
169 |
| Theory (Context.Proof _) => "in local theory mode" |
|
170 |
| Proof _ => "in proof mode" |
|
171 |
| Skipped_Proof _ => "in skipped proof mode"); |
|
|
5946
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
172 |
|
|
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
173 |
|
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
174 |
(* current node *) |
| 5828 | 175 |
|
| 69883 | 176 |
fun is_toplevel state = (case node_of state of Toplevel => true | _ => false); |
| 5828 | 177 |
|
| 69883 | 178 |
fun is_theory state = |
179 |
not (is_toplevel state) andalso is_some (theory_node (node_of state)); |
|
| 18589 | 180 |
|
| 69883 | 181 |
fun is_proof state = |
182 |
not (is_toplevel state) andalso is_some (proof_node (node_of state)); |
|
| 5828 | 183 |
|
| 69883 | 184 |
fun is_skipped_proof state = |
185 |
not (is_toplevel state) andalso skipped_proof_node (node_of state); |
|
| 30801 | 186 |
|
| 69883 | 187 |
fun proper_node_of state = if is_toplevel state then raise UNDEF else node_of state; |
188 |
fun proper_node_case f g state = cases_proper_node f g (proper_node_of state); |
|
189 |
||
190 |
val context_of = proper_node_case Context.proof_of Proof.context_of; |
|
191 |
val generic_theory_of = proper_node_case I (Context.Proof o Proof.context_of); |
|
192 |
val theory_of = proper_node_case Context.theory_of Proof.theory_of; |
|
193 |
val proof_of = proper_node_case (fn _ => error "No proof state") I; |
|
| 17208 | 194 |
|
| 18589 | 195 |
fun proof_position_of state = |
| 69883 | 196 |
(case proper_node_of state of |
| 33390 | 197 |
Proof (prf, _) => Proof_Node.position prf |
| 60094 | 198 |
| _ => ~1); |
| 6664 | 199 |
|
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
200 |
fun is_end_theory (State ((Toplevel, _), (SOME _, _))) = true |
| 68505 | 201 |
| is_end_theory _ = false; |
202 |
||
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
203 |
fun end_theory _ (State ((Toplevel, _), (SOME thy, _))) = thy |
| 68869 | 204 |
| end_theory pos _ = error ("Malformed theory" ^ Position.here pos);
|
|
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
205 |
|
| 5828 | 206 |
|
| 67381 | 207 |
(* presentation context *) |
208 |
||
209 |
structure Presentation_State = Proof_Data |
|
210 |
( |
|
211 |
type T = state option; |
|
212 |
fun init _ = NONE; |
|
213 |
); |
|
214 |
||
| 69883 | 215 |
fun presentation_context0 (State ((_, pr_ctxt), _)) = pr_ctxt; |
|
67642
10ff1f077119
more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents:
67641
diff
changeset
|
216 |
|
|
10ff1f077119
more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents:
67641
diff
changeset
|
217 |
fun presentation_context (state as State (current, _)) = |
|
10ff1f077119
more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents:
67641
diff
changeset
|
218 |
presentation_context0 state |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
219 |
|> Presentation_State.put (SOME (State (current, (NONE, NONE)))); |
| 67381 | 220 |
|
221 |
fun presentation_state ctxt = |
|
222 |
(case Presentation_State.get ctxt of |
|
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
223 |
NONE => State (node_presentation (Theory (Context.Proof ctxt)), (NONE, NONE)) |
| 67381 | 224 |
| SOME state => state); |
225 |
||
226 |
||
| 16815 | 227 |
(* print state *) |
228 |
||
| 56893 | 229 |
fun pretty_context state = |
| 69883 | 230 |
if is_toplevel state then [] |
231 |
else |
|
232 |
let |
|
233 |
val gthy = |
|
234 |
(case node_of state of |
|
235 |
Toplevel => raise Match |
|
236 |
| Theory gthy => gthy |
|
237 |
| Proof (_, (_, gthy)) => gthy |
|
238 |
| Skipped_Proof (_, (_, gthy)) => gthy); |
|
239 |
val lthy = Context.cases Named_Target.theory_init I gthy; |
|
240 |
in Local_Theory.pretty lthy end; |
|
| 16815 | 241 |
|
|
56887
1ca814da47ae
clarified print_state, which goes back to TTY loop before Proof General, and before separate print_context;
wenzelm
parents:
56867
diff
changeset
|
242 |
fun pretty_state state = |
| 69883 | 243 |
(case node_of state of |
244 |
Toplevel => [] |
|
245 |
| Theory _ => [] |
|
246 |
| Proof (prf, _) => Proof.pretty_state (Proof_Node.current prf) |
|
247 |
| Skipped_Proof (d, _) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]);
|
|
| 56867 | 248 |
|
| 61208 | 249 |
val string_of_state = pretty_state #> Pretty.chunks #> Pretty.string_of; |
| 16815 | 250 |
|
| 37858 | 251 |
fun pretty_abstract state = Pretty.str ("<Isar " ^ str_of_state state ^ ">");
|
252 |
||
|
62819
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
wenzelm
parents:
62663
diff
changeset
|
253 |
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_abstract); |
| 62663 | 254 |
|
| 16815 | 255 |
|
| 15668 | 256 |
|
| 5828 | 257 |
(** toplevel transitions **) |
258 |
||
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
259 |
(* presentation *) |
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
260 |
|
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
261 |
type presentation = state -> Latex.text option; |
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
262 |
fun presentation g : presentation = SOME o g; |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
263 |
val no_presentation: presentation = K NONE; |
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
264 |
|
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
265 |
|
| 69888 | 266 |
(* primitive transitions *) |
267 |
||
268 |
datatype trans = |
|
269 |
(*init theory*) |
|
270 |
Init of unit -> theory | |
|
271 |
(*formal exit of theory*) |
|
272 |
Exit | |
|
| 76411 | 273 |
(*keep state unchanged*) |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
274 |
Keep of bool -> presentation | |
| 69888 | 275 |
(*node transaction and presentation*) |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
276 |
Transaction of (bool -> node -> node_presentation) * presentation; |
| 69888 | 277 |
|
278 |
local |
|
| 7022 | 279 |
|
|
31476
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
31431
diff
changeset
|
280 |
exception FAILURE of state * exn; |
|
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
31431
diff
changeset
|
281 |
|
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
282 |
fun apply_presentation g (st as State (node, (prev_thy, _))) = |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
283 |
State (node, (prev_thy, g st)); |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
284 |
|
| 76809 | 285 |
fun no_update f node state = |
286 |
Runtime.controlled_execution (try generic_theory_of state) |
|
287 |
(fn () => |
|
288 |
let |
|
289 |
val prev_thy = previous_theory_of state; |
|
290 |
val state' = State (node_presentation node, (prev_thy, NONE)); |
|
291 |
in apply_presentation f state' end) () |
|
292 |
||
293 |
fun update f g node = |
|
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
294 |
let |
| 69886 | 295 |
val node_pr = node_presentation node; |
| 69883 | 296 |
val context = cases_proper_node I (Context.Proof o Proof.context_of) node; |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
297 |
fun make_state node_pr' = State (node_pr', (get_theory node, NONE)); |
| 6689 | 298 |
|
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
299 |
val (st', err) = |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
300 |
(Runtime.controlled_execution (SOME context) (f #> make_state #> apply_presentation g) node, |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
301 |
NONE) handle exn => (make_state node_pr, SOME exn); |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
302 |
in |
|
52696
38466f4f3483
immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents:
52565
diff
changeset
|
303 |
(case err of |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
304 |
NONE => st' |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
305 |
| SOME exn => raise FAILURE (st', exn)) |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
306 |
end; |
| 6689 | 307 |
|
| 69887 | 308 |
fun apply_tr int trans state = |
309 |
(case (trans, node_of state) of |
|
310 |
(Init f, Toplevel) => |
|
311 |
Runtime.controlled_execution NONE (fn () => |
|
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
312 |
State (node_presentation (Theory (Context.Theory (f ()))), (NONE, NONE))) () |
| 69887 | 313 |
| (Exit, node as Theory (Context.Theory thy)) => |
| 69883 | 314 |
let |
315 |
val State ((node', pr_ctxt), _) = |
|
| 76809 | 316 |
node |> update |
|
71023
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
wenzelm
parents:
70398
diff
changeset
|
317 |
(fn _ => |
|
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
wenzelm
parents:
70398
diff
changeset
|
318 |
node_presentation |
|
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
wenzelm
parents:
70398
diff
changeset
|
319 |
(Theory (Context.Theory (tap Thm.expose_theory (Theory.end_theory thy))))) |
|
74833
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
320 |
no_presentation; |
|
fe9e590ae52f
output for document commands like 'section', 'text' is defined in user-space, as part of the command transaction;
wenzelm
parents:
74831
diff
changeset
|
321 |
in State ((Toplevel, pr_ctxt), (get_theory node', NONE)) end |
| 76809 | 322 |
| (Keep f, node) => no_update (fn x => f int x) node state |
| 69887 | 323 |
| (Transaction _, Toplevel) => raise UNDEF |
| 76809 | 324 |
| (Transaction (f, g), node) => update (fn x => f int x) g node |
| 69887 | 325 |
| _ => raise UNDEF); |
| 5828 | 326 |
|
| 32792 | 327 |
fun apply_union _ [] state = raise FAILURE (state, UNDEF) |
328 |
| apply_union int (tr :: trs) state = |
|
329 |
apply_union int trs state |
|
330 |
handle Runtime.UNDEF => apply_tr int tr state |
|
| 76810 | 331 |
| FAILURE (alt_state, Runtime.UNDEF) => apply_tr int tr alt_state |
| 6689 | 332 |
| exn as FAILURE _ => raise exn |
333 |
| exn => raise FAILURE (state, exn); |
|
334 |
||
| 70134 | 335 |
fun apply_markers name markers (state as State ((node, pr_ctxt), prev_thy)) = |
| 69887 | 336 |
let |
337 |
val state' = |
|
338 |
Runtime.controlled_execution (try generic_theory_of state) |
|
| 70134 | 339 |
(fn () => State ((node, fold (Document_Marker.evaluate name) markers pr_ctxt), prev_thy)) (); |
| 69887 | 340 |
in (state', NONE) end |
341 |
handle exn => (state, SOME exn); |
|
342 |
||
| 6689 | 343 |
in |
344 |
||
| 70134 | 345 |
fun apply_trans int name markers trans state = |
346 |
(apply_union int trans state |> apply_markers name markers) |
|
| 15531 | 347 |
handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); |
| 6689 | 348 |
|
349 |
end; |
|
| 5828 | 350 |
|
351 |
||
352 |
(* datatype transition *) |
|
353 |
||
354 |
datatype transition = Transition of |
|
| 69887 | 355 |
{name: string, (*command name*)
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
356 |
pos: Position.T, (*source position*) |
| 69887 | 357 |
markers: Input.source list, (*semantic document markers*) |
358 |
timing: Time.time, (*prescient timing information*) |
|
359 |
trans: trans list}; (*primitive transitions (union)*) |
|
| 5828 | 360 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
361 |
fun make_transition (name, pos, markers, timing, trans) = |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
362 |
Transition {name = name, pos = pos, markers = markers, timing = timing, trans = trans};
|
| 5828 | 363 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
364 |
fun map_transition f (Transition {name, pos, markers, timing, trans}) =
|
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
365 |
make_transition (f (name, pos, markers, timing, trans)); |
| 5828 | 366 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
367 |
val empty = make_transition ("", Position.none, [], Time.zeroTime, []);
|
| 5828 | 368 |
|
369 |
||
370 |
(* diagnostics *) |
|
371 |
||
| 27427 | 372 |
fun name_of (Transition {name, ...}) = name;
|
| 28105 | 373 |
fun pos_of (Transition {pos, ...}) = pos;
|
| 69877 | 374 |
fun timing_of (Transition {timing, ...}) = timing;
|
| 5828 | 375 |
|
| 60076 | 376 |
fun command_msg msg tr = |
377 |
msg ^ "command " ^ quote (Markup.markup Markup.keyword1 (name_of tr)) ^ |
|
378 |
Position.here (pos_of tr); |
|
| 5828 | 379 |
|
| 60076 | 380 |
fun at_command tr = command_msg "At " tr; |
381 |
fun type_error tr = command_msg "Bad context for " tr; |
|
| 5828 | 382 |
|
383 |
||
384 |
(* modify transitions *) |
|
385 |
||
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
386 |
fun name name = map_transition (fn (_, pos, markers, timing, trans) => |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
387 |
(name, pos, markers, timing, trans)); |
| 9010 | 388 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
389 |
fun position pos = map_transition (fn (name, _, markers, timing, trans) => |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
390 |
(name, pos, markers, timing, trans)); |
| 69887 | 391 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
392 |
fun markers markers = map_transition (fn (name, pos, _, timing, trans) => |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
393 |
(name, pos, markers, timing, trans)); |
| 14923 | 394 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
395 |
fun timing timing = map_transition (fn (name, pos, markers, _, trans) => |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
396 |
(name, pos, markers, timing, trans)); |
| 69877 | 397 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
398 |
fun add_trans tr = map_transition (fn (name, pos, markers, timing, trans) => |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
399 |
(name, pos, markers, timing, tr :: trans)); |
| 16607 | 400 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
401 |
val reset_trans = map_transition (fn (name, pos, markers, timing, _) => |
|
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
402 |
(name, pos, markers, timing, [])); |
| 5828 | 403 |
|
404 |
||
| 21007 | 405 |
(* basic transitions *) |
| 5828 | 406 |
|
|
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
407 |
fun init_theory f = add_trans (Init f); |
|
37977
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents:
37953
diff
changeset
|
408 |
|
|
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
409 |
fun is_init (Transition {trans = [Init _], ...}) = true
|
|
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
410 |
| is_init _ = false; |
|
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
411 |
|
|
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
412 |
fun modify_init f tr = if is_init tr then init_theory f (reset_trans tr) else tr; |
|
37977
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents:
37953
diff
changeset
|
413 |
|
| 6689 | 414 |
val exit = add_trans Exit; |
|
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
415 |
|
|
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
416 |
fun present_transaction f g = add_trans (Transaction (f, g)); |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
417 |
fun transaction f = present_transaction f no_presentation; |
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
418 |
fun transaction0 f = present_transaction (node_presentation oo f) no_presentation; |
| 5828 | 419 |
|
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
420 |
fun present f = add_trans (Keep (K (presentation f))); |
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
421 |
fun keep f = add_trans (Keep (fn _ => fn st => let val () = f st in NONE end)); |
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
422 |
fun keep' f = add_trans (Keep (fn int => fn st => let val () = f int st in NONE end)); |
| 5828 | 423 |
|
|
60190
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
424 |
fun keep_proof f = |
|
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
425 |
keep (fn st => |
|
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
426 |
if is_proof st then f st |
|
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
427 |
else if is_skipped_proof st then () |
|
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
428 |
else warning "No proof state"); |
|
906de96ba68a
allow diagnostic proof commands with skip_proofs;
wenzelm
parents:
60189
diff
changeset
|
429 |
|
| 74672 | 430 |
val ignoredN = "<ignored>"; |
431 |
val malformedN = "<malformed>"; |
|
432 |
||
433 |
fun is_ignored tr = name_of tr = ignoredN; |
|
434 |
fun is_malformed tr = name_of tr = malformedN; |
|
| 48772 | 435 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
436 |
fun ignored pos = |
| 74672 | 437 |
empty |> name ignoredN |> position pos |> keep (fn _ => ()); |
|
73098
8a20737e4ebf
support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents:
72505
diff
changeset
|
438 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
439 |
fun malformed pos msg = |
| 74672 | 440 |
empty |> name malformedN |> position pos |> keep (fn _ => error msg); |
| 27840 | 441 |
|
| 21007 | 442 |
|
|
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
443 |
(* theory transitions *) |
|
44304
7ee000ce5390
maintain recent future proofs at transaction boundaries;
wenzelm
parents:
44270
diff
changeset
|
444 |
|
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
445 |
fun generic_theory f = transaction (fn _ => |
| 69886 | 446 |
(fn Theory gthy => node_presentation (Theory (f gthy)) |
| 26491 | 447 |
| _ => raise UNDEF)); |
448 |
||
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
449 |
fun theory' f = present_transaction (fn int => |
| 69878 | 450 |
(fn Theory (Context.Theory thy) => |
|
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
451 |
let val thy' = thy |
|
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
452 |
|> Sign.new_group |
|
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
453 |
|> f int |
|
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
454 |
|> Sign.reset_group; |
| 69886 | 455 |
in node_presentation (Theory (Context.Theory thy')) end |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
456 |
| _ => raise UNDEF)); |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
457 |
|
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
458 |
fun theory f = theory' (K f) no_presentation; |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
459 |
|
| 72434 | 460 |
fun begin_main_target begin f = transaction (fn _ => |
| 69878 | 461 |
(fn Theory (Context.Theory thy) => |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
462 |
let |
| 20985 | 463 |
val lthy = f thy; |
| 72453 | 464 |
val gthy = |
465 |
if begin |
|
466 |
then Context.Proof lthy |
|
| 72505 | 467 |
else Target_Context.end_named_cmd lthy; |
|
56897
c668735fb8b5
print results as "state", to avoid intrusion into the source text;
wenzelm
parents:
56895
diff
changeset
|
468 |
val _ = |
| 60245 | 469 |
(case Local_Theory.pretty lthy of |
470 |
[] => () |
|
471 |
| prts => Output.state (Pretty.string_of (Pretty.chunks prts))); |
|
| 69878 | 472 |
in (Theory gthy, lthy) end |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
473 |
| _ => raise UNDEF)); |
| 17076 | 474 |
|
| 72434 | 475 |
val end_main_target = transaction (fn _ => |
| 72505 | 476 |
(fn Theory (Context.Proof lthy) => (Theory (Target_Context.end_named_cmd lthy), lthy) |
| 21007 | 477 |
| _ => raise UNDEF)); |
478 |
||
| 72434 | 479 |
fun begin_nested_target f = transaction0 (fn _ => |
| 69878 | 480 |
(fn Theory gthy => |
|
72449
e1ee4a9902bd
centralized case distinction for beginning and ending nested targets in one place
haftmann
parents:
72436
diff
changeset
|
481 |
let |
| 72453 | 482 |
val lthy' = f gthy; |
|
72449
e1ee4a9902bd
centralized case distinction for beginning and ending nested targets in one place
haftmann
parents:
72436
diff
changeset
|
483 |
in Theory (Context.Proof lthy') end |
| 47069 | 484 |
| _ => raise UNDEF)); |
485 |
||
| 72434 | 486 |
val end_nested_target = transaction (fn _ => |
| 69878 | 487 |
(fn Theory (Context.Proof lthy) => |
| 72505 | 488 |
(case try Target_Context.end_nested_cmd lthy of |
489 |
SOME gthy' => (Theory gthy', lthy) |
|
| 47069 | 490 |
| NONE => raise UNDEF) |
491 |
| _ => raise UNDEF)); |
|
492 |
||
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
493 |
fun restricted_context (SOME (strict, scope)) = |
|
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
494 |
Proof_Context.map_naming (Name_Space.restricted strict scope) |
|
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
495 |
| restricted_context NONE = I; |
|
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59923
diff
changeset
|
496 |
|
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
497 |
fun local_theory' restricted target f = present_transaction (fn int => |
| 69878 | 498 |
(fn Theory gthy => |
| 21294 | 499 |
let |
| 72453 | 500 |
val (finish, lthy) = Target_Context.switch_named_cmd target gthy; |
| 47274 | 501 |
val lthy' = lthy |
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
502 |
|> restricted_context restricted |
|
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
503 |
|> Local_Theory.new_group |
|
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
504 |
|> f int |
|
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
505 |
|> Local_Theory.reset_group; |
| 69878 | 506 |
in (Theory (finish lthy'), lthy') end |
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
507 |
| _ => raise UNDEF)); |
| 15668 | 508 |
|
|
74964
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
509 |
fun local_theory restricted target f = |
|
77a96ed74340
allow general command transactions with presentation;
wenzelm
parents:
74833
diff
changeset
|
510 |
local_theory' restricted target (K f) no_presentation; |
| 21007 | 511 |
|
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
512 |
fun present_local_theory target g = present_transaction (fn _ => |
| 69878 | 513 |
(fn Theory gthy => |
| 72453 | 514 |
let val (finish, lthy) = Target_Context.switch_named_cmd target gthy; |
| 69878 | 515 |
in (Theory (finish lthy), lthy) end |
|
74831
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
516 |
| _ => raise UNDEF)) |
|
32490add64b4
tuned signature: more explicit types for presentation;
wenzelm
parents:
74672
diff
changeset
|
517 |
(presentation g); |
| 21007 | 518 |
|
519 |
||
520 |
(* proof transitions *) |
|
521 |
||
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
522 |
fun end_proof f = transaction (fn int => |
|
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
523 |
(fn Proof (prf, (finish, _)) => |
| 33390 | 524 |
let val state = Proof_Node.current prf in |
| 21007 | 525 |
if can (Proof.assert_bottom true) state then |
526 |
let |
|
527 |
val ctxt' = f int state; |
|
528 |
val gthy' = finish ctxt'; |
|
| 69878 | 529 |
in (Theory gthy', ctxt') end |
| 21007 | 530 |
else raise UNDEF |
531 |
end |
|
| 69886 | 532 |
| Skipped_Proof (0, (gthy, _)) => node_presentation (Theory gthy) |
| 21007 | 533 |
| _ => raise UNDEF)); |
534 |
||
| 21294 | 535 |
local |
536 |
||
| 69882 | 537 |
fun begin_proof init_proof = transaction0 (fn int => |
| 69878 | 538 |
(fn Theory gthy => |
| 21294 | 539 |
let |
| 69882 | 540 |
val (finish, prf) = init_proof int gthy; |
| 67157 | 541 |
val document = Options.default_string "document"; |
542 |
val skip = (document = "" orelse document = "false") andalso Goal.skip_proofs_enabled (); |
|
| 58795 | 543 |
val schematic_goal = try Proof.schematic_goal prf; |
| 47274 | 544 |
val _ = |
| 58795 | 545 |
if skip andalso schematic_goal = SOME true then |
| 47274 | 546 |
warning "Cannot skip proof of schematic goal statement" |
547 |
else (); |
|
| 21294 | 548 |
in |
| 58795 | 549 |
if skip andalso schematic_goal = SOME false then |
| 51555 | 550 |
Skipped_Proof (0, (finish (Proof.global_skip_proof true prf), gthy)) |
| 47274 | 551 |
else Proof (Proof_Node.init prf, (finish, gthy)) |
| 21294 | 552 |
end |
553 |
| _ => raise UNDEF)); |
|
554 |
||
555 |
in |
|
556 |
||
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
557 |
fun local_theory_to_proof' restricted target f = begin_proof |
| 47274 | 558 |
(fn int => fn gthy => |
|
59923
b21c82422d65
support private scope for individual local theory commands;
wenzelm
parents:
59472
diff
changeset
|
559 |
let |
| 72453 | 560 |
val (finish, lthy) = Target_Context.switch_named_cmd target gthy; |
|
59923
b21c82422d65
support private scope for individual local theory commands;
wenzelm
parents:
59472
diff
changeset
|
561 |
val prf = lthy |
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
562 |
|> restricted_context restricted |
|
59923
b21c82422d65
support private scope for individual local theory commands;
wenzelm
parents:
59472
diff
changeset
|
563 |
|> Local_Theory.new_group |
|
b21c82422d65
support private scope for individual local theory commands;
wenzelm
parents:
59472
diff
changeset
|
564 |
|> f int; |
|
b21c82422d65
support private scope for individual local theory commands;
wenzelm
parents:
59472
diff
changeset
|
565 |
in (finish o Local_Theory.reset_group, prf) end); |
|
24780
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
566 |
|
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
567 |
fun local_theory_to_proof restricted target f = |
|
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
568 |
local_theory_to_proof' restricted target (K f); |
| 21294 | 569 |
|
570 |
fun theory_to_proof f = begin_proof |
|
| 47274 | 571 |
(fn _ => fn gthy => |
|
56057
ad6bd8030d88
more explicit Sign.change_check -- detect structural mistakes where they emerge, not at later theory merges;
wenzelm
parents:
56055
diff
changeset
|
572 |
(Context.Theory o Sign.reset_group o Sign.change_check o Proof_Context.theory_of, |
|
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
573 |
(case gthy of |
| 52788 | 574 |
Context.Theory thy => f (Sign.new_group thy) |
|
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
575 |
| _ => raise UNDEF))); |
| 21294 | 576 |
|
577 |
end; |
|
578 |
||
| 69878 | 579 |
val forget_proof = transaction0 (fn _ => |
|
58798
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
wenzelm
parents:
58795
diff
changeset
|
580 |
(fn Proof (prf, (_, orig_gthy)) => |
|
68876
cefaac3d24ff
no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68875
diff
changeset
|
581 |
if Proof.is_notepad (Proof_Node.current prf) then raise UNDEF |
| 69878 | 582 |
else Theory orig_gthy |
583 |
| Skipped_Proof (_, (_, orig_gthy)) => Theory orig_gthy |
|
| 21007 | 584 |
| _ => raise UNDEF)); |
585 |
||
| 69878 | 586 |
fun proofs' f = transaction0 (fn int => |
|
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
587 |
(fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x) |
| 51555 | 588 |
| skip as Skipped_Proof _ => skip |
| 16815 | 589 |
| _ => raise UNDEF)); |
| 15668 | 590 |
|
|
49863
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
wenzelm
parents:
49062
diff
changeset
|
591 |
fun proof' f = proofs' ((Seq.single o Seq.Result) oo f); |
|
17904
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents:
17513
diff
changeset
|
592 |
val proofs = proofs' o K; |
| 6689 | 593 |
val proof = proof' o K; |
| 16815 | 594 |
|
|
60693
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
595 |
|
|
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
596 |
(* skipped proofs *) |
|
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
597 |
|
| 69878 | 598 |
fun actual_proof f = transaction0 (fn _ => |
| 21007 | 599 |
(fn Proof (prf, x) => Proof (f prf, x) |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
600 |
| _ => raise UNDEF)); |
| 16815 | 601 |
|
| 69878 | 602 |
fun skip_proof f = transaction0 (fn _ => |
|
60693
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
603 |
(fn skip as Skipped_Proof _ => (f (); skip) |
| 18563 | 604 |
| _ => raise UNDEF)); |
605 |
||
| 69878 | 606 |
val skip_proof_open = transaction0 (fn _ => |
|
60693
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
607 |
(fn Skipped_Proof (d, x) => Skipped_Proof (d + 1, x) |
|
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
608 |
| _ => raise UNDEF)); |
|
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
609 |
|
| 69878 | 610 |
val skip_proof_close = transaction0 (fn _ => |
611 |
(fn Skipped_Proof (0, (gthy, _)) => Theory gthy |
|
|
60693
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents:
60403
diff
changeset
|
612 |
| Skipped_Proof (d, x) => Skipped_Proof (d - 1, x) |
|
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
613 |
| _ => raise UNDEF)); |
| 5828 | 614 |
|
615 |
||
616 |
||
617 |
(** toplevel transactions **) |
|
618 |
||
|
52527
dbac84eab3bc
separate exec_id assignment for Command.print states, without affecting result of eval;
wenzelm
parents:
52499
diff
changeset
|
619 |
(* runtime position *) |
| 27427 | 620 |
|
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
621 |
fun exec_id id (tr as Transition {pos, ...}) =
|
|
73098
8a20737e4ebf
support more command positions, analogous to Command.core_range in Isabelle/Scala;
wenzelm
parents:
72505
diff
changeset
|
622 |
let val put_id = Position.put_id (Document_ID.print id) |
|
73106
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
wenzelm
parents:
73098
diff
changeset
|
623 |
in position (put_id pos) tr end; |
| 25799 | 624 |
|
| 25960 | 625 |
fun setmp_thread_position (Transition {pos, ...}) f x =
|
|
25819
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
wenzelm
parents:
25809
diff
changeset
|
626 |
Position.setmp_thread_data pos f x; |
| 25799 | 627 |
|
628 |
||
| 5828 | 629 |
(* apply transitions *) |
630 |
||
| 6664 | 631 |
local |
632 |
||
|
76073
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
633 |
fun show_state (st', opt_err) = |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
634 |
let |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
635 |
val enabled = is_none opt_err andalso Options.default_bool \<^system_option>\<open>show_states\<close>; |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
636 |
val opt_err' = |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
637 |
if enabled then |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
638 |
(case Exn.capture (Output.state o string_of_state) st' of |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
639 |
Exn.Exn exn => SOME exn |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
640 |
| Exn.Res _ => opt_err) |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
641 |
else opt_err; |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
642 |
in (st', opt_err') end; |
|
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
643 |
|
| 70134 | 644 |
fun app int (tr as Transition {name, markers, trans, ...}) =
|
| 67932 | 645 |
setmp_thread_position tr |
| 70134 | 646 |
(Timing.protocol (name_of tr) (pos_of tr) (apply_trans int name markers trans) |
| 76810 | 647 |
##> Option.map (fn Runtime.UNDEF => ERROR (type_error tr) | exn => exn) |
|
76073
951abf9db857
option "show_states" for more verbosity of batch-builds;
wenzelm
parents:
75077
diff
changeset
|
648 |
#> show_state); |
| 6664 | 649 |
|
650 |
in |
|
| 5828 | 651 |
|
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
652 |
fun transition int tr st = |
|
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
653 |
let |
|
60895
501be4aa75b4
default ML context for all command transactions, e.g. relevant for debugging and toplevel pretty-printing;
wenzelm
parents:
60693
diff
changeset
|
654 |
val (st', opt_err) = |
|
67642
10ff1f077119
more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents:
67641
diff
changeset
|
655 |
Context.setmp_generic_context (try (Context.Proof o presentation_context0) st) |
|
60895
501be4aa75b4
default ML context for all command transactions, e.g. relevant for debugging and toplevel pretty-printing;
wenzelm
parents:
60693
diff
changeset
|
656 |
(fn () => app int tr st) (); |
|
59055
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
657 |
val opt_err' = opt_err |> Option.map |
|
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
658 |
(fn Runtime.EXCURSION_FAIL exn_info => exn_info |
|
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
659 |
| exn => (Runtime.exn_context (try context_of st) exn, at_command tr)); |
|
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
660 |
in (st', opt_err') end; |
| 6664 | 661 |
|
662 |
end; |
|
| 5828 | 663 |
|
664 |
||
|
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
665 |
(* managed commands *) |
| 5828 | 666 |
|
| 51323 | 667 |
fun command_errors int tr st = |
668 |
(case transition int tr st of |
|
|
59055
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
669 |
(st', NONE) => ([], SOME st') |
| 65948 | 670 |
| (_, SOME (exn, _)) => (Runtime.exn_messages exn, NONE)); |
| 51323 | 671 |
|
|
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
672 |
fun command_exception int tr st = |
|
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
673 |
(case transition int tr st of |
|
59055
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
674 |
(st', NONE) => st' |
|
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
675 |
| (_, SOME (exn, info)) => |
| 62505 | 676 |
if Exn.is_interrupt exn then Exn.reraise exn |
|
59055
5a7157b8e870
more informative failure of protocol commands, with exception trace;
wenzelm
parents:
59032
diff
changeset
|
677 |
else raise Runtime.EXCURSION_FAIL (exn, info)); |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
678 |
|
| 58848 | 679 |
val command = command_exception false; |
|
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
680 |
|
|
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
681 |
|
| 56937 | 682 |
(* reset state *) |
683 |
||
684 |
local |
|
685 |
||
686 |
fun reset_state check trans st = |
|
687 |
if check st then NONE |
|
688 |
else #2 (command_errors false (trans empty) st); |
|
689 |
||
690 |
in |
|
691 |
||
|
68876
cefaac3d24ff
no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68875
diff
changeset
|
692 |
val reset_theory = reset_state is_theory forget_proof; |
| 56937 | 693 |
|
694 |
val reset_proof = |
|
695 |
reset_state is_proof |
|
| 69878 | 696 |
(transaction0 (fn _ => |
697 |
(fn Theory gthy => Skipped_Proof (0, (gthy, gthy)) |
|
| 56937 | 698 |
| _ => raise UNDEF))); |
699 |
||
|
68877
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
700 |
val reset_notepad = |
|
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
701 |
reset_state |
|
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
702 |
(fn st => |
|
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
703 |
(case try proof_of st of |
|
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
704 |
SOME state => not (Proof.is_notepad state) orelse can Proof.end_notepad state |
|
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
705 |
| NONE => true)) |
| 68878 | 706 |
(proof Proof.reset_notepad); |
|
68877
33d78e5e0a00
more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents:
68876
diff
changeset
|
707 |
|
| 56937 | 708 |
end; |
709 |
||
710 |
||
|
46959
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
wenzelm
parents:
45666
diff
changeset
|
711 |
(* scheduled proof result *) |
|
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
712 |
|
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
713 |
datatype result = |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
714 |
Result of transition * state | |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
715 |
Result_List of result list | |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
716 |
Result_Future of result future; |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
717 |
|
|
73687
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
718 |
fun join_results result = |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
719 |
let |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
720 |
fun add (tr, st') res = |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
721 |
(case res of |
| 76415 | 722 |
[] => [(make_state NONE, tr, st')] |
|
73687
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
723 |
| (_, _, st) :: _ => (st, tr, st') :: res); |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
724 |
fun acc (Result r) = add r |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
725 |
| acc (Result_List rs) = fold acc rs |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
726 |
| acc (Result_Future x) = acc (Future.join x); |
|
54fe8cc0e1c6
clarified signature: provide access to previous state;
wenzelm
parents:
73614
diff
changeset
|
727 |
in rev (acc result []) end; |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
728 |
|
| 51323 | 729 |
local |
730 |
||
| 47417 | 731 |
structure Result = Proof_Data |
|
28974
d6b190efa01a
excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents:
28645
diff
changeset
|
732 |
( |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
733 |
type T = result; |
| 59150 | 734 |
fun init _ = Result_List []; |
|
28974
d6b190efa01a
excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents:
28645
diff
changeset
|
735 |
); |
|
d6b190efa01a
excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents:
28645
diff
changeset
|
736 |
|
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
737 |
val get_result = Result.get o Proof.context_of; |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
738 |
val put_result = Proof.map_context o Result.put; |
| 51324 | 739 |
|
|
66169
8cfa8c7ee1f6
keep original bottom-up order of proof forks, which potentially reduces thread congestion due to Proofterm.consolidate;
wenzelm
parents:
65948
diff
changeset
|
740 |
fun timing_estimate elem = |
| 68839 | 741 |
let val trs = tl (Thy_Element.flat_element elem) |
| 69877 | 742 |
in fold (fn tr => fn t => timing_of tr + t) trs Time.zeroTime end; |
|
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
743 |
|
|
68130
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents:
67932
diff
changeset
|
744 |
fun future_proofs_enabled estimate st = |
| 51324 | 745 |
(case try proof_of st of |
746 |
NONE => false |
|
747 |
| SOME state => |
|
| 70398 | 748 |
not (Proofterm.proofs_enabled ()) andalso |
| 51324 | 749 |
not (Proof.is_relevant state) andalso |
750 |
(if can (Proof.assert_bottom true) state |
|
|
68130
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents:
67932
diff
changeset
|
751 |
then Future.proofs_enabled 1 |
|
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents:
67932
diff
changeset
|
752 |
else Future.proofs_enabled 2 orelse Future.proofs_enabled_timing estimate)); |
| 51278 | 753 |
|
| 69887 | 754 |
val empty_markers = markers []; |
755 |
val empty_trans = reset_trans #> keep (K ()); |
|
756 |
||
757 |
in |
|
758 |
||
759 |
fun fork_presentation tr = (tr |> empty_markers, tr |> empty_trans); |
|
760 |
||
| 58923 | 761 |
fun atom_result keywords tr st = |
| 51323 | 762 |
let |
763 |
val st' = |
|
|
68130
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents:
67932
diff
changeset
|
764 |
if Future.proofs_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then |
| 69887 | 765 |
let |
766 |
val (tr1, tr2) = fork_presentation tr; |
|
767 |
val _ = |
|
768 |
Execution.fork {name = "Toplevel.diag", pos = pos_of tr, pri = ~1}
|
|
769 |
(fn () => command tr1 st); |
|
770 |
in command tr2 st end |
|
| 51323 | 771 |
else command tr st; |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
772 |
in (Result (tr, st'), st') end; |
| 51323 | 773 |
|
| 68839 | 774 |
fun element_result keywords (Thy_Element.Element (tr, NONE)) st = atom_result keywords tr st |
775 |
| element_result keywords (elem as Thy_Element.Element (head_tr, SOME element_rest)) st = |
|
|
48633
7cd32f9d4293
recovered comination of Toplevel.skip_proofs and Goal.parallel_proofs from 9679bab23f93 (NB: skip_proofs leads to failure of Toplevel.proof_of);
wenzelm
parents:
47881
diff
changeset
|
776 |
let |
| 58923 | 777 |
val (head_result, st') = atom_result keywords head_tr st; |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
778 |
val (body_elems, end_tr) = element_rest; |
|
66169
8cfa8c7ee1f6
keep original bottom-up order of proof forks, which potentially reduces thread congestion due to Proofterm.consolidate;
wenzelm
parents:
65948
diff
changeset
|
779 |
val estimate = timing_estimate elem; |
| 51324 | 780 |
in |
|
68130
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents:
67932
diff
changeset
|
781 |
if not (future_proofs_enabled estimate st') |
| 51324 | 782 |
then |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
783 |
let |
| 68839 | 784 |
val proof_trs = maps Thy_Element.flat_element body_elems @ [end_tr]; |
| 58923 | 785 |
val (proof_results, st'') = fold_map (atom_result keywords) proof_trs st'; |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
786 |
in (Result_List (head_result :: proof_results), st'') end |
| 51324 | 787 |
else |
788 |
let |
|
| 69887 | 789 |
val (end_tr1, end_tr2) = fork_presentation end_tr; |
790 |
||
| 51324 | 791 |
val finish = Context.Theory o Proof_Context.theory_of; |
|
28974
d6b190efa01a
excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents:
28645
diff
changeset
|
792 |
|
|
51605
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
793 |
val future_proof = |
|
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
794 |
Proof.future_proof (fn state => |
| 53192 | 795 |
Execution.fork |
|
66169
8cfa8c7ee1f6
keep original bottom-up order of proof forks, which potentially reduces thread congestion due to Proofterm.consolidate;
wenzelm
parents:
65948
diff
changeset
|
796 |
{name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
|
|
51605
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
797 |
(fn () => |
|
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
798 |
let |
| 69883 | 799 |
val State ((Proof (prf, (_, orig_gthy)), _), prev_thy) = st'; |
| 69878 | 800 |
val node' = Proof (Proof_Node.apply (K state) prf, (finish, orig_gthy)); |
| 69887 | 801 |
val (results, result_state) = |
| 69886 | 802 |
State (node_presentation node', prev_thy) |
| 69887 | 803 |
|> fold_map (element_result keywords) body_elems ||> command end_tr1; |
804 |
in (Result_List results, presentation_context0 result_state) end)) |
|
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
805 |
#> (fn (res, state') => state' |> put_result (Result_Future res)); |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
806 |
|
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
807 |
val forked_proof = |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
808 |
proof (future_proof #> |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
809 |
(fn state => state |> Proof.local_done_proof |> put_result (get_result state))) o |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
810 |
end_proof (fn _ => future_proof #> |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
811 |
(fn state => state |> Proof.global_done_proof |> Result.put (get_result state))); |
|
28974
d6b190efa01a
excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents:
28645
diff
changeset
|
812 |
|
| 69887 | 813 |
val st'' = st' |> command (head_tr |> reset_trans |> forked_proof); |
814 |
val end_st = st'' |> command end_tr2; |
|
815 |
val end_result = Result (end_tr, end_st); |
|
| 51324 | 816 |
val result = |
|
67642
10ff1f077119
more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents:
67641
diff
changeset
|
817 |
Result_List [head_result, Result.get (presentation_context0 st''), end_result]; |
| 69887 | 818 |
in (result, end_st) end |
| 51324 | 819 |
end; |
|
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
820 |
|
| 6664 | 821 |
end; |
| 51323 | 822 |
|
823 |
end; |