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