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