author | wenzelm |
Wed, 22 Jan 2014 15:11:10 +0100 | |
changeset 55109 | ecff9e26360c |
parent 54678 | 87910da843d5 |
child 55446 | e77f2858bd59 |
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 |
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
11 |
val toplevel: 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 |
30398 | 17 |
val presentation_context_of: state -> Proof.context |
30801 | 18 |
val previous_context_of: state -> Proof.context 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 |
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
24 |
val end_theory: Position.T -> state -> theory |
16815 | 25 |
val print_state_context: state -> unit |
26 |
val print_state: bool -> state -> unit |
|
37858 | 27 |
val pretty_abstract: state -> Pretty.T |
32738 | 28 |
val quiet: bool Unsynchronized.ref |
29 |
val interact: bool Unsynchronized.ref |
|
30 |
val timing: bool Unsynchronized.ref |
|
31 |
val profiling: int Unsynchronized.ref |
|
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
32 |
val debug: bool Unsynchronized.ref |
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
33 |
val debugging: ('a -> 'b) -> 'a -> 'b |
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
34 |
val controlled_execution: ('a -> 'b) -> 'a -> 'b |
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
35 |
val program: (unit -> 'a) -> 'a |
33604
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
36 |
val thread: bool -> (unit -> unit) -> Thread.thread |
16682 | 37 |
type transition |
5828 | 38 |
val empty: transition |
38888
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
wenzelm
parents:
38876
diff
changeset
|
39 |
val print_of: transition -> bool |
27427 | 40 |
val name_of: transition -> string |
28105 | 41 |
val pos_of: transition -> Position.T |
5828 | 42 |
val name: string -> transition -> transition |
43 |
val position: Position.T -> transition -> transition |
|
44 |
val interactive: bool -> transition -> transition |
|
38888
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
wenzelm
parents:
38876
diff
changeset
|
45 |
val set_print: bool -> transition -> transition |
5828 | 46 |
val print: transition -> transition |
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
47 |
val init_theory: (unit -> theory) -> transition -> transition |
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
48 |
val is_init: transition -> bool |
44186
806f0ec1a43d
simplified Toplevel.init_theory: discontinued special master argument;
wenzelm
parents:
44185
diff
changeset
|
49 |
val modify_init: (unit -> theory) -> transition -> transition |
6689 | 50 |
val exit: transition -> transition |
5828 | 51 |
val keep: (state -> unit) -> transition -> transition |
7612 | 52 |
val keep': (bool -> state -> unit) -> transition -> transition |
5828 | 53 |
val imperative: (unit -> unit) -> transition -> transition |
27840 | 54 |
val ignored: Position.T -> transition |
51268
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents:
51241
diff
changeset
|
55 |
val is_ignored: transition -> bool |
27840 | 56 |
val malformed: Position.T -> string -> transition |
48772 | 57 |
val is_malformed: transition -> bool |
26491 | 58 |
val generic_theory: (generic_theory -> generic_theory) -> transition -> transition |
7612 | 59 |
val theory': (bool -> theory -> theory) -> transition -> transition |
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
60 |
val theory: (theory -> theory) -> transition -> transition |
20985 | 61 |
val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition |
21007 | 62 |
val end_local_theory: transition -> transition |
47069 | 63 |
val open_target: (generic_theory -> local_theory) -> transition -> transition |
64 |
val close_target: transition -> transition |
|
45488
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
65 |
val local_theory': (xstring * Position.T) option -> (bool -> local_theory -> local_theory) -> |
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
66 |
transition -> transition |
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
67 |
val local_theory: (xstring * Position.T) option -> (local_theory -> local_theory) -> |
29380 | 68 |
transition -> transition |
45488
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
69 |
val present_local_theory: (xstring * Position.T) option -> (state -> unit) -> |
24453 | 70 |
transition -> transition |
45488
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
71 |
val local_theory_to_proof': (xstring * Position.T) option -> |
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
72 |
(bool -> local_theory -> Proof.state) -> transition -> transition |
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
73 |
val local_theory_to_proof: (xstring * Position.T) option -> |
6d71d9e52369
pass positions for named targets, for formal links in the document model;
wenzelm
parents:
44304
diff
changeset
|
74 |
(local_theory -> Proof.state) -> transition -> transition |
17363 | 75 |
val theory_to_proof: (theory -> Proof.state) -> transition -> transition |
21007 | 76 |
val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition |
77 |
val forget_proof: transition -> transition |
|
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
78 |
val present_proof: (state -> unit) -> transition -> transition |
49863
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
wenzelm
parents:
49062
diff
changeset
|
79 |
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
|
80 |
val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition |
49863
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
wenzelm
parents:
49062
diff
changeset
|
81 |
val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition |
21177 | 82 |
val proof: (Proof.state -> Proof.state) -> transition -> transition |
33390 | 83 |
val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition |
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
84 |
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
|
85 |
val skip_proof_to_theory: (int -> bool) -> transition -> transition |
52536 | 86 |
val exec_id: Document_ID.exec -> transition -> transition |
9512 | 87 |
val unknown_theory: transition -> transition |
88 |
val unknown_proof: transition -> transition |
|
89 |
val unknown_context: transition -> transition |
|
28425 | 90 |
val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b |
28103
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
wenzelm
parents:
28095
diff
changeset
|
91 |
val add_hook: (transition -> state -> state -> unit) -> unit |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
92 |
val get_timing: transition -> Time.time option |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
93 |
val put_timing: Time.time option -> transition -> transition |
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
94 |
val transition: bool -> transition -> state -> (state * (exn * string) option) option |
51323 | 95 |
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
|
96 |
val command_exception: bool -> transition -> state -> state |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
97 |
type result |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
98 |
val join_results: result -> (transition * state) list |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
99 |
val element_result: transition Thy_Syntax.element -> state -> result * state |
5828 | 100 |
end; |
101 |
||
6965 | 102 |
structure Toplevel: TOPLEVEL = |
5828 | 103 |
struct |
104 |
||
105 |
(** toplevel state **) |
|
106 |
||
31476
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
31431
diff
changeset
|
107 |
exception UNDEF = Runtime.UNDEF; |
19063 | 108 |
|
109 |
||
21294 | 110 |
(* local theory wrappers *) |
5828 | 111 |
|
38350
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
haftmann
parents:
38236
diff
changeset
|
112 |
val loc_init = Named_Target.context_cmd; |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
113 |
val loc_exit = Local_Theory.assert_bottom true #> Local_Theory.exit_global; |
21294 | 114 |
|
47274 | 115 |
fun loc_begin loc (Context.Theory thy) = |
116 |
(Context.Theory o loc_exit, loc_init (the_default ("-", Position.none) loc) thy) |
|
117 |
| loc_begin NONE (Context.Proof lthy) = |
|
118 |
(Context.Proof o Local_Theory.restore, lthy) |
|
119 |
| loc_begin (SOME loc) (Context.Proof lthy) = |
|
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
120 |
(Context.Proof o Named_Target.reinit lthy, loc_init loc (loc_exit (Local_Theory.assert_nonbrittle lthy))); |
21294 | 121 |
|
122 |
||
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
123 |
(* datatype node *) |
21294 | 124 |
|
5828 | 125 |
datatype node = |
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
126 |
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
|
127 |
(*theory with presentation context*) | |
33390 | 128 |
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
|
129 |
(*proof node, finish, original theory*) | |
51555 | 130 |
Skipped_Proof of int * (generic_theory * generic_theory); |
27564
fc6d34e49e17
replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents:
27500
diff
changeset
|
131 |
(*proof depth, resulting theory, original theory*) |
5828 | 132 |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
133 |
val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE; |
18589 | 134 |
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; |
51555 | 135 |
val skipped_proof_node = fn Skipped_Proof _ => true | _ => false; |
18589 | 136 |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
137 |
fun cases_node f _ (Theory (gthy, _)) = f gthy |
33390 | 138 |
| cases_node _ g (Proof (prf, _)) = g (Proof_Node.current prf) |
51555 | 139 |
| cases_node f _ (Skipped_Proof (_, (gthy, _))) = f gthy; |
19063 | 140 |
|
29066 | 141 |
val context_node = cases_node Context.proof_of Proof.context_of; |
142 |
||
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
143 |
|
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
144 |
(* datatype state *) |
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
145 |
|
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
146 |
datatype state = State of node option * node option; (*current, previous*) |
5828 | 147 |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
148 |
val toplevel = State (NONE, NONE); |
5828 | 149 |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
150 |
fun is_toplevel (State (NONE, _)) = true |
7732 | 151 |
| is_toplevel _ = false; |
152 |
||
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
153 |
fun level (State (NONE, _)) = 0 |
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
154 |
| level (State (SOME (Theory _), _)) = 0 |
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
155 |
| level (State (SOME (Proof (prf, _)), _)) = Proof.level (Proof_Node.current prf) |
51555 | 156 |
| level (State (SOME (Skipped_Proof (d, _)), _)) = d + 1; (*different notion of proof depth!*) |
17076 | 157 |
|
52565 | 158 |
fun str_of_state (State (NONE, SOME (Theory (Context.Theory thy, _)))) = |
159 |
"at top level, result theory " ^ quote (Context.theory_name thy) |
|
160 |
| str_of_state (State (NONE, _)) = "at top level" |
|
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
161 |
| str_of_state (State (SOME (Theory (Context.Theory _, _)), _)) = "in theory mode" |
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
162 |
| str_of_state (State (SOME (Theory (Context.Proof _, _)), _)) = "in local theory mode" |
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
163 |
| str_of_state (State (SOME (Proof _), _)) = "in proof mode" |
51555 | 164 |
| str_of_state (State (SOME (Skipped_Proof _), _)) = "in skipped proof mode"; |
5946
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
165 |
|
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents:
5939
diff
changeset
|
166 |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
167 |
(* current node *) |
5828 | 168 |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
169 |
fun node_of (State (NONE, _)) = raise UNDEF |
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
170 |
| node_of (State (SOME node, _)) = node; |
5828 | 171 |
|
18589 | 172 |
fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state)); |
173 |
fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state)); |
|
51555 | 174 |
fun is_skipped_proof state = not (is_toplevel state) andalso skipped_proof_node (node_of state); |
18589 | 175 |
|
19063 | 176 |
fun node_case f g state = cases_node f g (node_of state); |
5828 | 177 |
|
30398 | 178 |
fun presentation_context_of state = |
179 |
(case try node_of state of |
|
180 |
SOME (Theory (_, SOME ctxt)) => ctxt |
|
181 |
| SOME node => context_node node |
|
182 |
| NONE => raise UNDEF); |
|
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
183 |
|
30801 | 184 |
fun previous_context_of (State (_, NONE)) = NONE |
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
185 |
| previous_context_of (State (_, SOME prev)) = SOME (context_node prev); |
30801 | 186 |
|
21506 | 187 |
val context_of = node_case Context.proof_of Proof.context_of; |
22089 | 188 |
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
|
189 |
val theory_of = node_case Context.theory_of Proof.theory_of; |
18589 | 190 |
val proof_of = node_case (fn _ => raise UNDEF) I; |
17208 | 191 |
|
18589 | 192 |
fun proof_position_of state = |
193 |
(case node_of state of |
|
33390 | 194 |
Proof (prf, _) => Proof_Node.position prf |
18589 | 195 |
| _ => raise UNDEF); |
6664 | 196 |
|
43667
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
197 |
fun end_theory _ (State (NONE, SOME (Theory (Context.Theory thy, _)))) = thy |
48992 | 198 |
| end_theory pos (State (NONE, _)) = error ("Bad theory" ^ Position.here pos) |
199 |
| end_theory pos (State (SOME _, _)) = error ("Unfinished theory" ^ Position.here pos); |
|
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
200 |
|
5828 | 201 |
|
16815 | 202 |
(* print state *) |
203 |
||
38388 | 204 |
val pretty_context = Local_Theory.pretty o Context.cases (Named_Target.theory_init) I; |
16815 | 205 |
|
23640
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
wenzelm
parents:
23619
diff
changeset
|
206 |
fun print_state_context state = |
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
207 |
(case try node_of state of |
21506 | 208 |
NONE => [] |
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
209 |
| SOME (Theory (gthy, _)) => pretty_context gthy |
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
210 |
| SOME (Proof (_, (_, gthy))) => pretty_context gthy |
51555 | 211 |
| SOME (Skipped_Proof (_, (gthy, _))) => pretty_context gthy) |
23640
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
wenzelm
parents:
23619
diff
changeset
|
212 |
|> Pretty.chunks |> Pretty.writeln; |
16815 | 213 |
|
23640
baec2e674461
toplevel prompt/print_state: proper markup, removed hooks;
wenzelm
parents:
23619
diff
changeset
|
214 |
fun print_state prf_only state = |
23701 | 215 |
(case try node_of state of |
216 |
NONE => [] |
|
217 |
| SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy |
|
218 |
| SOME (Proof (prf, _)) => |
|
33390 | 219 |
Proof.pretty_state (Proof_Node.position prf) (Proof_Node.current prf) |
51555 | 220 |
| SOME (Skipped_Proof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]) |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49863
diff
changeset
|
221 |
|> Pretty.markup_chunks Markup.state |> Pretty.writeln; |
16815 | 222 |
|
37858 | 223 |
fun pretty_abstract state = Pretty.str ("<Isar " ^ str_of_state state ^ ">"); |
224 |
||
16815 | 225 |
|
15668 | 226 |
|
5828 | 227 |
(** toplevel transitions **) |
228 |
||
53403
c09f4005d6bd
some explicit indication of Proof General legacy;
wenzelm
parents:
53192
diff
changeset
|
229 |
val quiet = Unsynchronized.ref false; (*Proof General legacy*) |
c09f4005d6bd
some explicit indication of Proof General legacy;
wenzelm
parents:
53192
diff
changeset
|
230 |
val interact = Unsynchronized.ref false; (*Proof General legacy*) |
c09f4005d6bd
some explicit indication of Proof General legacy;
wenzelm
parents:
53192
diff
changeset
|
231 |
val timing = Unsynchronized.ref false; (*Proof General legacy*) |
32738 | 232 |
val profiling = Unsynchronized.ref 0; |
16682 | 233 |
|
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
234 |
|
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
235 |
(* controlled execution *) |
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
236 |
|
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
237 |
val debug = Runtime.debug; |
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
238 |
fun debugging f = Runtime.debugging ML_Compiler.exn_message f; |
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
239 |
fun controlled_execution f = Runtime.controlled_execution ML_Compiler.exn_message f; |
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
240 |
|
33604
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
241 |
fun program body = |
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
242 |
(body |
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
243 |
|> controlled_execution |
54387 | 244 |
|> Runtime.toplevel_error ML_Compiler.exn_error_message) (); |
33604
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
245 |
|
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
246 |
fun thread interrupts body = |
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
247 |
Thread.fork |
39232
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
wenzelm
parents:
38888
diff
changeset
|
248 |
(((fn () => body () handle exn => if Exn.is_interrupt exn then () else reraise exn) |
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
249 |
|> debugging |
33604
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
wenzelm
parents:
33553
diff
changeset
|
250 |
|> Runtime.toplevel_error |
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39513
diff
changeset
|
251 |
(fn exn => |
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39513
diff
changeset
|
252 |
Output.urgent_message ("## INTERNAL ERROR ##\n" ^ ML_Compiler.exn_message exn))), |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37208
diff
changeset
|
253 |
Simple_Thread.attributes interrupts); |
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
254 |
|
5828 | 255 |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
256 |
(* node transactions -- maintaining stable checkpoints *) |
7022 | 257 |
|
31476
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
31431
diff
changeset
|
258 |
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
|
259 |
|
6689 | 260 |
local |
261 |
||
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
262 |
fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE) |
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
263 |
| reset_presentation node = node; |
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
264 |
|
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
265 |
fun map_theory f (Theory (gthy, ctxt)) = |
33671 | 266 |
Theory (Context.mapping f (Local_Theory.raw_theory f) gthy, ctxt) |
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
267 |
| map_theory _ node = node; |
6689 | 268 |
|
269 |
in |
|
270 |
||
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
271 |
fun apply_transaction f g node = |
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
272 |
let |
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
273 |
val cont_node = reset_presentation node; |
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
274 |
fun state_error e nd = (State (SOME nd, SOME node), e); |
6689 | 275 |
|
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
276 |
val (result, err) = |
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
277 |
cont_node |
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
278 |
|> controlled_execution f |
26624
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
279 |
|> state_error NONE |
770265032999
transaction/init: ensure stable theory (non-draft);
wenzelm
parents:
26621
diff
changeset
|
280 |
handle exn => state_error (SOME exn) cont_node; |
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
281 |
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
|
282 |
(case err of |
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
|
283 |
NONE => tap g result |
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
|
284 |
| SOME exn => raise FAILURE (result, exn)) |
20128
8f0e07d7cf92
keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents:
19996
diff
changeset
|
285 |
end; |
6689 | 286 |
|
43667
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
287 |
val exit_transaction = |
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
288 |
apply_transaction |
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
289 |
(fn Theory (Context.Theory thy, _) => Theory (Context.Theory (Theory.end_theory thy), NONE) |
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
290 |
| node => node) (K ()) |
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
291 |
#> (fn State (node', _) => State (NONE, node')); |
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
292 |
|
6689 | 293 |
end; |
294 |
||
295 |
||
296 |
(* primitive transitions *) |
|
297 |
||
5828 | 298 |
datatype trans = |
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
299 |
Init of unit -> theory | (*init theory*) |
37953
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
300 |
Exit | (*formal exit of theory*) |
ddc3b72f9a42
simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents:
37951
diff
changeset
|
301 |
Keep of bool -> state -> unit | (*peek at state*) |
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
302 |
Transaction of (bool -> node -> node) * (state -> unit); (*node transaction and presentation*) |
21958
9dfd1ca4c0a0
refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents:
21861
diff
changeset
|
303 |
|
6689 | 304 |
local |
5828 | 305 |
|
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
306 |
fun apply_tr _ (Init f) (State (NONE, _)) = |
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
307 |
State (SOME (Theory (Context.Theory (controlled_execution f ()), NONE)), NONE) |
43667
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
308 |
| apply_tr _ Exit (State (SOME (state as Theory (Context.Theory _, _)), _)) = |
6d73cceb1503
explicit exit_transaction with Theory.end_theory (which could include sanity checks as in HOL-SPARK for example);
wenzelm
parents:
43665
diff
changeset
|
309 |
exit_transaction state |
32792 | 310 |
| apply_tr int (Keep f) state = |
53709
84522727f9d3
improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents:
53403
diff
changeset
|
311 |
controlled_execution (fn x => tap (f int) x) state |
32792 | 312 |
| apply_tr int (Transaction (f, g)) (State (SOME state, _)) = |
313 |
apply_transaction (fn x => f int x) g state |
|
314 |
| apply_tr _ _ _ = raise UNDEF; |
|
5828 | 315 |
|
32792 | 316 |
fun apply_union _ [] state = raise FAILURE (state, UNDEF) |
317 |
| apply_union int (tr :: trs) state = |
|
318 |
apply_union int trs state |
|
319 |
handle Runtime.UNDEF => apply_tr int tr state |
|
320 |
| FAILURE (alt_state, UNDEF) => apply_tr int tr alt_state |
|
6689 | 321 |
| exn as FAILURE _ => raise exn |
322 |
| exn => raise FAILURE (state, exn); |
|
323 |
||
324 |
in |
|
325 |
||
32792 | 326 |
fun apply_trans int trs state = (apply_union int trs state, NONE) |
15531 | 327 |
handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); |
6689 | 328 |
|
329 |
end; |
|
5828 | 330 |
|
331 |
||
332 |
(* datatype transition *) |
|
333 |
||
334 |
datatype transition = Transition of |
|
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
335 |
{name: string, (*command name*) |
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
336 |
pos: Position.T, (*source position*) |
53403
c09f4005d6bd
some explicit indication of Proof General legacy;
wenzelm
parents:
53192
diff
changeset
|
337 |
int_only: bool, (*interactive-only*) (* TTY / Proof General legacy*) |
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
338 |
print: bool, (*print result state*) |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
339 |
timing: Time.time option, (*prescient timing information*) |
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
340 |
trans: trans list}; (*primitive transitions (union)*) |
5828 | 341 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
342 |
fun make_transition (name, pos, int_only, print, timing, trans) = |
51217
65ab2c4f4c32
support for prescient timing information within command transactions;
wenzelm
parents:
51216
diff
changeset
|
343 |
Transition {name = name, pos = pos, int_only = int_only, print = print, |
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
344 |
timing = timing, trans = trans}; |
5828 | 345 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
346 |
fun map_transition f (Transition {name, pos, int_only, print, timing, trans}) = |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
347 |
make_transition (f (name, pos, int_only, print, timing, trans)); |
5828 | 348 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
349 |
val empty = make_transition ("", Position.none, false, false, NONE, []); |
5828 | 350 |
|
351 |
||
352 |
(* diagnostics *) |
|
353 |
||
38888
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
wenzelm
parents:
38876
diff
changeset
|
354 |
fun print_of (Transition {print, ...}) = print; |
27427 | 355 |
fun name_of (Transition {name, ...}) = name; |
28105 | 356 |
fun pos_of (Transition {pos, ...}) = pos; |
5828 | 357 |
|
48993 | 358 |
fun command_msg msg tr = msg ^ "command " ^ quote (name_of tr) ^ Position.here (pos_of tr); |
38875
c7a66b584147
tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents:
38721
diff
changeset
|
359 |
fun at_command tr = command_msg "At " tr; |
5828 | 360 |
|
361 |
fun type_error tr state = |
|
18685 | 362 |
ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state); |
5828 | 363 |
|
364 |
||
365 |
(* modify transitions *) |
|
366 |
||
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
367 |
fun name name = map_transition (fn (_, pos, int_only, print, timing, trans) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
368 |
(name, pos, int_only, print, timing, trans)); |
9010 | 369 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
370 |
fun position pos = map_transition (fn (name, _, int_only, print, timing, trans) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
371 |
(name, pos, int_only, print, timing, trans)); |
5828 | 372 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
373 |
fun interactive int_only = map_transition (fn (name, pos, _, print, timing, trans) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
374 |
(name, pos, int_only, print, timing, trans)); |
14923 | 375 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
376 |
fun add_trans tr = map_transition (fn (name, pos, int_only, print, timing, trans) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
377 |
(name, pos, int_only, print, timing, tr :: trans)); |
16607 | 378 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
379 |
val reset_trans = map_transition (fn (name, pos, int_only, print, timing, _) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
380 |
(name, pos, int_only, print, timing, [])); |
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
381 |
|
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
382 |
fun set_print print = map_transition (fn (name, pos, int_only, _, timing, trans) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
383 |
(name, pos, int_only, print, timing, trans)); |
28453
06702e7acd1d
excursion/unit_result: no print for forked end, finish into global theory, pick resul from presentation context;
wenzelm
parents:
28451
diff
changeset
|
384 |
|
06702e7acd1d
excursion/unit_result: no print for forked end, finish into global theory, pick resul from presentation context;
wenzelm
parents:
28451
diff
changeset
|
385 |
val print = set_print true; |
5828 | 386 |
|
387 |
||
21007 | 388 |
(* basic transitions *) |
5828 | 389 |
|
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
390 |
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
|
391 |
|
44187
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
392 |
fun is_init (Transition {trans = [Init _], ...}) = true |
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
393 |
| is_init _ = false; |
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
394 |
|
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents:
44186
diff
changeset
|
395 |
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
|
396 |
|
6689 | 397 |
val exit = add_trans Exit; |
7612 | 398 |
val keep' = add_trans o Keep; |
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
399 |
|
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
400 |
fun present_transaction f g = add_trans (Transaction (f, g)); |
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
401 |
fun transaction f = present_transaction f (K ()); |
5828 | 402 |
|
7612 | 403 |
fun keep f = add_trans (Keep (fn _ => f)); |
5828 | 404 |
fun imperative f = keep (fn _ => f ()); |
405 |
||
27840 | 406 |
fun ignored pos = empty |> name "<ignored>" |> position pos |> imperative I; |
51268
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents:
51241
diff
changeset
|
407 |
fun is_ignored tr = name_of tr = "<ignored>"; |
48772 | 408 |
|
409 |
val malformed_name = "<malformed>"; |
|
27840 | 410 |
fun malformed pos msg = |
48772 | 411 |
empty |> name malformed_name |> position pos |> imperative (fn () => error msg); |
412 |
fun is_malformed tr = name_of tr = malformed_name; |
|
27840 | 413 |
|
21007 | 414 |
val unknown_theory = imperative (fn () => warning "Unknown theory context"); |
415 |
val unknown_proof = imperative (fn () => warning "Unknown proof context"); |
|
416 |
val unknown_context = imperative (fn () => warning "Unknown context"); |
|
15668 | 417 |
|
21007 | 418 |
|
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
419 |
(* theory transitions *) |
44304
7ee000ce5390
maintain recent future proofs at transaction boundaries;
wenzelm
parents:
44270
diff
changeset
|
420 |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
421 |
fun generic_theory f = transaction (fn _ => |
26491 | 422 |
(fn Theory (gthy, _) => Theory (f gthy, NONE) |
423 |
| _ => raise UNDEF)); |
|
424 |
||
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
425 |
fun theory' f = transaction (fn int => |
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
426 |
(fn Theory (Context.Theory thy, _) => |
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
427 |
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
|
428 |
|> Sign.new_group |
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
429 |
|> f int |
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
430 |
|> Sign.reset_group; |
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
431 |
in Theory (Context.Theory thy', NONE) end |
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
432 |
| _ => raise UNDEF)); |
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
433 |
|
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
434 |
fun theory f = theory' (K f); |
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
435 |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
436 |
fun begin_local_theory begin f = transaction (fn _ => |
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
437 |
(fn Theory (Context.Theory thy, _) => |
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
438 |
let |
20985 | 439 |
val lthy = f thy; |
21294 | 440 |
val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy); |
441 |
in Theory (gthy, SOME lthy) end |
|
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
442 |
| _ => raise UNDEF)); |
17076 | 443 |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
444 |
val end_local_theory = transaction (fn _ => |
21294 | 445 |
(fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy) |
21007 | 446 |
| _ => raise UNDEF)); |
447 |
||
47069 | 448 |
fun open_target f = transaction (fn _ => |
449 |
(fn Theory (gthy, _) => |
|
450 |
let val lthy = f gthy |
|
451 |
in Theory (Context.Proof lthy, SOME lthy) end |
|
452 |
| _ => raise UNDEF)); |
|
453 |
||
454 |
val close_target = transaction (fn _ => |
|
455 |
(fn Theory (Context.Proof lthy, _) => |
|
456 |
(case try Local_Theory.close_target lthy of |
|
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
457 |
SOME ctxt' => |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
458 |
let |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
459 |
val gthy' = |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
460 |
if can Local_Theory.assert ctxt' |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
461 |
then Context.Proof ctxt' |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
462 |
else Context.Theory (Proof_Context.theory_of ctxt'); |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
50201
diff
changeset
|
463 |
in Theory (gthy', SOME lthy) end |
47069 | 464 |
| NONE => raise UNDEF) |
465 |
| _ => raise UNDEF)); |
|
466 |
||
467 |
||
21007 | 468 |
local |
469 |
||
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
470 |
fun local_theory_presentation loc f = present_transaction (fn int => |
21294 | 471 |
(fn Theory (gthy, _) => |
472 |
let |
|
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
473 |
val (finish, lthy) = loc_begin loc gthy; |
47274 | 474 |
val lthy' = lthy |
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
475 |
|> Local_Theory.new_group |
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
476 |
|> f int |
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
477 |
|> Local_Theory.reset_group; |
21294 | 478 |
in Theory (finish lthy', SOME lthy') end |
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
479 |
| _ => raise UNDEF)); |
15668 | 480 |
|
21007 | 481 |
in |
482 |
||
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
483 |
fun local_theory' loc f = local_theory_presentation loc f (K ()); |
29380 | 484 |
fun local_theory loc f = local_theory' loc (K f); |
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
485 |
fun present_local_theory loc = local_theory_presentation loc (K I); |
18955 | 486 |
|
21007 | 487 |
end; |
488 |
||
489 |
||
490 |
(* proof transitions *) |
|
491 |
||
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
492 |
fun end_proof f = transaction (fn int => |
24795
6f5cb7885fd7
print_state_context: local theory context, not proof context;
wenzelm
parents:
24780
diff
changeset
|
493 |
(fn Proof (prf, (finish, _)) => |
33390 | 494 |
let val state = Proof_Node.current prf in |
21007 | 495 |
if can (Proof.assert_bottom true) state then |
496 |
let |
|
497 |
val ctxt' = f int state; |
|
498 |
val gthy' = finish ctxt'; |
|
499 |
in Theory (gthy', SOME ctxt') end |
|
500 |
else raise UNDEF |
|
501 |
end |
|
51555 | 502 |
| Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE) |
21007 | 503 |
| _ => raise UNDEF)); |
504 |
||
21294 | 505 |
local |
506 |
||
47274 | 507 |
fun begin_proof init = transaction (fn int => |
21294 | 508 |
(fn Theory (gthy, _) => |
509 |
let |
|
47274 | 510 |
val (finish, prf) = init int gthy; |
52499 | 511 |
val skip = Goal.skip_proofs_enabled (); |
40960 | 512 |
val (is_goal, no_skip) = |
513 |
(true, Proof.schematic_goal prf) handle ERROR _ => (false, true); |
|
47274 | 514 |
val _ = |
515 |
if is_goal andalso skip andalso no_skip then |
|
516 |
warning "Cannot skip proof of schematic goal statement" |
|
517 |
else (); |
|
21294 | 518 |
in |
40960 | 519 |
if skip andalso not no_skip then |
51555 | 520 |
Skipped_Proof (0, (finish (Proof.global_skip_proof true prf), gthy)) |
47274 | 521 |
else Proof (Proof_Node.init prf, (finish, gthy)) |
21294 | 522 |
end |
523 |
| _ => raise UNDEF)); |
|
524 |
||
525 |
in |
|
526 |
||
24780
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
527 |
fun local_theory_to_proof' loc f = begin_proof |
47274 | 528 |
(fn int => fn gthy => |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
529 |
let val (finish, lthy) = loc_begin loc gthy |
49012
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents:
49011
diff
changeset
|
530 |
in (finish o Local_Theory.reset_group, f int (Local_Theory.new_group lthy)) end); |
24780
47bb1e380d83
local_theory transactions: more careful treatment of context position;
wenzelm
parents:
24634
diff
changeset
|
531 |
|
24453 | 532 |
fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f); |
21294 | 533 |
|
534 |
fun theory_to_proof f = begin_proof |
|
47274 | 535 |
(fn _ => fn gthy => |
52788 | 536 |
(Context.Theory o Sign.reset_group o Proof_Context.theory_of, |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
537 |
(case gthy of |
52788 | 538 |
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
|
539 |
| _ => raise UNDEF))); |
21294 | 540 |
|
541 |
end; |
|
542 |
||
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
543 |
val forget_proof = transaction (fn _ => |
21007 | 544 |
(fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) |
51555 | 545 |
| Skipped_Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE) |
21007 | 546 |
| _ => raise UNDEF)); |
547 |
||
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
548 |
val present_proof = present_transaction (fn _ => |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
549 |
(fn Proof (prf, x) => Proof (Proof_Node.apply I prf, x) |
51555 | 550 |
| skip as Skipped_Proof _ => skip |
30366
e3d788b9dffb
simplified presentation: built into transaction, pass state directly;
wenzelm
parents:
29516
diff
changeset
|
551 |
| _ => raise UNDEF)); |
21177 | 552 |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
553 |
fun proofs' f = transaction (fn int => |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49042
diff
changeset
|
554 |
(fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x) |
51555 | 555 |
| skip as Skipped_Proof _ => skip |
16815 | 556 |
| _ => raise UNDEF)); |
15668 | 557 |
|
49863
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
wenzelm
parents:
49062
diff
changeset
|
558 |
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
|
559 |
val proofs = proofs' o K; |
6689 | 560 |
val proof = proof' o K; |
16815 | 561 |
|
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
562 |
fun actual_proof f = transaction (fn _ => |
21007 | 563 |
(fn Proof (prf, x) => Proof (f prf, x) |
20963
a7fd8f05a2be
added type global_theory -- theory or local_theory;
wenzelm
parents:
20928
diff
changeset
|
564 |
| _ => raise UNDEF)); |
16815 | 565 |
|
53403
c09f4005d6bd
some explicit indication of Proof General legacy;
wenzelm
parents:
53192
diff
changeset
|
566 |
(*Proof General legacy*) |
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
567 |
fun skip_proof f = transaction (fn _ => |
51555 | 568 |
(fn Skipped_Proof (h, x) => Skipped_Proof (f h, x) |
18563 | 569 |
| _ => raise UNDEF)); |
570 |
||
53403
c09f4005d6bd
some explicit indication of Proof General legacy;
wenzelm
parents:
53192
diff
changeset
|
571 |
(*Proof General legacy*) |
27601
6683cdb94af8
simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents:
27583
diff
changeset
|
572 |
fun skip_proof_to_theory pred = transaction (fn _ => |
51555 | 573 |
(fn Skipped_Proof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF |
33725
a8481da77270
implicit name space grouping for theory/local_theory transactions;
wenzelm
parents:
33671
diff
changeset
|
574 |
| _ => raise UNDEF)); |
5828 | 575 |
|
576 |
||
577 |
||
578 |
(** toplevel transactions **) |
|
579 |
||
52527
dbac84eab3bc
separate exec_id assignment for Command.print states, without affecting result of eval;
wenzelm
parents:
52499
diff
changeset
|
580 |
(* runtime position *) |
27427 | 581 |
|
52536 | 582 |
fun exec_id id (tr as Transition {pos, ...}) = |
583 |
position (Position.put_id (Document_ID.print id) pos) tr; |
|
25799 | 584 |
|
25960 | 585 |
fun setmp_thread_position (Transition {pos, ...}) f x = |
25819
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
wenzelm
parents:
25809
diff
changeset
|
586 |
Position.setmp_thread_data pos f x; |
25799 | 587 |
|
588 |
||
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
589 |
(* post-transition hooks *) |
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
590 |
|
37905 | 591 |
local |
592 |
val hooks = Unsynchronized.ref ([]: (transition -> state -> state -> unit) list); |
|
593 |
in |
|
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
594 |
|
32738 | 595 |
fun add_hook f = CRITICAL (fn () => Unsynchronized.change hooks (cons f)); |
33223 | 596 |
fun get_hooks () = ! hooks; |
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
597 |
|
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
598 |
end; |
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
599 |
|
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
600 |
|
5828 | 601 |
(* apply transitions *) |
602 |
||
51217
65ab2c4f4c32
support for prescient timing information within command transactions;
wenzelm
parents:
51216
diff
changeset
|
603 |
fun get_timing (Transition {timing, ...}) = timing; |
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
604 |
fun put_timing timing = map_transition (fn (name, pos, int_only, print, _, trans) => |
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
605 |
(name, pos, int_only, print, timing, trans)); |
51217
65ab2c4f4c32
support for prescient timing information within command transactions;
wenzelm
parents:
51216
diff
changeset
|
606 |
|
6664 | 607 |
local |
608 |
||
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
609 |
fun app int (tr as Transition {trans, print, ...}) = |
25819
e6feb08b7f4b
replaced thread_properties by simplified version in position.ML;
wenzelm
parents:
25809
diff
changeset
|
610 |
setmp_thread_position tr (fn state => |
25799 | 611 |
let |
51595 | 612 |
val timing_start = Timing.start (); |
25799 | 613 |
|
51595 | 614 |
val (result, opt_err) = |
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51605
diff
changeset
|
615 |
state |> (apply_trans int trans |> ! profiling > 0 ? profile (! profiling)); |
26621
78b3ad7af5d5
eliminated unused name_of, source, source_of, print', print3, three_buffersN;
wenzelm
parents:
26602
diff
changeset
|
616 |
val _ = if int andalso not (! quiet) andalso print then print_state false result else (); |
51216 | 617 |
|
51595 | 618 |
val timing_result = Timing.result timing_start; |
51662
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
wenzelm
parents:
51661
diff
changeset
|
619 |
val timing_props = |
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
wenzelm
parents:
51661
diff
changeset
|
620 |
Markup.command_timing :: (Markup.nameN, name_of tr) :: Position.properties_of (pos_of tr); |
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
wenzelm
parents:
51661
diff
changeset
|
621 |
val _ = Timing.protocol_message timing_props timing_result; |
51595 | 622 |
in |
623 |
(result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_err) |
|
624 |
end); |
|
6664 | 625 |
|
626 |
in |
|
5828 | 627 |
|
26602
5534b6a6b810
made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents:
26491
diff
changeset
|
628 |
fun transition int tr st = |
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
629 |
let |
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
630 |
val hooks = get_hooks (); |
28103
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
wenzelm
parents:
28095
diff
changeset
|
631 |
fun apply_hooks st' = hooks |> List.app (fn f => (try (fn () => f tr st st') (); ())); |
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
632 |
|
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
633 |
val ctxt = try context_of st; |
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
634 |
val res = |
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
635 |
(case app int tr st of |
38888
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
wenzelm
parents:
38876
diff
changeset
|
636 |
(_, SOME Runtime.TERMINATE) => NONE |
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
wenzelm
parents:
38876
diff
changeset
|
637 |
| (st', SOME (Runtime.EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info) |
31476
c5d2899b6de9
moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents:
31431
diff
changeset
|
638 |
| (st', SOME exn) => SOME (st', SOME (Runtime.exn_context ctxt exn, at_command tr)) |
28103
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
wenzelm
parents:
28095
diff
changeset
|
639 |
| (st', NONE) => SOME (st', NONE)); |
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
wenzelm
parents:
28095
diff
changeset
|
640 |
val _ = (case res of SOME (st', NONE) => apply_hooks st' | _ => ()); |
28095
7eaf0813bdc3
added add_hook interface for post-transition hooks;
wenzelm
parents:
27859
diff
changeset
|
641 |
in res end; |
6664 | 642 |
|
643 |
end; |
|
5828 | 644 |
|
645 |
||
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
646 |
(* managed commands *) |
5828 | 647 |
|
51323 | 648 |
fun command_errors int tr st = |
649 |
(case transition int tr st of |
|
650 |
SOME (st', NONE) => ([], SOME st') |
|
651 |
| SOME (_, SOME (exn, _)) => (ML_Compiler.exn_messages_ids exn, NONE) |
|
652 |
| NONE => (ML_Compiler.exn_messages_ids Runtime.TERMINATE, NONE)); |
|
653 |
||
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
654 |
fun command_exception int tr st = |
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
655 |
(case transition int tr st of |
28425 | 656 |
SOME (st', NONE) => st' |
39285 | 657 |
| SOME (_, SOME (exn, info)) => |
658 |
if Exn.is_interrupt exn then reraise exn else raise Runtime.EXCURSION_FAIL (exn, info) |
|
38888
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
wenzelm
parents:
38876
diff
changeset
|
659 |
| NONE => raise Runtime.EXCURSION_FAIL (Runtime.TERMINATE, at_command tr)); |
27576
7afff36043e6
eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents:
27564
diff
changeset
|
660 |
|
51323 | 661 |
fun command tr = command_exception (! interact) tr; |
51284
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents:
51282
diff
changeset
|
662 |
|
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
663 |
|
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
|
664 |
(* scheduled proof result *) |
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
665 |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
666 |
datatype result = |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
667 |
Result of transition * state | |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
668 |
Result_List of result list | |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
669 |
Result_Future of result future; |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
670 |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
671 |
fun join_results (Result x) = [x] |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
672 |
| join_results (Result_List xs) = maps join_results xs |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
673 |
| join_results (Result_Future x) = join_results (Future.join x); |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
674 |
|
51323 | 675 |
local |
676 |
||
47417 | 677 |
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
|
678 |
( |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
679 |
type T = result; |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
680 |
val empty: T = Result_List []; |
47417 | 681 |
fun init _ = empty; |
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
|
682 |
); |
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
|
683 |
|
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
684 |
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
|
685 |
val put_result = Proof.map_context o Result.put; |
51324 | 686 |
|
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
687 |
fun timing_estimate include_head elem = |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
688 |
let |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
689 |
val trs = Thy_Syntax.flat_element elem |> not include_head ? tl; |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
690 |
val timings = map get_timing trs; |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
691 |
in |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
692 |
if forall is_some timings then |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
693 |
SOME (fold (curry Time.+ o the) timings Time.zeroTime) |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
694 |
else NONE |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
695 |
end; |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
696 |
|
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
697 |
fun priority NONE = ~1 |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
698 |
| priority (SOME estimate) = |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
699 |
Int.min (Real.floor (Real.max (Math.log10 (Time.toReal estimate), ~3.0)) - 3, ~1); |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
700 |
|
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
701 |
fun proof_future_enabled estimate st = |
51324 | 702 |
(case try proof_of st of |
703 |
NONE => false |
|
704 |
| SOME state => |
|
705 |
not (Proof.is_relevant state) andalso |
|
706 |
(if can (Proof.assert_bottom true) state |
|
53189
ee8b8dafef0e
discontinued parallel_subproofs_saturation and related internal counters (superseded by parallel_subproofs_threshold and timing information);
wenzelm
parents:
52788
diff
changeset
|
707 |
then Goal.future_enabled 1 |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
708 |
else |
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
709 |
(case estimate of |
53189
ee8b8dafef0e
discontinued parallel_subproofs_saturation and related internal counters (superseded by parallel_subproofs_threshold and timing information);
wenzelm
parents:
52788
diff
changeset
|
710 |
NONE => Goal.future_enabled 2 |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
711 |
| SOME t => Goal.future_enabled_timing t))); |
51278 | 712 |
|
51323 | 713 |
fun atom_result tr st = |
714 |
let |
|
715 |
val st' = |
|
53189
ee8b8dafef0e
discontinued parallel_subproofs_saturation and related internal counters (superseded by parallel_subproofs_threshold and timing information);
wenzelm
parents:
52788
diff
changeset
|
716 |
if Goal.future_enabled 1 andalso Keyword.is_diag (name_of tr) then |
53192 | 717 |
(Execution.fork |
51605
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
718 |
{name = "Toplevel.diag", pos = pos_of tr, |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
719 |
pri = priority (timing_estimate true (Thy_Syntax.atom tr))} |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
720 |
(fn () => command tr st); st) |
51323 | 721 |
else command tr st; |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
722 |
in (Result (tr, st'), st') end; |
51323 | 723 |
|
724 |
in |
|
725 |
||
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
726 |
fun element_result (Thy_Syntax.Element (tr, NONE)) st = atom_result tr st |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
727 |
| element_result (elem as Thy_Syntax.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
|
728 |
let |
51324 | 729 |
val (head_result, st') = atom_result head_tr st; |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
730 |
val (body_elems, end_tr) = element_rest; |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
731 |
val estimate = timing_estimate false elem; |
51324 | 732 |
in |
51423
e5f9a6d9ca82
clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents:
51332
diff
changeset
|
733 |
if not (proof_future_enabled estimate st') |
51324 | 734 |
then |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
735 |
let |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
736 |
val proof_trs = maps Thy_Syntax.flat_element body_elems @ [end_tr]; |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
737 |
val (proof_results, st'') = fold_map atom_result proof_trs st'; |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
738 |
in (Result_List (head_result :: proof_results), st'') end |
51324 | 739 |
else |
740 |
let |
|
741 |
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
|
742 |
|
51605
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
743 |
val future_proof = |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
744 |
Proof.future_proof (fn state => |
53192 | 745 |
Execution.fork |
51605
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
746 |
{name = "Toplevel.future_proof", pos = pos_of head_tr, pri = priority estimate} |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
747 |
(fn () => |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
748 |
let |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
749 |
val State (SOME (Proof (prf, (_, orig_gthy))), prev) = st'; |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
750 |
val prf' = Proof_Node.apply (K state) prf; |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
751 |
val (result, result_state) = |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
752 |
State (SOME (Proof (prf', (finish, orig_gthy))), prev) |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
753 |
|> fold_map element_result body_elems ||> command end_tr; |
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents:
51595
diff
changeset
|
754 |
in (Result_List result, presentation_context_of result_state) end)) |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
755 |
#> (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
|
756 |
|
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
757 |
val forked_proof = |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
758 |
proof (future_proof #> |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
759 |
(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
|
760 |
end_proof (fn _ => future_proof #> |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
761 |
(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
|
762 |
|
51324 | 763 |
val st'' = st' |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
764 |
|> command (head_tr |> set_print false |> reset_trans |> forked_proof); |
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
765 |
val end_result = Result (end_tr, st''); |
51324 | 766 |
val result = |
51332
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents:
51324
diff
changeset
|
767 |
Result_List [head_result, Result.get (presentation_context_of st''), end_result]; |
51324 | 768 |
in (result, st'') end |
769 |
end; |
|
28433
b3dab95f098f
begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents:
28425
diff
changeset
|
770 |
|
6664 | 771 |
end; |
51323 | 772 |
|
773 |
end; |