| author | wenzelm | 
| Tue, 09 Dec 2014 18:29:45 +0100 | |
| changeset 59117 | caddfa6ca534 | 
| parent 58950 | d07464875dd4 | 
| child 59150 | 71b416020f42 | 
| permissions | -rw-r--r-- | 
| 5820 | 1 | (* Title: Pure/Isar/proof.ML | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | ||
| 19000 | 4 | The Isar/VM proof language interpreter: maintains a structured flow of | 
| 5 | context elements, goals, refinements, and facts. | |
| 5820 | 6 | *) | 
| 7 | ||
| 8 | signature PROOF = | |
| 9 | sig | |
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32856diff
changeset | 10 | type context = Proof.context | 
| 23639 | 11 | type method = Method.method | 
| 5820 | 12 | type state | 
| 17359 | 13 | val init: context -> state | 
| 14 | val level: state -> int | |
| 15 | val assert_bottom: bool -> state -> state | |
| 5820 | 16 | val context_of: state -> context | 
| 17 | val theory_of: state -> theory | |
| 13377 | 18 | val map_context: (context -> context) -> state -> state | 
| 40642 
99c6ce92669b
added useful function map_context_result to signature
 bulwahn parents: 
40132diff
changeset | 19 | val map_context_result : (context -> 'a * context) -> state -> 'a * state | 
| 28278 | 20 | val map_contexts: (context -> context) -> state -> state | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37186diff
changeset | 21 | val propagate_ml_env: state -> state | 
| 30757 
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
 wenzelm parents: 
30557diff
changeset | 22 | val bind_terms: (indexname * term option) list -> state -> state | 
| 26251 
b8c6259d366b
put_facts: do_props, i.e. facts are indexed by proposition again;
 wenzelm parents: 
25958diff
changeset | 23 | val put_thms: bool -> string * thm list option -> state -> state | 
| 6091 | 24 | val the_facts: state -> thm list | 
| 9469 | 25 | val the_fact: state -> thm | 
| 47068 | 26 | val set_facts: thm list -> state -> state | 
| 27 | val reset_facts: state -> state | |
| 6891 | 28 | val assert_forward: state -> state | 
| 17359 | 29 | val assert_chain: state -> state | 
| 9469 | 30 | val assert_forward_or_chain: state -> state | 
| 5820 | 31 | val assert_backward: state -> state | 
| 8206 | 32 | val assert_no_chain: state -> state | 
| 5820 | 33 | val enter_forward: state -> state | 
| 58798 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 34 | val has_bottom_goal: state -> bool | 
| 12423 | 35 | val pretty_state: int -> state -> Pretty.T list | 
| 17112 | 36 | val refine: Method.text -> state -> state Seq.seq | 
| 37 | val refine_end: Method.text -> state -> state Seq.seq | |
| 18908 | 38 | val refine_insert: thm list -> state -> state | 
| 17359 | 39 | val refine_goals: (context -> thm -> unit) -> context -> thm list -> state -> state Seq.seq | 
| 33288 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 40 |   val raw_goal: state -> {context: context, facts: thm list, goal: thm}
 | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 41 |   val goal: state -> {context: context, facts: thm list, goal: thm}
 | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 42 |   val simple_goal: state -> {context: context, goal: thm}
 | 
| 38721 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38333diff
changeset | 43 | val status_markup: state -> Markup.T | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 44 | val let_bind: (term list * term) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 45 | val let_bind_cmd: (string list * string) list -> state -> state | 
| 36505 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 46 | val write: Syntax.mode -> (term * mixfix) list -> state -> state | 
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 47 | val write_cmd: Syntax.mode -> (string * mixfix) list -> state -> state | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 48 | val fix: (binding * typ option * mixfix) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 49 | val fix_cmd: (binding * string option * mixfix) list -> state -> state | 
| 20224 
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
 wenzelm parents: 
20208diff
changeset | 50 | val assm: Assumption.export -> | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 51 | (Thm.binding * (term * term list) list) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 52 | val assm_cmd: Assumption.export -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 53 | (Attrib.binding * (string * string list) list) list -> state -> state | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 54 | val assume: (Thm.binding * (term * term list) list) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 55 | val assume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 56 | val presume: (Thm.binding * (term * term list) list) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 57 | val presume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 58 | val def: (Thm.binding * ((binding * mixfix) * (term * term list))) list -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 59 | val def_cmd: (Attrib.binding * ((binding * mixfix) * (string * string list))) list -> state -> state | 
| 17359 | 60 | val chain: state -> state | 
| 61 | val chain_facts: thm list -> state -> state | |
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 62 | val note_thmss: (Thm.binding * (thm list * attribute list) list) list -> state -> state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 63 | val note_thmss_cmd: (Attrib.binding * (Facts.ref * Token.src list) list) list -> state -> state | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 64 | val from_thmss: ((thm list * attribute list) list) list -> state -> state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 65 | val from_thmss_cmd: ((Facts.ref * Token.src list) list) list -> state -> state | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 66 | val with_thmss: ((thm list * attribute list) list) list -> state -> state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 67 | val with_thmss_cmd: ((Facts.ref * Token.src list) list) list -> state -> state | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 68 | val using: ((thm list * attribute list) list) list -> state -> state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 69 | val using_cmd: ((Facts.ref * Token.src list) list) list -> state -> state | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 70 | val unfolding: ((thm list * attribute list) list) list -> state -> state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 71 | val unfolding_cmd: ((Facts.ref * Token.src list) list) list -> state -> state | 
| 53378 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53192diff
changeset | 72 | val invoke_case: (string * Position.T) * binding option list * attribute list -> | 
| 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53192diff
changeset | 73 | state -> state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 74 | val invoke_case_cmd: (string * Position.T) * binding option list * Token.src list -> | 
| 53378 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53192diff
changeset | 75 | state -> state | 
| 17359 | 76 | val begin_block: state -> state | 
| 77 | val next_block: state -> state | |
| 20309 | 78 | val end_block: state -> state | 
| 49042 | 79 | val begin_notepad: context -> state | 
| 80 | val end_notepad: state -> context | |
| 17112 | 81 | val proof: Method.text option -> state -> state Seq.seq | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 82 | val proof_results: Method.text_range option -> state -> state Seq.result Seq.seq | 
| 49865 | 83 | val defer: int -> state -> state | 
| 84 | val prefer: int -> state -> state | |
| 17112 | 85 | val apply: Method.text -> state -> state Seq.seq | 
| 86 | val apply_end: Method.text -> state -> state Seq.seq | |
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 87 | val apply_results: Method.text_range -> state -> state Seq.result Seq.seq | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 88 | val apply_end_results: Method.text_range -> state -> state Seq.result Seq.seq | 
| 17359 | 89 | val local_goal: (context -> ((string * string) * (string * thm list) list) -> unit) -> | 
| 49042 | 90 | (context -> 'a -> attribute) -> | 
| 45327 | 91 |     ('b list -> context -> (term list list * (context -> context)) * context) ->
 | 
| 29383 | 92 | string -> Method.text option -> (thm list list -> state -> state) -> | 
| 29581 | 93 | ((binding * 'a list) * 'b) list -> state -> state | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 94 | val local_qed: Method.text_range option * bool -> state -> state | 
| 21442 | 95 | val theorem: Method.text option -> (thm list list -> context -> context) -> | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 96 | (term * term list) list list -> context -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 97 | val theorem_cmd: Method.text option -> (thm list list -> context -> context) -> | 
| 21362 
3a2ab1dce297
simplified Proof.theorem(_i) interface -- removed target support;
 wenzelm parents: 
21274diff
changeset | 98 | (string * string list) list list -> context -> state | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 99 | val global_qed: Method.text_range option * bool -> state -> context | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 100 | val local_terminal_proof: Method.text_range * Method.text_range option -> state -> state | 
| 29383 | 101 | val local_default_proof: state -> state | 
| 102 | val local_immediate_proof: state -> state | |
| 103 | val local_skip_proof: bool -> state -> state | |
| 104 | val local_done_proof: state -> state | |
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 105 | val global_terminal_proof: Method.text_range * Method.text_range option -> state -> context | 
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20341diff
changeset | 106 | val global_default_proof: state -> context | 
| 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20341diff
changeset | 107 | val global_immediate_proof: state -> context | 
| 29383 | 108 | val global_skip_proof: bool -> state -> context | 
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20341diff
changeset | 109 | val global_done_proof: state -> context | 
| 29383 | 110 | val have: Method.text option -> (thm list list -> state -> state) -> | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 111 | (Thm.binding * (term * term list) list) list -> bool -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 112 | val have_cmd: Method.text option -> (thm list list -> state -> state) -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 113 | (Attrib.binding * (string * string list) list) list -> bool -> state -> state | 
| 29383 | 114 | val show: Method.text option -> (thm list list -> state -> state) -> | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 115 | (Thm.binding * (term * term list) list) list -> bool -> state -> state | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 116 | val show_cmd: Method.text option -> (thm list list -> state -> state) -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 117 | (Attrib.binding * (string * string list) list) list -> bool -> state -> state | 
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 118 | val schematic_goal: state -> bool | 
| 51226 
1973089f1dba
proper check of Proof.is_relevant (again, cf. c3e99efacb67 and df8fc0567a3d);
 wenzelm parents: 
51222diff
changeset | 119 | val is_relevant: state -> bool | 
| 51318 | 120 |   val future_proof: (state -> ('a * context) future) -> state -> 'a future * state
 | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 121 | val local_future_terminal_proof: Method.text_range * Method.text_range option -> bool -> | 
| 49866 | 122 | state -> state | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 123 | val global_future_terminal_proof: Method.text_range * Method.text_range option -> bool -> | 
| 49866 | 124 | state -> context | 
| 5820 | 125 | end; | 
| 126 | ||
| 13377 | 127 | structure Proof: PROOF = | 
| 5820 | 128 | struct | 
| 129 | ||
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32856diff
changeset | 130 | type context = Proof.context; | 
| 17112 | 131 | type method = Method.method; | 
| 16813 | 132 | |
| 5820 | 133 | |
| 134 | (** proof state **) | |
| 135 | ||
| 17359 | 136 | (* datatype state *) | 
| 5820 | 137 | |
| 17112 | 138 | datatype mode = Forward | Chain | Backward; | 
| 5820 | 139 | |
| 17359 | 140 | datatype state = | 
| 141 | State of node Stack.T | |
| 142 | and node = | |
| 7176 | 143 | Node of | 
| 144 |    {context: context,
 | |
| 145 | facts: thm list option, | |
| 146 | mode: mode, | |
| 17359 | 147 | goal: goal option} | 
| 148 | and goal = | |
| 149 | Goal of | |
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 150 |    {statement: (string * Position.T) * term list list * term,
 | 
| 28410 | 151 | (*goal kind and statement (starting with vars), initial proposition*) | 
| 25958 | 152 | using: thm list, (*goal facts*) | 
| 153 | goal: thm, (*subgoals ==> statement*) | |
| 17859 | 154 | before_qed: Method.text option, | 
| 18124 | 155 | after_qed: | 
| 29383 | 156 | (thm list list -> state -> state) * | 
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20341diff
changeset | 157 | (thm list list -> context -> context)}; | 
| 17359 | 158 | |
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 159 | fun make_goal (statement, using, goal, before_qed, after_qed) = | 
| 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 160 |   Goal {statement = statement, using = using, goal = goal,
 | 
| 17859 | 161 | before_qed = before_qed, after_qed = after_qed}; | 
| 5820 | 162 | |
| 7176 | 163 | fun make_node (context, facts, mode, goal) = | 
| 164 |   Node {context = context, facts = facts, mode = mode, goal = goal};
 | |
| 165 | ||
| 17359 | 166 | fun map_node f (Node {context, facts, mode, goal}) =
 | 
| 167 | make_node (f (context, facts, mode, goal)); | |
| 5820 | 168 | |
| 21727 | 169 | val init_context = | 
| 42360 | 170 | Proof_Context.set_stmt true #> | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46728diff
changeset | 171 | Proof_Context.map_naming (K Name_Space.local_naming); | 
| 21727 | 172 | |
| 21466 
6ffb8f455b84
init: enter inner statement mode, which prevents local notes from being named internally;
 wenzelm parents: 
21451diff
changeset | 173 | fun init ctxt = | 
| 21727 | 174 | State (Stack.init (make_node (init_context ctxt, NONE, Forward, NONE))); | 
| 5820 | 175 | |
| 58796 | 176 | fun top (State stack) = Stack.top stack |> (fn Node node => node); | 
| 177 | fun map_top f (State stack) = State (Stack.map_top (map_node f) stack); | |
| 178 | fun map_all f (State stack) = State (Stack.map_all (map_node f) stack); | |
| 12045 | 179 | |
| 5820 | 180 | |
| 181 | ||
| 182 | (** basic proof state operations **) | |
| 183 | ||
| 17359 | 184 | (* block structure *) | 
| 185 | ||
| 58796 | 186 | fun open_block (State stack) = State (Stack.push stack); | 
| 17359 | 187 | |
| 58796 | 188 | fun close_block (State stack) = State (Stack.pop stack) | 
| 47060 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 wenzelm parents: 
47005diff
changeset | 189 | handle List.Empty => error "Unbalanced block parentheses"; | 
| 17359 | 190 | |
| 58796 | 191 | fun level (State stack) = Stack.level stack; | 
| 17359 | 192 | |
| 193 | fun assert_bottom b state = | |
| 47065 | 194 | let val b' = level state <= 2 in | 
| 195 | if b andalso not b' then error "Not at bottom of proof" | |
| 196 | else if not b andalso b' then error "Already at bottom of proof" | |
| 17359 | 197 | else state | 
| 198 | end; | |
| 199 | ||
| 200 | ||
| 5820 | 201 | (* context *) | 
| 202 | ||
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 203 | val context_of = #context o top; | 
| 42360 | 204 | val theory_of = Proof_Context.theory_of o context_of; | 
| 5820 | 205 | |
| 47431 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 206 | fun map_node_context f = | 
| 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 207 | map_node (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal)); | 
| 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 208 | |
| 17359 | 209 | fun map_context f = | 
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 210 | map_top (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal)); | 
| 5820 | 211 | |
| 17359 | 212 | fun map_context_result f state = | 
| 17859 | 213 | f (context_of state) ||> (fn ctxt => map_context (K ctxt) state); | 
| 5820 | 214 | |
| 28278 | 215 | fun map_contexts f = map_all (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal)); | 
| 216 | ||
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37186diff
changeset | 217 | fun propagate_ml_env state = map_contexts | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37186diff
changeset | 218 | (Context.proof_map (ML_Env.inherit (Context.Proof (context_of state)))) state; | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37186diff
changeset | 219 | |
| 42360 | 220 | val bind_terms = map_context o Proof_Context.bind_terms; | 
| 221 | val put_thms = map_context oo Proof_Context.put_thms; | |
| 5820 | 222 | |
| 223 | ||
| 224 | (* facts *) | |
| 225 | ||
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 226 | val get_facts = #facts o top; | 
| 17359 | 227 | |
| 228 | fun the_facts state = | |
| 229 | (case get_facts state of SOME facts => facts | |
| 18678 | 230 | | NONE => error "No current facts available"); | 
| 5820 | 231 | |
| 9469 | 232 | fun the_fact state = | 
| 17359 | 233 | (case the_facts state of [thm] => thm | 
| 18678 | 234 | | _ => error "Single theorem expected"); | 
| 7605 | 235 | |
| 17359 | 236 | fun put_facts facts = | 
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 237 | map_top (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) #> | 
| 33386 | 238 | put_thms true (Auto_Bind.thisN, facts); | 
| 5820 | 239 | |
| 47068 | 240 | val set_facts = put_facts o SOME; | 
| 241 | val reset_facts = put_facts NONE; | |
| 242 | ||
| 17359 | 243 | fun these_factss more_facts (named_factss, state) = | 
| 47068 | 244 | (named_factss, state |> set_facts (maps snd named_factss @ more_facts)); | 
| 5820 | 245 | |
| 17359 | 246 | fun export_facts inner outer = | 
| 247 | (case get_facts inner of | |
| 47068 | 248 | NONE => reset_facts outer | 
| 17359 | 249 | | SOME thms => | 
| 250 | thms | |
| 42360 | 251 | |> Proof_Context.export (context_of inner) (context_of outer) | 
| 47068 | 252 | |> (fn ths => set_facts ths outer)); | 
| 5820 | 253 | |
| 254 | ||
| 255 | (* mode *) | |
| 256 | ||
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 257 | val get_mode = #mode o top; | 
| 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 258 | fun put_mode mode = map_top (fn (ctxt, facts, _, goal) => (ctxt, facts, mode, goal)); | 
| 5820 | 259 | |
| 17359 | 260 | val mode_name = (fn Forward => "state" | Chain => "chain" | Backward => "prove"); | 
| 5820 | 261 | |
| 262 | fun assert_mode pred state = | |
| 263 | let val mode = get_mode state in | |
| 264 | if pred mode then state | |
| 18678 | 265 |     else error ("Illegal application of proof command in " ^ quote (mode_name mode) ^ " mode")
 | 
| 5820 | 266 | end; | 
| 267 | ||
| 19308 | 268 | val assert_forward = assert_mode (fn mode => mode = Forward); | 
| 269 | val assert_chain = assert_mode (fn mode => mode = Chain); | |
| 270 | val assert_forward_or_chain = assert_mode (fn mode => mode = Forward orelse mode = Chain); | |
| 271 | val assert_backward = assert_mode (fn mode => mode = Backward); | |
| 272 | val assert_no_chain = assert_mode (fn mode => mode <> Chain); | |
| 5820 | 273 | |
| 17359 | 274 | val enter_forward = put_mode Forward; | 
| 275 | val enter_chain = put_mode Chain; | |
| 276 | val enter_backward = put_mode Backward; | |
| 5820 | 277 | |
| 17359 | 278 | |
| 279 | (* current goal *) | |
| 280 | ||
| 281 | fun current_goal state = | |
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 282 | (case top state of | 
| 17359 | 283 |     {context, goal = SOME (Goal goal), ...} => (context, goal)
 | 
| 47065 | 284 | | _ => error "No current goal"); | 
| 5820 | 285 | |
| 17359 | 286 | fun assert_current_goal g state = | 
| 287 | let val g' = can current_goal state in | |
| 47065 | 288 | if g andalso not g' then error "No goal in this block" | 
| 289 | else if not g andalso g' then error "Goal present in this block" | |
| 17359 | 290 | else state | 
| 291 | end; | |
| 6776 | 292 | |
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 293 | fun put_goal goal = map_top (fn (ctxt, using, mode, _) => (ctxt, using, mode, goal)); | 
| 17359 | 294 | |
| 47068 | 295 | val set_goal = put_goal o SOME; | 
| 296 | val reset_goal = put_goal NONE; | |
| 297 | ||
| 17859 | 298 | val before_qed = #before_qed o #2 o current_goal; | 
| 299 | ||
| 17359 | 300 | |
| 58798 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 301 | (* bottom goal *) | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 302 | |
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 303 | fun has_bottom_goal (State stack) = | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 304 | let | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 305 |     fun bottom [Node {goal = SOME _, ...}, Node {goal = NONE, ...}] = true
 | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 306 |       | bottom [Node {goal, ...}] = is_some goal
 | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 307 | | bottom [] = false | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 308 | | bottom (_ :: rest) = bottom rest; | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 309 | in bottom (op :: (Stack.dest stack)) end; | 
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 310 | |
| 
49ed5eea15d4
'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
 wenzelm parents: 
58797diff
changeset | 311 | |
| 17359 | 312 | (* nested goal *) | 
| 5820 | 313 | |
| 58797 | 314 | fun map_goal f g h (State stack) = | 
| 315 | (case Stack.dest stack of | |
| 316 |     (Node {context, facts, mode, goal = SOME goal}, node :: nodes) =>
 | |
| 17359 | 317 | let | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 318 |         val Goal {statement, using, goal, before_qed, after_qed} = goal;
 | 
| 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 319 | val goal' = make_goal (g (statement, using, goal, before_qed, after_qed)); | 
| 47431 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 320 | val node' = map_node_context h node; | 
| 58797 | 321 | val stack' = Stack.make (make_node (f context, facts, mode, SOME goal')) (node' :: nodes); | 
| 322 | in State stack' end | |
| 323 | | (nd, node :: nodes) => | |
| 47431 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 324 | let | 
| 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 325 | val nd' = map_node_context f nd; | 
| 58797 | 326 | val State stack' = map_goal f g h (State (Stack.make node nodes)); | 
| 327 | val (node', nodes') = Stack.dest stack'; | |
| 328 | in State (Stack.make nd' (node' :: nodes')) end | |
| 329 | | _ => State stack); | |
| 5820 | 330 | |
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 331 | fun provide_goal goal = map_goal I (fn (statement, using, _, before_qed, after_qed) => | 
| 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 332 | (statement, using, goal, before_qed, after_qed)) I; | 
| 19188 | 333 | |
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 334 | fun using_facts using = map_goal I (fn (statement, _, goal, before_qed, after_qed) => | 
| 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 335 | (statement, using, goal, before_qed, after_qed)) I; | 
| 17359 | 336 | |
| 337 | local | |
| 338 | fun find i state = | |
| 339 | (case try current_goal state of | |
| 340 | SOME (ctxt, goal) => (ctxt, (i, goal)) | |
| 18678 | 341 | | NONE => find (i + 1) (close_block state handle ERROR _ => error "No goal present")); | 
| 17359 | 342 | in val find_goal = find 0 end; | 
| 343 | ||
| 344 | fun get_goal state = | |
| 345 |   let val (ctxt, (_, {using, goal, ...})) = find_goal state
 | |
| 346 | in (ctxt, (using, goal)) end; | |
| 5820 | 347 | |
| 348 | ||
| 349 | ||
| 12423 | 350 | (** pretty_state **) | 
| 5820 | 351 | |
| 15531 | 352 | fun pretty_facts _ _ NONE = [] | 
| 51584 | 353 | | pretty_facts ctxt s (SOME ths) = [Proof_Display.pretty_goal_facts ctxt s ths, Pretty.str ""]; | 
| 6756 | 354 | |
| 17359 | 355 | fun pretty_state nr state = | 
| 5820 | 356 | let | 
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
47815diff
changeset | 357 |     val {context = ctxt, facts, mode, goal = _} = top state;
 | 
| 42717 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 wenzelm parents: 
42496diff
changeset | 358 | val verbose = Config.get ctxt Proof_Context.verbose; | 
| 5820 | 359 | |
| 49860 | 360 | fun prt_goal (SOME (_, (_, | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 361 |       {statement = ((_, pos), _, _), using, goal, before_qed = _, after_qed = _}))) =
 | 
| 51584 | 362 | pretty_facts ctxt "using" | 
| 17359 | 363 | (if mode <> Backward orelse null using then NONE else SOME using) @ | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 364 | [Proof_Display.pretty_goal_header goal] @ Goal_Display.pretty_goals ctxt goal | 
| 17359 | 365 | | prt_goal NONE = []; | 
| 6848 | 366 | |
| 17359 | 367 | val prt_ctxt = | 
| 42717 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 wenzelm parents: 
42496diff
changeset | 368 | if verbose orelse mode = Forward then Proof_Context.pretty_context ctxt | 
| 42360 | 369 | else if mode = Backward then Proof_Context.pretty_ctxt ctxt | 
| 7575 | 370 | else []; | 
| 56912 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56898diff
changeset | 371 | |
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56898diff
changeset | 372 | val position_markup = Position.markup (Position.thread_data ()) Markup.position; | 
| 17359 | 373 | in | 
| 56912 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56898diff
changeset | 374 | [Pretty.block | 
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56898diff
changeset | 375 | [Pretty.mark_str (position_markup, "proof"), | 
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56898diff
changeset | 376 |         Pretty.str (" (" ^ mode_name mode ^ "): depth " ^ string_of_int (level state div 2 - 1))],
 | 
| 17359 | 377 | Pretty.str ""] @ | 
| 378 | (if null prt_ctxt then [] else prt_ctxt @ [Pretty.str ""]) @ | |
| 42717 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 wenzelm parents: 
42496diff
changeset | 379 | (if verbose orelse mode = Forward then | 
| 51584 | 380 | pretty_facts ctxt "" facts @ prt_goal (try find_goal state) | 
| 381 | else if mode = Chain then pretty_facts ctxt "picking" facts | |
| 17359 | 382 | else prt_goal (try find_goal state)) | 
| 383 | end; | |
| 5820 | 384 | |
| 385 | ||
| 386 | ||
| 387 | (** proof steps **) | |
| 388 | ||
| 17359 | 389 | (* refine via method *) | 
| 5820 | 390 | |
| 8234 | 391 | local | 
| 392 | ||
| 16146 | 393 | fun goalN i = "goal" ^ string_of_int i; | 
| 394 | fun goals st = map goalN (1 upto Thm.nprems_of st); | |
| 395 | ||
| 396 | fun no_goal_cases st = map (rpair NONE) (goals st); | |
| 397 | ||
| 398 | fun goal_cases st = | |
| 47815 | 399 | Rule_Cases.make_common | 
| 400 | (Thm.theory_of_thm st, Thm.prop_of st) (map (rpair [] o rpair []) (goals st)); | |
| 16146 | 401 | |
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58004diff
changeset | 402 | fun apply_method text ctxt state = | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 403 |   #2 (#2 (find_goal state)) |> (fn {statement, using, goal, before_qed, after_qed} =>
 | 
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58004diff
changeset | 404 | Method.evaluate text ctxt using goal | 
| 58004 | 405 | |> Seq.map (fn (meth_cases, goal') => | 
| 6848 | 406 | state | 
| 16146 | 407 | |> map_goal | 
| 53378 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53192diff
changeset | 408 | (Proof_Context.update_cases false (no_goal_cases goal @ goal_cases goal') #> | 
| 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53192diff
changeset | 409 | Proof_Context.update_cases true meth_cases) | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 410 | (K (statement, using, goal', before_qed, after_qed)) I)); | 
| 5820 | 411 | |
| 8234 | 412 | in | 
| 413 | ||
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58004diff
changeset | 414 | fun refine text state = apply_method text (context_of state) state; | 
| 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58004diff
changeset | 415 | fun refine_end text state = apply_method text (#1 (find_goal state)) state; | 
| 58004 | 416 | |
| 417 | fun refine_insert ths = | |
| 418 | Seq.hd o refine (Method.Basic (K (Method.insert ths))); | |
| 18908 | 419 | |
| 8234 | 420 | end; | 
| 421 | ||
| 5820 | 422 | |
| 17359 | 423 | (* refine via sub-proof *) | 
| 424 | ||
| 46466 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 425 | local | 
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 426 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 427 | fun finish_tac _ 0 = K all_tac | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 428 | | finish_tac ctxt n = | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 429 | Goal.norm_hhf_tac ctxt THEN' | 
| 30557 
a28d83e903ce
goal_tac: finish marked assumptions from left to right -- corresponds better with the strategy of etac, with significant performance gains in some situations;
 wenzelm parents: 
30510diff
changeset | 430 | SUBGOAL (fn (goal, i) => | 
| 
a28d83e903ce
goal_tac: finish marked assumptions from left to right -- corresponds better with the strategy of etac, with significant performance gains in some situations;
 wenzelm parents: 
30510diff
changeset | 431 | if can Logic.unprotect (Logic.strip_assums_concl goal) then | 
| 58837 | 432 | eresolve_tac [Drule.protectI] i THEN finish_tac ctxt (n - 1) i | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 433 | else finish_tac ctxt (n - 1) (i + 1)); | 
| 30557 
a28d83e903ce
goal_tac: finish marked assumptions from left to right -- corresponds better with the strategy of etac, with significant performance gains in some situations;
 wenzelm parents: 
30510diff
changeset | 434 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 435 | fun goal_tac ctxt rule = | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 436 | Goal.norm_hhf_tac ctxt THEN' | 
| 58837 | 437 | resolve_tac [rule] THEN' | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 438 | finish_tac ctxt (Thm.nprems_of rule); | 
| 11816 | 439 | |
| 46466 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 440 | fun FINDGOAL tac st = | 
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 441 | let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n) | 
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 442 | in find 1 (Thm.nprems_of st) st end; | 
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 443 | |
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 444 | in | 
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 445 | |
| 19915 | 446 | fun refine_goals print_rule inner raw_rules state = | 
| 447 | let | |
| 448 | val (outer, (_, goal)) = get_goal state; | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 449 | fun refine rule st = (print_rule outer rule; FINDGOAL (goal_tac outer rule) st); | 
| 19915 | 450 | in | 
| 451 | raw_rules | |
| 42360 | 452 | |> Proof_Context.goal_export inner outer | 
| 47068 | 453 | |> (fn rules => Seq.lift provide_goal (EVERY (map refine rules) goal) state) | 
| 19915 | 454 | end; | 
| 17359 | 455 | |
| 46466 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 456 | end; | 
| 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45666diff
changeset | 457 | |
| 6932 | 458 | |
| 49846 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
49748diff
changeset | 459 | (* conclude goal *) | 
| 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
49748diff
changeset | 460 | |
| 28627 
63663cfa297c
conclude_goal: precise goal context, include all sorts from context into statement,  check shyps of result;
 wenzelm parents: 
28458diff
changeset | 461 | fun conclude_goal ctxt goal propss = | 
| 5820 | 462 | let | 
| 42360 | 463 | val thy = Proof_Context.theory_of ctxt; | 
| 5820 | 464 | |
| 54981 | 465 | val _ = | 
| 466 | Theory.subthy (theory_of_thm goal, thy) orelse error "Bad background theory of goal state"; | |
| 49860 | 467 | val _ = Thm.no_prems goal orelse error (Proof_Display.string_of_goal ctxt goal); | 
| 20224 
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
 wenzelm parents: 
20208diff
changeset | 468 | |
| 54567 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53380diff
changeset | 469 |     fun lost_structure () = error ("Lost goal structure:\n" ^ Display.string_of_thm ctxt goal);
 | 
| 5820 | 470 | |
| 54567 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53380diff
changeset | 471 | val th = | 
| 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53380diff
changeset | 472 | (Goal.conclude (if length (flat propss) > 1 then Thm.norm_proof goal else goal) | 
| 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53380diff
changeset | 473 | handle THM _ => lost_structure ()) | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58892diff
changeset | 474 | |> Drule.flexflex_unique (SOME ctxt) | 
| 54567 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53380diff
changeset | 475 | |> Thm.check_shyps (Variable.sorts_of ctxt) | 
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 476 | |> Thm.check_hyps (Context.Proof ctxt); | 
| 23418 | 477 | |
| 478 | val goal_propss = filter_out null propss; | |
| 479 | val results = | |
| 480 | Conjunction.elim_balanced (length goal_propss) th | |
| 481 | |> map2 Conjunction.elim_balanced (map length goal_propss) | |
| 482 | handle THM _ => lost_structure (); | |
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58892diff
changeset | 483 | val _ = | 
| 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58892diff
changeset | 484 | Unify.matches_list (Context.Proof ctxt) (flat goal_propss) (map Thm.prop_of (flat results)) | 
| 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58892diff
changeset | 485 |         orelse error ("Proved a different theorem:\n" ^ Display.string_of_thm ctxt th);
 | 
| 23418 | 486 | |
| 487 | fun recover_result ([] :: pss) thss = [] :: recover_result pss thss | |
| 488 | | recover_result (_ :: pss) (ths :: thss) = ths :: recover_result pss thss | |
| 489 | | recover_result [] [] = [] | |
| 490 | | recover_result _ _ = lost_structure (); | |
| 491 | in recover_result propss results end; | |
| 5820 | 492 | |
| 49859 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 493 | val finished_goal_error = "Failed to finish proof"; | 
| 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 494 | |
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 495 | fun finished_goal pos state = | 
| 49859 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 496 | let val (ctxt, (_, goal)) = get_goal state in | 
| 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 497 | if Thm.no_prems goal then Seq.Result state | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 498 | else | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 499 | Seq.Error (fn () => | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 500 | finished_goal_error ^ Position.here pos ^ ":\n" ^ | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 501 | Proof_Display.string_of_goal ctxt goal) | 
| 49859 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 502 | end; | 
| 49846 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
49748diff
changeset | 503 | |
| 5820 | 504 | |
| 33288 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 505 | (* goal views -- corresponding to methods *) | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 506 | |
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 507 | fun raw_goal state = | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 508 | let val (ctxt, (facts, goal)) = get_goal state | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 509 |   in {context = ctxt, facts = facts, goal = goal} end;
 | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 510 | |
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 511 | val goal = raw_goal o refine_insert []; | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 512 | |
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 513 | fun simple_goal state = | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 514 | let | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 515 | val (_, (facts, _)) = get_goal state; | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 516 | val (ctxt, (_, goal)) = get_goal (refine_insert facts state); | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 517 |   in {context = ctxt, goal = goal} end;
 | 
| 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 518 | |
| 38721 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38333diff
changeset | 519 | fun status_markup state = | 
| 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38333diff
changeset | 520 | (case try goal state of | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49912diff
changeset | 521 |     SOME {goal, ...} => Markup.proof_state (Thm.nprems_of goal)
 | 
| 38721 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38333diff
changeset | 522 | | NONE => Markup.empty); | 
| 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38333diff
changeset | 523 | |
| 49866 | 524 | fun method_error kind pos state = | 
| 525 | Seq.single (Proof_Display.method_error kind pos (raw_goal state)); | |
| 49861 | 526 | |
| 33288 
bd3f30da7bc1
replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
 wenzelm parents: 
33165diff
changeset | 527 | |
| 5820 | 528 | |
| 529 | (*** structured proof commands ***) | |
| 530 | ||
| 17112 | 531 | (** context elements **) | 
| 5820 | 532 | |
| 36324 | 533 | (* let bindings *) | 
| 5820 | 534 | |
| 16813 | 535 | local | 
| 536 | ||
| 17359 | 537 | fun gen_bind bind args state = | 
| 5820 | 538 | state | 
| 539 | |> assert_forward | |
| 36324 | 540 | |> map_context (bind true args #> snd) | 
| 47068 | 541 | |> reset_facts; | 
| 5820 | 542 | |
| 16813 | 543 | in | 
| 544 | ||
| 42360 | 545 | val let_bind = gen_bind Proof_Context.match_bind_i; | 
| 546 | val let_bind_cmd = gen_bind Proof_Context.match_bind; | |
| 5820 | 547 | |
| 16813 | 548 | end; | 
| 549 | ||
| 5820 | 550 | |
| 36505 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 551 | (* concrete syntax *) | 
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 552 | |
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 553 | local | 
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 554 | |
| 36507 
c966a1aab860
'write': actually observe the proof structure (like 'let' or 'fix');
 wenzelm parents: 
36505diff
changeset | 555 | fun gen_write prep_arg mode args = | 
| 
c966a1aab860
'write': actually observe the proof structure (like 'let' or 'fix');
 wenzelm parents: 
36505diff
changeset | 556 | assert_forward | 
| 42360 | 557 | #> map_context (fn ctxt => ctxt |> Proof_Context.notation true mode (map (prep_arg ctxt) args)) | 
| 47068 | 558 | #> reset_facts; | 
| 36505 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 559 | |
| 55955 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 560 | fun read_arg ctxt (c, mx) = | 
| 56002 | 561 |   (case Proof_Context.read_const {proper = false, strict = false} ctxt c of
 | 
| 55955 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 562 | Free (x, _) => | 
| 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 563 | let val T = Proof_Context.infer_type ctxt (x, Mixfix.mixfixT mx) | 
| 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 564 | in (Free (x, T), mx) end | 
| 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 565 | | t => (t, mx)); | 
| 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 566 | |
| 36505 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 567 | in | 
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 568 | |
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 569 | val write = gen_write (K I); | 
| 55955 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 wenzelm parents: 
55954diff
changeset | 570 | val write_cmd = gen_write read_arg; | 
| 36505 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 571 | |
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 572 | end; | 
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 573 | |
| 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 wenzelm parents: 
36354diff
changeset | 574 | |
| 17359 | 575 | (* fix *) | 
| 9196 | 576 | |
| 12714 | 577 | local | 
| 578 | ||
| 30763 
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
 wenzelm parents: 
30761diff
changeset | 579 | fun gen_fix prep_vars args = | 
| 17359 | 580 | assert_forward | 
| 45597 | 581 | #> map_context (fn ctxt => snd (Proof_Context.add_fixes (fst (prep_vars args ctxt)) ctxt)) | 
| 47068 | 582 | #> reset_facts; | 
| 5820 | 583 | |
| 16813 | 584 | in | 
| 585 | ||
| 45597 | 586 | val fix = gen_fix Proof_Context.cert_vars; | 
| 587 | val fix_cmd = gen_fix Proof_Context.read_vars; | |
| 5820 | 588 | |
| 16813 | 589 | end; | 
| 590 | ||
| 5820 | 591 | |
| 17359 | 592 | (* assume etc. *) | 
| 5820 | 593 | |
| 16813 | 594 | local | 
| 595 | ||
| 17112 | 596 | fun gen_assume asm prep_att exp args state = | 
| 5820 | 597 | state | 
| 598 | |> assert_forward | |
| 47815 | 599 | |> map_context_result (asm exp (Attrib.map_specs (map (prep_att (context_of state))) args)) | 
| 17359 | 600 | |> these_factss [] |> #2; | 
| 6932 | 601 | |
| 16813 | 602 | in | 
| 603 | ||
| 42360 | 604 | val assm = gen_assume Proof_Context.add_assms_i (K I); | 
| 47815 | 605 | val assm_cmd = gen_assume Proof_Context.add_assms Attrib.attribute_cmd; | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 606 | val assume = assm Assumption.assume_export; | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 607 | val assume_cmd = assm_cmd Assumption.assume_export; | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 608 | val presume = assm Assumption.presume_export; | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 609 | val presume_cmd = assm_cmd Assumption.presume_export; | 
| 5820 | 610 | |
| 16813 | 611 | end; | 
| 612 | ||
| 7271 | 613 | |
| 17359 | 614 | (* def *) | 
| 11891 | 615 | |
| 16813 | 616 | local | 
| 617 | ||
| 20913 | 618 | fun gen_def prep_att prep_vars prep_binds args state = | 
| 11891 | 619 | let | 
| 620 | val _ = assert_forward state; | |
| 20913 | 621 | val (raw_name_atts, (raw_vars, raw_rhss)) = args |> split_list ||> split_list; | 
| 47815 | 622 | val name_atts = map (apsnd (map (prep_att (context_of state)))) raw_name_atts; | 
| 11891 | 623 | in | 
| 20913 | 624 | state | 
| 625 | |> map_context_result (prep_vars (map (fn (x, mx) => (x, NONE, mx)) raw_vars)) | |
| 626 | |>> map (fn (x, _, mx) => (x, mx)) | |
| 627 | |-> (fn vars => | |
| 628 | map_context_result (prep_binds false (map swap raw_rhss)) | |
| 49748 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
49063diff
changeset | 629 | #-> (fn rhss => | 
| 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
49063diff
changeset | 630 | let | 
| 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
49063diff
changeset | 631 | val defs = (vars ~~ (name_atts ~~ rhss)) |> map (fn ((x, mx), ((a, atts), rhs)) => | 
| 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
49063diff
changeset | 632 | ((x, mx), ((Thm.def_binding_optional x a, atts), rhs))); | 
| 
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
 wenzelm parents: 
49063diff
changeset | 633 | in map_context_result (Local_Defs.add_defs defs) end)) | 
| 47068 | 634 | |-> (set_facts o map (#2 o #2)) | 
| 11891 | 635 | end; | 
| 636 | ||
| 16813 | 637 | in | 
| 638 | ||
| 42360 | 639 | val def = gen_def (K I) Proof_Context.cert_vars Proof_Context.match_bind_i; | 
| 47815 | 640 | val def_cmd = gen_def Attrib.attribute_cmd Proof_Context.read_vars Proof_Context.match_bind; | 
| 11891 | 641 | |
| 16813 | 642 | end; | 
| 643 | ||
| 11891 | 644 | |
| 8374 | 645 | |
| 17112 | 646 | (** facts **) | 
| 5820 | 647 | |
| 17359 | 648 | (* chain *) | 
| 5820 | 649 | |
| 24011 | 650 | fun clean_facts ctxt = | 
| 47068 | 651 | set_facts (filter_out Thm.is_dummy (the_facts ctxt)) ctxt; | 
| 24011 | 652 | |
| 17359 | 653 | val chain = | 
| 654 | assert_forward | |
| 24011 | 655 | #> clean_facts | 
| 17359 | 656 | #> enter_chain; | 
| 5820 | 657 | |
| 17359 | 658 | fun chain_facts facts = | 
| 47068 | 659 | set_facts facts | 
| 17359 | 660 | #> chain; | 
| 5820 | 661 | |
| 662 | ||
| 17359 | 663 | (* note etc. *) | 
| 17112 | 664 | |
| 28965 | 665 | fun no_binding args = map (pair (Binding.empty, [])) args; | 
| 17112 | 666 | |
| 667 | local | |
| 668 | ||
| 30760 
0fffc66b10d7
simplified references to facts, eliminated external note_thmss;
 wenzelm parents: 
30757diff
changeset | 669 | fun gen_thmss more_facts opt_chain opt_result prep_atts prep_fact args state = | 
| 17112 | 670 | state | 
| 671 | |> assert_forward | |
| 47815 | 672 | |> map_context_result (fn ctxt => ctxt |> Proof_Context.note_thmss "" | 
| 673 | (Attrib.map_facts_refs (map (prep_atts ctxt)) (prep_fact ctxt) args)) | |
| 17112 | 674 | |> these_factss (more_facts state) | 
| 17359 | 675 | ||> opt_chain | 
| 676 | |> opt_result; | |
| 17112 | 677 | |
| 678 | in | |
| 679 | ||
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 680 | val note_thmss = gen_thmss (K []) I #2 (K I) (K I); | 
| 47815 | 681 | val note_thmss_cmd = gen_thmss (K []) I #2 Attrib.attribute_cmd Proof_Context.get_fact; | 
| 17112 | 682 | |
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 683 | val from_thmss = gen_thmss (K []) chain #2 (K I) (K I) o no_binding; | 
| 47815 | 684 | val from_thmss_cmd = | 
| 685 | gen_thmss (K []) chain #2 Attrib.attribute_cmd Proof_Context.get_fact o no_binding; | |
| 17359 | 686 | |
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 687 | val with_thmss = gen_thmss the_facts chain #2 (K I) (K I) o no_binding; | 
| 47815 | 688 | val with_thmss_cmd = | 
| 689 | gen_thmss the_facts chain #2 Attrib.attribute_cmd Proof_Context.get_fact o no_binding; | |
| 30760 
0fffc66b10d7
simplified references to facts, eliminated external note_thmss;
 wenzelm parents: 
30757diff
changeset | 690 | |
| 
0fffc66b10d7
simplified references to facts, eliminated external note_thmss;
 wenzelm parents: 
30757diff
changeset | 691 | val local_results = gen_thmss (K []) I I (K I) (K I) o map (apsnd Thm.simple_fact); | 
| 17359 | 692 | |
| 17112 | 693 | end; | 
| 694 | ||
| 695 | ||
| 18548 | 696 | (* using/unfolding *) | 
| 17112 | 697 | |
| 698 | local | |
| 699 | ||
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45327diff
changeset | 700 | fun gen_using f g prep_att prep_fact args state = | 
| 17112 | 701 | state | 
| 702 | |> assert_backward | |
| 21442 | 703 | |> map_context_result | 
| 47815 | 704 | (fn ctxt => ctxt |> Proof_Context.note_thmss "" | 
| 705 | (Attrib.map_facts_refs (map (prep_att ctxt)) (prep_fact ctxt) (no_binding args))) | |
| 18843 | 706 | |> (fn (named_facts, state') => | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 707 | state' |> map_goal I (fn (statement, using, goal, before_qed, after_qed) => | 
| 18843 | 708 | let | 
| 709 | val ctxt = context_of state'; | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19475diff
changeset | 710 | val ths = maps snd named_facts; | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 711 | in (statement, f ctxt ths using, g ctxt ths goal, before_qed, after_qed) end) I); | 
| 18548 | 712 | |
| 24050 | 713 | fun append_using _ ths using = using @ filter_out Thm.is_dummy ths; | 
| 35624 | 714 | fun unfold_using ctxt ths = map (Local_Defs.unfold ctxt ths); | 
| 715 | val unfold_goals = Local_Defs.unfold_goals; | |
| 17112 | 716 | |
| 717 | in | |
| 718 | ||
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 719 | val using = gen_using append_using (K (K I)) (K I) (K I); | 
| 47815 | 720 | val using_cmd = gen_using append_using (K (K I)) Attrib.attribute_cmd Proof_Context.get_fact; | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 721 | val unfolding = gen_using unfold_using unfold_goals (K I) (K I); | 
| 47815 | 722 | val unfolding_cmd = gen_using unfold_using unfold_goals Attrib.attribute_cmd Proof_Context.get_fact; | 
| 17112 | 723 | |
| 724 | end; | |
| 725 | ||
| 726 | ||
| 727 | (* case *) | |
| 728 | ||
| 729 | local | |
| 730 | ||
| 57486 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 wenzelm parents: 
56933diff
changeset | 731 | fun gen_invoke_case internal prep_att ((name, pos), xs, raw_atts) state = | 
| 17112 | 732 | let | 
| 47815 | 733 | val atts = map (prep_att (context_of state)) raw_atts; | 
| 19078 | 734 | val (asms, state') = state |> map_context_result (fn ctxt => | 
| 57486 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 wenzelm parents: 
56933diff
changeset | 735 | ctxt |> Proof_Context.apply_case (Proof_Context.check_case ctxt internal (name, pos) xs)); | 
| 53380 
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
 wenzelm parents: 
53378diff
changeset | 736 | val assumptions = | 
| 
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
 wenzelm parents: 
53378diff
changeset | 737 | asms |> map (fn (b, ts) => ((Binding.set_pos pos b, atts), map (rpair []) ts)); | 
| 17112 | 738 | in | 
| 739 | state' | |
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36322diff
changeset | 740 | |> assume assumptions | 
| 33386 | 741 | |> bind_terms Auto_Bind.no_facts | 
| 53378 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53192diff
changeset | 742 | |> `the_facts |-> (fn thms => note_thmss [((Binding.make (name, pos), []), [(thms, [])])]) | 
| 17112 | 743 | end; | 
| 744 | ||
| 745 | in | |
| 746 | ||
| 57486 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 wenzelm parents: 
56933diff
changeset | 747 | val invoke_case = gen_invoke_case true (K I); | 
| 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 wenzelm parents: 
56933diff
changeset | 748 | val invoke_case_cmd = gen_invoke_case false Attrib.attribute_cmd; | 
| 17112 | 749 | |
| 750 | end; | |
| 751 | ||
| 752 | ||
| 753 | ||
| 17359 | 754 | (** proof structure **) | 
| 755 | ||
| 756 | (* blocks *) | |
| 757 | ||
| 758 | val begin_block = | |
| 759 | assert_forward | |
| 760 | #> open_block | |
| 47068 | 761 | #> reset_goal | 
| 17359 | 762 | #> open_block; | 
| 763 | ||
| 764 | val next_block = | |
| 765 | assert_forward | |
| 766 | #> close_block | |
| 767 | #> open_block | |
| 47068 | 768 | #> reset_goal | 
| 769 | #> reset_facts; | |
| 17359 | 770 | |
| 771 | fun end_block state = | |
| 772 | state | |
| 773 | |> assert_forward | |
| 40960 | 774 | |> assert_bottom false | 
| 17359 | 775 | |> close_block | 
| 776 | |> assert_current_goal false | |
| 777 | |> close_block | |
| 778 | |> export_facts state; | |
| 779 | ||
| 780 | ||
| 40960 | 781 | (* global notepad *) | 
| 782 | ||
| 783 | val begin_notepad = | |
| 784 | init | |
| 785 | #> open_block | |
| 786 | #> map_context (Variable.set_body true) | |
| 787 | #> open_block; | |
| 788 | ||
| 789 | val end_notepad = | |
| 790 | assert_forward | |
| 791 | #> assert_bottom true | |
| 792 | #> close_block | |
| 793 | #> assert_current_goal false | |
| 794 | #> close_block | |
| 795 | #> context_of; | |
| 796 | ||
| 797 | ||
| 17359 | 798 | (* sub-proofs *) | 
| 799 | ||
| 800 | fun proof opt_text = | |
| 801 | assert_backward | |
| 17859 | 802 | #> refine (the_default Method.default_text opt_text) | 
| 17359 | 803 | #> Seq.map (using_facts [] #> enter_forward); | 
| 804 | ||
| 49866 | 805 | fun proof_results arg = | 
| 806 | Seq.APPEND (proof (Method.text arg) #> Seq.make_results, | |
| 807 | method_error "initial" (Method.position arg)); | |
| 49863 
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
 wenzelm parents: 
49861diff
changeset | 808 | |
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 809 | fun end_proof bot (prev_pos, (opt_text, immed)) = | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 810 | let | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 811 | val (finish_text, terminal_pos, finished_pos) = | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 812 | (case opt_text of | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 813 | NONE => (Method.finish_text (NONE, immed), Position.none, prev_pos) | 
| 55709 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 wenzelm parents: 
54993diff
changeset | 814 | | SOME (text, (pos, end_pos)) => (Method.finish_text (SOME text, immed), pos, end_pos)); | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 815 | in | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 816 | Seq.APPEND (fn state => | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 817 | state | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 818 | |> assert_forward | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 819 | |> assert_bottom bot | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 820 | |> close_block | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 821 | |> assert_current_goal true | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 822 | |> using_facts [] | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 823 | |> `before_qed |-> (refine o the_default Method.succeed_text) | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 824 | |> Seq.maps (refine finish_text) | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 825 | |> Seq.make_results, method_error "terminal" terminal_pos) | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 826 | #> Seq.maps_results (Seq.single o finished_goal finished_pos) | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 827 | end; | 
| 17359 | 828 | |
| 29383 | 829 | fun check_result msg sq = | 
| 830 | (case Seq.pull sq of | |
| 831 | NONE => error msg | |
| 832 | | SOME (s, _) => s); | |
| 833 | ||
| 17359 | 834 | |
| 835 | (* unstructured refinement *) | |
| 836 | ||
| 49865 | 837 | fun defer i = | 
| 838 | assert_no_chain #> | |
| 839 | refine (Method.Basic (fn _ => METHOD (fn _ => ASSERT_SUBGOAL defer_tac i))) #> Seq.hd; | |
| 840 | ||
| 841 | fun prefer i = | |
| 842 | assert_no_chain #> | |
| 843 | refine (Method.Basic (fn _ => METHOD (fn _ => ASSERT_SUBGOAL prefer_tac i))) #> Seq.hd; | |
| 17359 | 844 | |
| 845 | fun apply text = assert_backward #> refine text #> Seq.map (using_facts []); | |
| 49866 | 846 | |
| 17359 | 847 | fun apply_end text = assert_forward #> refine_end text; | 
| 848 | ||
| 55709 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 wenzelm parents: 
54993diff
changeset | 849 | fun apply_results (text, (pos, _)) = | 
| 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 wenzelm parents: 
54993diff
changeset | 850 | Seq.APPEND (apply text #> Seq.make_results, method_error "" pos); | 
| 49866 | 851 | |
| 55709 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 wenzelm parents: 
54993diff
changeset | 852 | fun apply_end_results (text, (pos, _)) = | 
| 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 wenzelm parents: 
54993diff
changeset | 853 | Seq.APPEND (apply_end text #> Seq.make_results, method_error "" pos); | 
| 49863 
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
 wenzelm parents: 
49861diff
changeset | 854 | |
| 17359 | 855 | |
| 856 | ||
| 17112 | 857 | (** goals **) | 
| 858 | ||
| 17359 | 859 | (* generic goals *) | 
| 860 | ||
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 861 | local | 
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 862 | |
| 36322 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 863 | val is_var = | 
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 864 | can (dest_TVar o Logic.dest_type o Logic.dest_term) orf | 
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 865 | can (dest_Var o Logic.dest_term); | 
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 866 | |
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 867 | fun implicit_vars props = | 
| 19846 | 868 | let | 
| 36354 | 869 | val (var_props, _) = take_prefix is_var props; | 
| 36322 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 870 | val explicit_vars = fold Term.add_vars var_props []; | 
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 871 | val vars = filter_out (member (op =) explicit_vars) (fold Term.add_vars props []); | 
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 872 | in map (Logic.mk_term o Var) vars end; | 
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 873 | |
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 874 | fun refine_terms n = | 
| 58002 | 875 | refine (Method.Basic (K (NO_CASES o | 
| 876 | K (HEADGOAL (PRECISE_CONJUNCTS n | |
| 58837 | 877 | (HEADGOAL (CONJUNCTS (ALLGOALS (resolve_tac [Drule.termI]))))))))) | 
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 878 | #> Seq.hd; | 
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 879 | |
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 880 | in | 
| 17359 | 881 | |
| 17859 | 882 | fun generic_goal prepp kind before_qed after_qed raw_propp state = | 
| 5820 | 883 | let | 
| 17359 | 884 | val thy = theory_of state; | 
| 23418 | 885 | val cert = Thm.cterm_of thy; | 
| 17359 | 886 | val chaining = can assert_chain state; | 
| 25958 | 887 | val pos = Position.thread_data (); | 
| 17359 | 888 | |
| 889 | val ((propss, after_ctxt), goal_state) = | |
| 5936 | 890 | state | 
| 891 | |> assert_forward_or_chain | |
| 892 | |> enter_forward | |
| 17359 | 893 | |> open_block | 
| 45327 | 894 | |> map_context_result (prepp raw_propp); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19475diff
changeset | 895 | val props = flat propss; | 
| 15703 | 896 | |
| 36322 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 897 | val vars = implicit_vars props; | 
| 
81cba3921ba5
goals: simplified handling of implicit variables -- removed obsolete warning;
 wenzelm parents: 
35624diff
changeset | 898 | val propss' = vars :: propss; | 
| 23418 | 899 | val goal_propss = filter_out null propss'; | 
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 900 | val goal = | 
| 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 901 | cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss)) | 
| 28627 
63663cfa297c
conclude_goal: precise goal context, include all sorts from context into statement,  check shyps of result;
 wenzelm parents: 
28458diff
changeset | 902 | |> Thm.weaken_sorts (Variable.sorts_of (context_of goal_state)); | 
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 903 | val statement = ((kind, pos), propss', Thm.term_of goal); | 
| 18124 | 904 | val after_qed' = after_qed |>> (fn after_local => | 
| 905 | fn results => map_context after_ctxt #> after_local results); | |
| 5820 | 906 | in | 
| 17359 | 907 | goal_state | 
| 21727 | 908 | |> map_context (init_context #> Variable.set_body true) | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 909 | |> set_goal (make_goal (statement, [], Goal.init goal, before_qed, after_qed')) | 
| 42360 | 910 | |> map_context (Proof_Context.auto_bind_goal props) | 
| 21565 | 911 | |> chaining ? (`the_facts #-> using_facts) | 
| 47068 | 912 | |> reset_facts | 
| 17359 | 913 | |> open_block | 
| 47068 | 914 | |> reset_goal | 
| 5820 | 915 | |> enter_backward | 
| 23418 | 916 | |> not (null vars) ? refine_terms (length goal_propss) | 
| 32193 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 wenzelm parents: 
32187diff
changeset | 917 | |> null props ? (refine (Method.Basic Method.assumption) #> Seq.hd) | 
| 5820 | 918 | end; | 
| 919 | ||
| 29090 
bbfac5fd8d78
global_qed: refrain from ProofContext.auto_bind_facts, to avoid
 wenzelm parents: 
28981diff
changeset | 920 | fun generic_qed after_ctxt state = | 
| 12503 | 921 | let | 
| 36354 | 922 |     val (goal_ctxt, {statement = (_, stmt, _), goal, after_qed, ...}) = current_goal state;
 | 
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 923 | val outer_state = state |> close_block; | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 924 | val outer_ctxt = context_of outer_state; | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 925 | |
| 18503 
841137f20307
goal/qed: proper treatment of two levels of conjunctions;
 wenzelm parents: 
18475diff
changeset | 926 | val props = | 
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 927 | flat (tl stmt) | 
| 19915 | 928 | |> Variable.exportT_terms goal_ctxt outer_ctxt; | 
| 17359 | 929 | val results = | 
| 28627 
63663cfa297c
conclude_goal: precise goal context, include all sorts from context into statement,  check shyps of result;
 wenzelm parents: 
28458diff
changeset | 930 | tl (conclude_goal goal_ctxt goal stmt) | 
| 42360 | 931 | |> burrow (Proof_Context.export goal_ctxt outer_ctxt); | 
| 17359 | 932 | in | 
| 933 | outer_state | |
| 29090 
bbfac5fd8d78
global_qed: refrain from ProofContext.auto_bind_facts, to avoid
 wenzelm parents: 
28981diff
changeset | 934 | |> map_context (after_ctxt props) | 
| 34239 
e18b0f7b9902
after_qed: refrain from Position.setmp_thread_data, which causes duplication of results with several independent proof attempts;
 wenzelm parents: 
33389diff
changeset | 935 | |> pair (after_qed, results) | 
| 17359 | 936 | end; | 
| 937 | ||
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 938 | end; | 
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19585diff
changeset | 939 | |
| 9469 | 940 | |
| 17359 | 941 | (* local goals *) | 
| 942 | ||
| 17859 | 943 | fun local_goal print_results prep_att prepp kind before_qed after_qed stmt state = | 
| 17359 | 944 | let | 
| 21362 
3a2ab1dce297
simplified Proof.theorem(_i) interface -- removed target support;
 wenzelm parents: 
21274diff
changeset | 945 | val ((names, attss), propp) = | 
| 47815 | 946 | Attrib.map_specs (map (prep_att (context_of state))) stmt |> split_list |>> split_list; | 
| 17359 | 947 | |
| 18124 | 948 | fun after_qed' results = | 
| 18808 | 949 | local_results ((names ~~ attss) ~~ results) | 
| 17359 | 950 | #-> (fn res => tap (fn st => print_results (context_of st) ((kind, ""), res) : unit)) | 
| 18124 | 951 | #> after_qed results; | 
| 5820 | 952 | in | 
| 17359 | 953 | state | 
| 21362 
3a2ab1dce297
simplified Proof.theorem(_i) interface -- removed target support;
 wenzelm parents: 
21274diff
changeset | 954 | |> generic_goal prepp kind before_qed (after_qed', K I) propp | 
| 19900 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 wenzelm parents: 
19862diff
changeset | 955 | |> tap (Variable.warn_extra_tfrees (context_of state) o context_of) | 
| 5820 | 956 | end; | 
| 957 | ||
| 49866 | 958 | fun local_qeds arg = | 
| 959 | end_proof false arg | |
| 49859 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 960 | #> Seq.map_result (generic_qed Proof_Context.auto_bind_facts #-> | 
| 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 961 | (fn ((after_qed, _), results) => after_qed results)); | 
| 5820 | 962 | |
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 963 | fun local_qed arg = | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 964 | local_qeds (Position.none, arg) #> Seq.the_result finished_goal_error; | 
| 29383 | 965 | |
| 5820 | 966 | |
| 17359 | 967 | (* global goals *) | 
| 968 | ||
| 45327 | 969 | fun prepp_auto_fixes prepp args = | 
| 970 | prepp args #> | |
| 971 | (fn ((propss, a), ctxt) => ((propss, a), (fold o fold) Variable.auto_fixes propss ctxt)); | |
| 972 | ||
| 973 | fun global_goal prepp before_qed after_qed propp = | |
| 974 | init #> | |
| 975 | generic_goal (prepp_auto_fixes prepp) "" before_qed (K I, after_qed) propp; | |
| 17359 | 976 | |
| 42360 | 977 | val theorem = global_goal Proof_Context.bind_propp_schematic_i; | 
| 978 | val theorem_cmd = global_goal Proof_Context.bind_propp_schematic; | |
| 12065 | 979 | |
| 49866 | 980 | fun global_qeds arg = | 
| 981 | end_proof true arg | |
| 49859 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 982 | #> Seq.map_result (generic_qed (K I) #> (fn (((_, after_qed), results), state) => | 
| 
52da6a736c32
more informative error messages of initial/terminal proof methods;
 wenzelm parents: 
49848diff
changeset | 983 | after_qed results (context_of state))); | 
| 17112 | 984 | |
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 985 | fun global_qed arg = | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 986 | global_qeds (Position.none, arg) #> Seq.the_result finished_goal_error; | 
| 12959 | 987 | |
| 5820 | 988 | |
| 49907 | 989 | (* terminal proof steps *) | 
| 17112 | 990 | |
| 49890 
89eff795f757
skipped proofs appear as "bad" without counting as error;
 wenzelm parents: 
49889diff
changeset | 991 | local | 
| 
89eff795f757
skipped proofs appear as "bad" without counting as error;
 wenzelm parents: 
49889diff
changeset | 992 | |
| 49846 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
49748diff
changeset | 993 | fun terminal_proof qeds initial terminal = | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 994 | proof_results (SOME initial) #> Seq.maps_results (qeds (#2 (#2 initial), terminal)) | 
| 49863 
b5fb6e7f8d81
more informative errors for 'proof' and 'apply' steps;
 wenzelm parents: 
49861diff
changeset | 995 | #> Seq.the_result ""; | 
| 49848 
f222a054342e
more informative error of initial/terminal proof steps;
 wenzelm parents: 
49847diff
changeset | 996 | |
| 49890 
89eff795f757
skipped proofs appear as "bad" without counting as error;
 wenzelm parents: 
49889diff
changeset | 997 | in | 
| 
89eff795f757
skipped proofs appear as "bad" without counting as error;
 wenzelm parents: 
49889diff
changeset | 998 | |
| 29383 | 999 | fun local_terminal_proof (text, opt_text) = terminal_proof local_qeds text (opt_text, true); | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 1000 | val local_default_proof = local_terminal_proof ((Method.default_text, Position.no_range), NONE); | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 1001 | val local_immediate_proof = local_terminal_proof ((Method.this_text, Position.no_range), NONE); | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 1002 | val local_done_proof = terminal_proof local_qeds (Method.done_text, Position.no_range) (NONE, false); | 
| 17112 | 1003 | |
| 29383 | 1004 | fun global_terminal_proof (text, opt_text) = terminal_proof global_qeds text (opt_text, true); | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 1005 | val global_default_proof = global_terminal_proof ((Method.default_text, Position.no_range), NONE); | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 1006 | val global_immediate_proof = global_terminal_proof ((Method.this_text, Position.no_range), NONE); | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49867diff
changeset | 1007 | val global_done_proof = terminal_proof global_qeds (Method.done_text, Position.no_range) (NONE, false); | 
| 5820 | 1008 | |
| 49890 
89eff795f757
skipped proofs appear as "bad" without counting as error;
 wenzelm parents: 
49889diff
changeset | 1009 | end; | 
| 
89eff795f757
skipped proofs appear as "bad" without counting as error;
 wenzelm parents: 
49889diff
changeset | 1010 | |
| 6896 | 1011 | |
| 49907 | 1012 | (* skip proofs *) | 
| 1013 | ||
| 1014 | fun local_skip_proof int state = | |
| 1015 | local_terminal_proof ((Method.sorry_text int, Position.no_range), NONE) state before | |
| 51551 | 1016 | Skip_Proof.report (context_of state); | 
| 49907 | 1017 | |
| 1018 | fun global_skip_proof int state = | |
| 1019 | global_terminal_proof ((Method.sorry_text int, Position.no_range), NONE) state before | |
| 51551 | 1020 | Skip_Proof.report (context_of state); | 
| 49907 | 1021 | |
| 1022 | ||
| 17359 | 1023 | (* common goal statements *) | 
| 1024 | ||
| 1025 | local | |
| 1026 | ||
| 17859 | 1027 | fun gen_have prep_att prepp before_qed after_qed stmt int = | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56912diff
changeset | 1028 | local_goal (Proof_Display.print_results int (Position.thread_data ())) | 
| 46728 
85f8e3932712
display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
 wenzelm parents: 
46466diff
changeset | 1029 | prep_att prepp "have" before_qed after_qed stmt; | 
| 6896 | 1030 | |
| 17859 | 1031 | fun gen_show prep_att prepp before_qed after_qed stmt int state = | 
| 17359 | 1032 | let | 
| 32738 | 1033 | val testing = Unsynchronized.ref false; | 
| 1034 | val rule = Unsynchronized.ref (NONE: thm option); | |
| 17359 | 1035 | fun fail_msg ctxt = | 
| 50315 
cf9002ac1018
recovered error to finish proof (e.g. bad obtain export) from 223f18cfbb32;
 wenzelm parents: 
50201diff
changeset | 1036 | "Local statement fails to refine any pending goal" :: | 
| 33389 | 1037 | (case ! rule of NONE => [] | SOME th => [Proof_Display.string_of_rule ctxt "Failed" th]) | 
| 17359 | 1038 | |> cat_lines; | 
| 6896 | 1039 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56912diff
changeset | 1040 | val pos = Position.thread_data (); | 
| 17359 | 1041 | fun print_results ctxt res = | 
| 46728 
85f8e3932712
display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
 wenzelm parents: 
46466diff
changeset | 1042 | if ! testing then () | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56912diff
changeset | 1043 | else Proof_Display.print_results int pos ctxt res; | 
| 17359 | 1044 | fun print_rule ctxt th = | 
| 1045 | if ! testing then rule := SOME th | |
| 41181 
9240be8c8c69
show: display goal refinement rule as "state", to avoid odd Output.urgent_message and make its association with the goal more explicit;
 wenzelm parents: 
40960diff
changeset | 1046 | else if int then | 
| 56933 | 1047 | Proof_Display.string_of_rule ctxt "Successful" th | 
| 1048 | |> Markup.markup Markup.text_fold | |
| 1049 | |> Markup.markup Markup.state | |
| 1050 | |> writeln | |
| 17359 | 1051 | else (); | 
| 1052 | val test_proof = | |
| 50315 
cf9002ac1018
recovered error to finish proof (e.g. bad obtain export) from 223f18cfbb32;
 wenzelm parents: 
50201diff
changeset | 1053 | local_skip_proof true | 
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
39232diff
changeset | 1054 | |> Unsynchronized.setmp testing true | 
| 42042 | 1055 | |> Exn.interruptible_capture; | 
| 6896 | 1056 | |
| 18124 | 1057 | fun after_qed' results = | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19475diff
changeset | 1058 | refine_goals print_rule (context_of state) (flat results) | 
| 29383 | 1059 | #> check_result "Failed to refine any pending goal" | 
| 1060 | #> after_qed results; | |
| 17359 | 1061 | in | 
| 1062 | state | |
| 17859 | 1063 | |> local_goal print_results prep_att prepp "show" before_qed after_qed' stmt | 
| 21565 | 1064 | |> int ? (fn goal_state => | 
| 49907 | 1065 | (case test_proof (map_context (Context_Position.set_visible false) goal_state) of | 
| 50315 
cf9002ac1018
recovered error to finish proof (e.g. bad obtain export) from 223f18cfbb32;
 wenzelm parents: 
50201diff
changeset | 1066 | Exn.Res _ => goal_state | 
| 42042 | 1067 | | Exn.Exn exn => raise Exn.EXCEPTIONS ([exn, ERROR (fail_msg (context_of goal_state))]))) | 
| 17359 | 1068 | end; | 
| 1069 | ||
| 1070 | in | |
| 1071 | ||
| 42360 | 1072 | val have = gen_have (K I) Proof_Context.bind_propp_i; | 
| 47815 | 1073 | val have_cmd = gen_have Attrib.attribute_cmd Proof_Context.bind_propp; | 
| 42360 | 1074 | val show = gen_show (K I) Proof_Context.bind_propp_i; | 
| 47815 | 1075 | val show_cmd = gen_show Attrib.attribute_cmd Proof_Context.bind_propp; | 
| 6896 | 1076 | |
| 5820 | 1077 | end; | 
| 17359 | 1078 | |
| 28410 | 1079 | |
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1080 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1081 | (** future proofs **) | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1082 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1083 | (* relevant proof states *) | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1084 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1085 | fun schematic_goal state = | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1086 |   let val (_, (_, {statement = (_, _, prop), ...})) = find_goal state
 | 
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51551diff
changeset | 1087 | in Goal.is_schematic prop end; | 
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1088 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1089 | fun is_relevant state = | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1090 | (case try find_goal state of | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1091 | NONE => true | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1092 |   | SOME (_, (_, {statement = (_, _, prop), goal, ...})) =>
 | 
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51551diff
changeset | 1093 | Goal.is_schematic prop orelse not (Logic.protect prop aconv Thm.concl_of goal)); | 
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1094 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1095 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1096 | (* full proofs *) | 
| 28410 | 1097 | |
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1098 | local | 
| 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1099 | |
| 49912 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1100 | structure Result = Proof_Data | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1101 | ( | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1102 | type T = thm option; | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1103 | val empty = NONE; | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1104 | fun init _ = empty; | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1105 | ); | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1106 | |
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1107 | fun the_result ctxt = | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1108 | (case Result.get ctxt of | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1109 | NONE => error "No result of forked proof" | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1110 | | SOME th => th); | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1111 | |
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1112 | val set_result = Result.put o SOME; | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1113 | val reset_result = Result.put NONE; | 
| 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1114 | |
| 51318 | 1115 | in | 
| 1116 | ||
| 1117 | fun future_proof fork_proof state = | |
| 28410 | 1118 | let | 
| 29381 
45d77aeb1608
future_terminal_proof: no fork for interactive mode, assert_backward;
 wenzelm parents: 
29364diff
changeset | 1119 | val _ = assert_backward state; | 
| 28410 | 1120 | val (goal_ctxt, (_, goal)) = find_goal state; | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 1121 |     val {statement as (kind, _, prop), using, goal, before_qed, after_qed} = goal;
 | 
| 47431 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 1122 | val goal_tfrees = | 
| 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 1123 | fold Term.add_tfrees | 
| 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 1124 | (prop :: map Thm.term_of (Assumption.all_assms_of goal_ctxt)) []; | 
| 28410 | 1125 | |
| 29383 | 1126 | val _ = is_relevant state andalso error "Cannot fork relevant proof"; | 
| 1127 | ||
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1128 | val prop' = Logic.protect prop; | 
| 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1129 | val statement' = (kind, [[], [prop']], prop'); | 
| 52456 | 1130 | val goal' = Thm.adjust_maxidx_thm (Thm.maxidx_of goal) (Goal.protect (Thm.nprems_of goal) goal); | 
| 49912 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1131 | val after_qed' = (fn [[th]] => map_context (set_result th), fn [[th]] => set_result th); | 
| 28410 | 1132 | |
| 28973 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28627diff
changeset | 1133 | val result_ctxt = | 
| 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28627diff
changeset | 1134 | state | 
| 49912 
c6307ee2665d
more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
 wenzelm parents: 
49909diff
changeset | 1135 | |> map_context reset_result | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 1136 | |> map_goal I (K (statement', using, goal', before_qed, after_qed')) | 
| 47431 
d9dd4a9f18fc
more precise declaration of goal_tfrees in forked proof state;
 wenzelm parents: 
47416diff
changeset | 1137 | (fold (Variable.declare_typ o TFree) goal_tfrees) | 
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1138 | |> fork_proof; | 
| 28973 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28627diff
changeset | 1139 | |
| 51318 | 1140 | val future_thm = Future.map (the_result o snd) result_ctxt; | 
| 32062 
457f5bcd3d76
Proof.future_proof: declare all assumptions as well;
 wenzelm parents: 
32061diff
changeset | 1141 | val finished_goal = Goal.future_result goal_ctxt future_thm prop'; | 
| 28973 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28627diff
changeset | 1142 | val state' = | 
| 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28627diff
changeset | 1143 | state | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58837diff
changeset | 1144 | |> map_goal I (K (statement, using, finished_goal, NONE, after_qed)) I; | 
| 51318 | 1145 | in (Future.map fst result_ctxt, state') end; | 
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1146 | |
| 17359 | 1147 | end; | 
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1148 | |
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1149 | |
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1150 | (* terminal proofs *) (* FIXME avoid toplevel imitation -- include in PIDE/document *) | 
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1151 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1152 | local | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1153 | |
| 51605 
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
 wenzelm parents: 
51584diff
changeset | 1154 | fun future_terminal_proof proof1 proof2 done int state = | 
| 53189 
ee8b8dafef0e
discontinued parallel_subproofs_saturation and related internal counters (superseded by parallel_subproofs_threshold and timing information);
 wenzelm parents: 
52732diff
changeset | 1155 | if Goal.future_enabled 3 andalso not (is_relevant state) then | 
| 51318 | 1156 | state |> future_proof (fn state' => | 
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1157 | let | 
| 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1158 | val pos = Position.thread_data (); | 
| 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1159 | val props = Markup.command_timing :: (Markup.nameN, "by") :: Position.properties_of pos; | 
| 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1160 | in | 
| 53192 | 1161 |         Execution.fork {name = "Proof.future_terminal_proof", pos = pos, pri = ~1}
 | 
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1162 | (fn () => ((), Timing.protocol props proof2 state')) | 
| 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51606diff
changeset | 1163 | end) |> snd |> done | 
| 51318 | 1164 | else proof1 state; | 
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1165 | |
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1166 | in | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1167 | |
| 51318 | 1168 | fun local_future_terminal_proof meths = | 
| 51605 
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
 wenzelm parents: 
51584diff
changeset | 1169 | future_terminal_proof | 
| 51318 | 1170 | (local_terminal_proof meths) | 
| 1171 | (local_terminal_proof meths #> context_of) local_done_proof; | |
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1172 | |
| 51318 | 1173 | fun global_future_terminal_proof meths = | 
| 51605 
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
 wenzelm parents: 
51584diff
changeset | 1174 | future_terminal_proof | 
| 51318 | 1175 | (global_terminal_proof meths) | 
| 1176 | (global_terminal_proof meths) global_done_proof; | |
| 29350 | 1177 | |
| 29346 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1178 | end; | 
| 
fe6843aa4f5f
more precise is_relevant: requires original main goal, not initial goal state;
 wenzelm parents: 
29090diff
changeset | 1179 | |
| 29385 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1180 | end; | 
| 
c96b91bab2e6
future_proof: refined version covers local_future_proof and global_future_proof;
 wenzelm parents: 
29383diff
changeset | 1181 |