| author | ballarin |
| Fri, 16 Sep 2005 14:46:31 +0200 | |
| changeset 17437 | 9deaf32c83be |
| parent 17359 | 543735c6f424 |
| child 17450 | f2e0a211c4fc |
| permissions | -rw-r--r-- |
| 5820 | 1 |
(* Title: Pure/Isar/proof.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
| 15757 | 5 |
The Isar/VM proof language interpreter. |
| 5820 | 6 |
*) |
7 |
||
8 |
signature PROOF = |
|
9 |
sig |
|
| 17359 | 10 |
type context = Context.proof |
11 |
type method = Method.method |
|
| 5820 | 12 |
type state |
13 |
exception STATE of string * state |
|
| 17359 | 14 |
val init: context -> state |
15 |
val level: state -> int |
|
16 |
val assert_bottom: bool -> state -> state |
|
| 5820 | 17 |
val context_of: state -> context |
18 |
val theory_of: state -> theory |
|
| 16450 | 19 |
val sign_of: state -> theory (*obsolete*) |
| 13377 | 20 |
val map_context: (context -> context) -> state -> state |
| 7924 | 21 |
val warn_extra_tfrees: state -> state -> state |
| 17359 | 22 |
val put_thms: string * thm list option -> state -> state |
| 6091 | 23 |
val the_facts: state -> thm list |
| 9469 | 24 |
val the_fact: state -> thm |
| 6891 | 25 |
val assert_forward: state -> state |
| 17359 | 26 |
val assert_chain: state -> state |
| 9469 | 27 |
val assert_forward_or_chain: state -> state |
| 5820 | 28 |
val assert_backward: state -> state |
| 8206 | 29 |
val assert_no_chain: state -> state |
| 5820 | 30 |
val enter_forward: state -> state |
| 17359 | 31 |
val get_goal: state -> context * (thm list * thm) |
| 16539 | 32 |
val show_main_goal: bool ref |
| 6529 | 33 |
val verbose: bool ref |
| 12423 | 34 |
val pretty_state: int -> state -> Pretty.T list |
| 10360 | 35 |
val pretty_goals: bool -> state -> Pretty.T list |
| 17112 | 36 |
val refine: Method.text -> state -> state Seq.seq |
37 |
val refine_end: Method.text -> state -> state Seq.seq |
|
| 17359 | 38 |
val refine_goals: (context -> thm -> unit) -> context -> thm list -> state -> state Seq.seq |
| 5936 | 39 |
val match_bind: (string list * string) list -> state -> state |
40 |
val match_bind_i: (term list * term) list -> state -> state |
|
|
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
41 |
val let_bind: (string list * string) list -> state -> state |
|
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
42 |
val let_bind_i: (term list * term) list -> state -> state |
| 17359 | 43 |
val fix: (string list * string option) list -> state -> state |
44 |
val fix_i: (string list * typ option) list -> state -> state |
|
45 |
val assm: ProofContext.exporter -> |
|
46 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
|
47 |
state -> state |
|
48 |
val assm_i: ProofContext.exporter -> |
|
49 |
((string * context attribute list) * (term * (term list * term list)) list) list |
|
50 |
-> state -> state |
|
51 |
val assume: ((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
|
52 |
state -> state |
|
53 |
val assume_i: ((string * context attribute list) * (term * (term list * term list)) list) list |
|
54 |
-> state -> state |
|
55 |
val presume: ((string * Attrib.src list) * (string * (string list * string list)) list) list |
|
56 |
-> state -> state |
|
57 |
val presume_i: ((string * context attribute list) * (term * (term list * term list)) list) list |
|
58 |
-> state -> state |
|
59 |
val def: string * Attrib.src list -> string * (string * string list) -> state -> state |
|
60 |
val def_i: string * context attribute list -> string * (term * term list) -> state -> state |
|
61 |
val chain: state -> state |
|
62 |
val chain_facts: thm list -> state -> state |
|
63 |
val get_thmss: state -> (thmref * Attrib.src list) list -> thm list |
|
| 14564 | 64 |
val simple_note_thms: string -> thm list -> state -> state |
| 17359 | 65 |
val note_thmss: ((string * Attrib.src list) * |
| 17112 | 66 |
(thmref * Attrib.src list) list) list -> state -> state |
| 17359 | 67 |
val note_thmss_i: ((string * context attribute list) * |
| 9196 | 68 |
(thm list * context attribute list) list) list -> state -> state |
| 17112 | 69 |
val from_thmss: ((thmref * Attrib.src list) list) list -> state -> state |
70 |
val from_thmss_i: ((thm list * context attribute list) list) list -> state -> state |
|
71 |
val with_thmss: ((thmref * Attrib.src list) list) list -> state -> state |
|
72 |
val with_thmss_i: ((thm list * context attribute list) list) list -> state -> state |
|
73 |
val using_thmss: ((thmref * Attrib.src list) list) list -> state -> state |
|
| 12930 | 74 |
val using_thmss_i: ((thm list * context attribute list) list) list -> state -> state |
| 17112 | 75 |
val invoke_case: string * string option list * Attrib.src list -> state -> state |
76 |
val invoke_case_i: string * string option list * context attribute list -> state -> state |
|
| 17359 | 77 |
val begin_block: state -> state |
78 |
val next_block: state -> state |
|
79 |
val end_block: state -> state Seq.seq |
|
| 17112 | 80 |
val proof: Method.text option -> state -> state Seq.seq |
81 |
val defer: int option -> state -> state Seq.seq |
|
82 |
val prefer: int -> state -> state Seq.seq |
|
83 |
val apply: Method.text -> state -> state Seq.seq |
|
84 |
val apply_end: Method.text -> state -> state Seq.seq |
|
| 17359 | 85 |
val goal_names: string option -> string -> string list -> string |
| 17437 | 86 |
val generic_goal: |
87 |
(context * 'a -> context * (term list list * (context -> context))) -> |
|
88 |
string -> |
|
89 |
(context * thm list -> thm list list -> state -> state Seq.seq) * |
|
90 |
(context * thm list -> thm list list -> theory -> theory) -> |
|
91 |
'a -> state -> state |
|
| 17359 | 92 |
val local_goal: (context -> ((string * string) * (string * thm list) list) -> unit) -> |
93 |
(theory -> 'a -> context attribute) -> |
|
94 |
(context * 'b list -> context * (term list list * (context -> context))) -> |
|
95 |
string -> (context * thm list -> thm list list -> state -> state Seq.seq) -> |
|
96 |
((string * 'a list) * 'b) list -> state -> state |
|
97 |
val local_qed: Method.text option * bool -> state -> state Seq.seq |
|
| 17437 | 98 |
val auto_fix: context * (term list list * 'a) -> |
99 |
context * (term list list * 'a) |
|
| 17359 | 100 |
val global_goal: (context -> (string * string) * (string * thm list) list -> unit) -> |
101 |
(theory -> 'a -> theory attribute) -> |
|
102 |
(context * 'b list -> context * (term list list * (context -> context))) -> |
|
103 |
string -> (context * thm list -> thm list list -> theory -> theory) -> string option -> |
|
104 |
string * 'a list -> ((string * 'a list) * 'b) list -> context -> state |
|
105 |
val global_qed: Method.text option * bool -> state -> theory * context |
|
106 |
val local_terminal_proof: Method.text * Method.text option -> state -> state Seq.seq |
|
107 |
val local_default_proof: state -> state Seq.seq |
|
108 |
val local_immediate_proof: state -> state Seq.seq |
|
109 |
val local_done_proof: state -> state Seq.seq |
|
110 |
val local_skip_proof: bool -> state -> state Seq.seq |
|
111 |
val global_terminal_proof: Method.text * Method.text option -> state -> theory * context |
|
112 |
val global_default_proof: state -> theory * context |
|
113 |
val global_immediate_proof: state -> theory * context |
|
114 |
val global_done_proof: state -> theory * context |
|
115 |
val global_skip_proof: bool -> state -> theory * context |
|
116 |
val have: (context * thm list -> thm list list -> state -> state Seq.seq) -> |
|
117 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
|
118 |
bool -> state -> state |
|
119 |
val have_i: (context * thm list -> thm list list -> state -> state Seq.seq) -> |
|
120 |
((string * context attribute list) * (term * (term list * term list)) list) list -> |
|
121 |
bool -> state -> state |
|
122 |
val show: (context * thm list -> thm list list -> state -> state Seq.seq) -> |
|
123 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
|
124 |
bool -> state -> state |
|
125 |
val show_i: (context * thm list -> thm list list -> state -> state Seq.seq) -> |
|
126 |
((string * context attribute list) * (term * (term list * term list)) list) list -> |
|
127 |
bool -> state -> state |
|
128 |
val theorem: string -> (context * thm list -> thm list list -> theory -> theory) -> |
|
129 |
string option -> string * Attrib.src list -> |
|
130 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
|
131 |
context -> state |
|
132 |
val theorem_i: string -> (context * thm list -> thm list list -> theory -> theory) -> |
|
133 |
string option -> string * theory attribute list -> |
|
134 |
((string * theory attribute list) * (term * (term list * term list)) list) list -> |
|
135 |
context -> state |
|
| 5820 | 136 |
end; |
137 |
||
| 13377 | 138 |
structure Proof: PROOF = |
| 5820 | 139 |
struct |
140 |
||
| 16813 | 141 |
type context = ProofContext.context; |
| 17112 | 142 |
type method = Method.method; |
| 16813 | 143 |
|
| 5820 | 144 |
|
145 |
(** proof state **) |
|
146 |
||
| 17359 | 147 |
(* datatype state *) |
| 5820 | 148 |
|
| 17112 | 149 |
datatype mode = Forward | Chain | Backward; |
| 5820 | 150 |
|
| 17359 | 151 |
datatype state = |
152 |
State of node Stack.T |
|
153 |
and node = |
|
| 7176 | 154 |
Node of |
155 |
{context: context,
|
|
156 |
facts: thm list option, |
|
157 |
mode: mode, |
|
| 17359 | 158 |
goal: goal option} |
159 |
and goal = |
|
160 |
Goal of |
|
161 |
{statement: string * term list list, (*goal kind and statement*)
|
|
162 |
using: thm list, (*goal facts*) |
|
163 |
goal: thm, (*subgoals ==> statement*) |
|
164 |
after_qed: (* FIXME results only *) |
|
165 |
(context * thm list -> thm list list -> state -> state Seq.seq) * |
|
166 |
(context * thm list -> thm list list -> theory -> theory)}; |
|
167 |
||
168 |
exception STATE of string * state; |
|
169 |
||
170 |
fun make_goal (statement, using, goal, after_qed) = |
|
171 |
Goal {statement = statement, using = using, goal = goal, after_qed = after_qed};
|
|
| 5820 | 172 |
|
| 7176 | 173 |
fun make_node (context, facts, mode, goal) = |
174 |
Node {context = context, facts = facts, mode = mode, goal = goal};
|
|
175 |
||
| 17359 | 176 |
fun map_node f (Node {context, facts, mode, goal}) =
|
177 |
make_node (f (context, facts, mode, goal)); |
|
| 5820 | 178 |
|
| 17359 | 179 |
fun init ctxt = State (Stack.init (make_node (ctxt, NONE, Forward, NONE))); |
| 5820 | 180 |
|
| 17359 | 181 |
fun current (State st) = Stack.top st |> (fn Node node => node); |
182 |
fun map_current f (State st) = State (Stack.map (map_node f) st); |
|
| 12045 | 183 |
|
| 5820 | 184 |
|
185 |
||
186 |
(** basic proof state operations **) |
|
187 |
||
| 17359 | 188 |
(* block structure *) |
189 |
||
190 |
fun open_block (State st) = State (Stack.push st); |
|
191 |
||
192 |
fun close_block (state as State st) = State (Stack.pop st) |
|
193 |
handle Empty => raise STATE ("Unbalanced block parentheses", state);
|
|
194 |
||
195 |
fun level (State st) = Stack.level st; |
|
196 |
||
197 |
fun assert_bottom b state = |
|
198 |
let val b' = (level state <= 2) in |
|
199 |
if b andalso not b' then raise STATE ("Not at bottom of proof!", state)
|
|
200 |
else if not b andalso b' then raise STATE ("Already at bottom of proof!", state)
|
|
201 |
else state |
|
202 |
end; |
|
203 |
||
204 |
||
| 5820 | 205 |
(* context *) |
206 |
||
| 17359 | 207 |
val context_of = #context o current; |
| 5820 | 208 |
val theory_of = ProofContext.theory_of o context_of; |
| 16450 | 209 |
val sign_of = theory_of; |
| 5820 | 210 |
|
| 17359 | 211 |
fun map_context f = |
212 |
map_current (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal)); |
|
| 5820 | 213 |
|
| 17359 | 214 |
fun map_context_result f state = |
215 |
f (context_of state) |> swap ||> (fn ctxt => map_context (K ctxt) state); |
|
| 5820 | 216 |
|
| 7924 | 217 |
val warn_extra_tfrees = map_context o ProofContext.warn_extra_tfrees o context_of; |
| 10809 | 218 |
val add_binds_i = map_context o ProofContext.add_binds_i; |
| 6091 | 219 |
val put_thms = map_context o ProofContext.put_thms; |
| 8374 | 220 |
val get_case = ProofContext.get_case o context_of; |
| 5820 | 221 |
|
222 |
||
223 |
(* facts *) |
|
224 |
||
| 17359 | 225 |
val get_facts = #facts o current; |
226 |
||
227 |
fun the_facts state = |
|
228 |
(case get_facts state of SOME facts => facts |
|
229 |
| NONE => raise STATE ("No current facts available", state));
|
|
| 5820 | 230 |
|
| 9469 | 231 |
fun the_fact state = |
| 17359 | 232 |
(case the_facts state of [thm] => thm |
233 |
| _ => raise STATE ("Single theorem expected", state));
|
|
| 7605 | 234 |
|
| 17359 | 235 |
fun put_facts facts = |
236 |
map_current (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) |
|
237 |
#> put_thms (AutoBind.thisN, facts); |
|
| 5820 | 238 |
|
| 17359 | 239 |
fun these_factss more_facts (named_factss, state) = |
240 |
(named_factss, state |> put_facts (SOME (List.concat (map snd named_factss) @ more_facts))); |
|
| 5820 | 241 |
|
| 17359 | 242 |
fun export_facts inner outer = |
243 |
(case get_facts inner of |
|
244 |
NONE => Seq.single (put_facts NONE outer) |
|
245 |
| SOME thms => |
|
246 |
thms |
|
247 |
|> Seq.map_list (ProofContext.export false (context_of inner) (context_of outer)) |
|
248 |
|> Seq.map (fn ths => put_facts (SOME ths) outer)); |
|
| 5820 | 249 |
|
250 |
||
251 |
(* mode *) |
|
252 |
||
| 17359 | 253 |
val get_mode = #mode o current; |
| 5820 | 254 |
fun put_mode mode = map_current (fn (ctxt, facts, _, goal) => (ctxt, facts, mode, goal)); |
255 |
||
| 17359 | 256 |
val mode_name = (fn Forward => "state" | Chain => "chain" | Backward => "prove"); |
| 5820 | 257 |
|
258 |
fun assert_mode pred state = |
|
259 |
let val mode = get_mode state in |
|
260 |
if pred mode then state |
|
| 12010 | 261 |
else raise STATE ("Illegal application of proof command in "
|
262 |
^ quote (mode_name mode) ^ " mode", state) |
|
| 5820 | 263 |
end; |
264 |
||
265 |
val assert_forward = assert_mode (equal Forward); |
|
| 17359 | 266 |
val assert_chain = assert_mode (equal Chain); |
| 17112 | 267 |
val assert_forward_or_chain = assert_mode (equal Forward orf equal Chain); |
| 5820 | 268 |
val assert_backward = assert_mode (equal Backward); |
| 17112 | 269 |
val assert_no_chain = assert_mode (not_equal Chain); |
| 5820 | 270 |
|
| 17359 | 271 |
val enter_forward = put_mode Forward; |
272 |
val enter_chain = put_mode Chain; |
|
273 |
val enter_backward = put_mode Backward; |
|
| 5820 | 274 |
|
| 17359 | 275 |
|
276 |
(* current goal *) |
|
277 |
||
278 |
fun current_goal state = |
|
279 |
(case current state of |
|
280 |
{context, goal = SOME (Goal goal), ...} => (context, goal)
|
|
281 |
| _ => raise STATE ("No current goal!", state));
|
|
| 5820 | 282 |
|
| 17359 | 283 |
fun assert_current_goal g state = |
284 |
let val g' = can current_goal state in |
|
285 |
if g andalso not g' then raise STATE ("No goal in this block!", state)
|
|
286 |
else if not g andalso g' then raise STATE ("Goal present in this block!", state)
|
|
287 |
else state |
|
288 |
end; |
|
| 6776 | 289 |
|
| 17359 | 290 |
fun put_goal goal = map_current (fn (ctxt, using, mode, _) => (ctxt, using, mode, goal)); |
291 |
||
292 |
||
293 |
(* nested goal *) |
|
| 5820 | 294 |
|
| 17359 | 295 |
fun map_goal f g (State (Node {context, facts, mode, goal = SOME goal}, nodes)) =
|
296 |
let |
|
297 |
val Goal {statement, using, goal, after_qed} = goal;
|
|
298 |
val goal' = make_goal (g (statement, using, goal, after_qed)); |
|
299 |
in State (make_node (f context, facts, mode, SOME goal'), nodes) end |
|
300 |
| map_goal f g (State (nd, node :: nodes)) = |
|
301 |
let val State (node', nodes') = map_goal f g (State (node, nodes)) |
|
302 |
in map_context f (State (nd, node' :: nodes')) end |
|
303 |
| map_goal _ _ state = state; |
|
| 5820 | 304 |
|
| 17359 | 305 |
fun using_facts using = |
306 |
map_goal I (fn (statement, _, goal, after_qed) => (statement, using, goal, after_qed)); |
|
307 |
||
308 |
local |
|
309 |
fun find i state = |
|
310 |
(case try current_goal state of |
|
311 |
SOME (ctxt, goal) => (ctxt, (i, goal)) |
|
312 |
| NONE => find (i + 1) (close_block state |
|
313 |
handle STATE _ => raise STATE ("No goal present", state)));
|
|
314 |
in val find_goal = find 0 end; |
|
315 |
||
316 |
fun get_goal state = |
|
317 |
let val (ctxt, (_, {using, goal, ...})) = find_goal state
|
|
318 |
in (ctxt, (using, goal)) end; |
|
| 5820 | 319 |
|
320 |
||
321 |
||
| 12423 | 322 |
(** pretty_state **) |
| 5820 | 323 |
|
| 13869 | 324 |
val show_main_goal = ref false; |
| 16539 | 325 |
val verbose = ProofContext.verbose; |
| 13698 | 326 |
|
| 14876 | 327 |
val pretty_goals_local = Display.pretty_goals_aux o ProofContext.pp; |
| 12085 | 328 |
|
| 15531 | 329 |
fun pretty_facts _ _ NONE = [] |
330 |
| pretty_facts s ctxt (SOME ths) = |
|
| 12055 | 331 |
[Pretty.big_list (s ^ "this:") (map (ProofContext.pretty_thm ctxt) ths), Pretty.str ""]; |
| 6756 | 332 |
|
| 17359 | 333 |
fun pretty_state nr state = |
| 5820 | 334 |
let |
| 17359 | 335 |
val {context, facts, mode, goal = _} = current state;
|
336 |
val ref (_, _, bg) = Display.current_goals_markers; |
|
| 5820 | 337 |
|
338 |
fun levels_up 0 = "" |
|
| 17359 | 339 |
| levels_up 1 = "1 level up" |
340 |
| levels_up i = string_of_int i ^ " levels up"; |
|
| 5820 | 341 |
|
| 11556 | 342 |
fun subgoals 0 = "" |
| 17359 | 343 |
| subgoals 1 = "1 subgoal" |
344 |
| subgoals n = string_of_int n ^ " subgoals"; |
|
345 |
||
346 |
fun description strs = |
|
347 |
(case filter_out (equal "") strs of [] => "" |
|
348 |
| strs' => enclose " (" ")" (commas strs'));
|
|
| 12146 | 349 |
|
| 17359 | 350 |
fun prt_goal (SOME (ctxt, (i, {statement = (kind, _), using, goal, after_qed = _}))) =
|
351 |
pretty_facts "using " ctxt |
|
352 |
(if mode <> Backward orelse null using then NONE else SOME using) @ |
|
353 |
[Pretty.str ("goal" ^ description [kind, levels_up (i div 2),
|
|
354 |
subgoals (Thm.nprems_of goal)] ^ ":")] @ |
|
355 |
pretty_goals_local ctxt bg (true, ! show_main_goal) (! Display.goals_limit) goal |
|
356 |
| prt_goal NONE = []; |
|
| 6848 | 357 |
|
| 17359 | 358 |
val prt_ctxt = |
359 |
if ! verbose orelse mode = Forward then ProofContext.pretty_context context |
|
360 |
else if mode = Backward then ProofContext.pretty_asms context |
|
| 7575 | 361 |
else []; |
| 17359 | 362 |
in |
363 |
[Pretty.str ("proof (" ^ mode_name mode ^ "): step " ^ string_of_int nr ^
|
|
364 |
(if ! verbose then ", depth " ^ string_of_int (level state div 2 - 1) else "")), |
|
365 |
Pretty.str ""] @ |
|
366 |
(if null prt_ctxt then [] else prt_ctxt @ [Pretty.str ""]) @ |
|
367 |
(if ! verbose orelse mode = Forward then |
|
368 |
pretty_facts "" context facts @ prt_goal (try find_goal state) |
|
369 |
else if mode = Chain then pretty_facts "picking " context facts |
|
370 |
else prt_goal (try find_goal state)) |
|
371 |
end; |
|
| 5820 | 372 |
|
| 12085 | 373 |
fun pretty_goals main state = |
| 17359 | 374 |
let val (ctxt, (_, {goal, ...})) = find_goal state
|
375 |
in pretty_goals_local ctxt "" (false, main) (! Display.goals_limit) goal end; |
|
| 10320 | 376 |
|
| 5820 | 377 |
|
378 |
||
379 |
(** proof steps **) |
|
380 |
||
| 17359 | 381 |
(* refine via method *) |
| 5820 | 382 |
|
| 8234 | 383 |
local |
384 |
||
| 16146 | 385 |
fun goalN i = "goal" ^ string_of_int i; |
386 |
fun goals st = map goalN (1 upto Thm.nprems_of st); |
|
387 |
||
388 |
fun no_goal_cases st = map (rpair NONE) (goals st); |
|
389 |
||
390 |
fun goal_cases st = |
|
| 16450 | 391 |
RuleCases.make true NONE (Thm.theory_of_thm st, Thm.prop_of st) (goals st); |
| 16146 | 392 |
|
| 16450 | 393 |
fun check_theory thy state = |
394 |
if subthy (thy, theory_of state) then state |
|
395 |
else raise STATE ("Bad theory of method result: " ^ Context.str_of_thy thy, state);
|
|
| 5820 | 396 |
|
| 17112 | 397 |
fun apply_method current_context meth_fun state = |
| 6848 | 398 |
let |
| 17359 | 399 |
val (goal_ctxt, (_, {statement, using, goal, after_qed})) = find_goal state;
|
| 17112 | 400 |
val meth = meth_fun (if current_context then context_of state else goal_ctxt); |
| 16146 | 401 |
in |
| 17359 | 402 |
Method.apply meth using goal |> Seq.map (fn (goal', meth_cases) => |
| 6848 | 403 |
state |
| 17359 | 404 |
|> check_theory (Thm.theory_of_thm goal') |
| 16146 | 405 |
|> map_goal |
| 17359 | 406 |
(ProofContext.add_cases (no_goal_cases goal @ goal_cases goal' @ meth_cases)) |
407 |
(K (statement, using, goal', after_qed))) |
|
| 16146 | 408 |
end; |
| 5820 | 409 |
|
| 17112 | 410 |
fun apply_text cc text state = |
411 |
let |
|
412 |
val thy = theory_of state; |
|
413 |
||
414 |
fun eval (Method.Basic m) = apply_method cc m |
|
415 |
| eval (Method.Source src) = apply_method cc (Method.method thy src) |
|
416 |
| eval (Method.Then txts) = Seq.EVERY (map eval txts) |
|
417 |
| eval (Method.Orelse txts) = Seq.FIRST (map eval txts) |
|
418 |
| eval (Method.Try txt) = Seq.TRY (eval txt) |
|
419 |
| eval (Method.Repeat1 txt) = Seq.REPEAT1 (eval txt); |
|
420 |
in eval text state end; |
|
421 |
||
| 8234 | 422 |
in |
423 |
||
| 17112 | 424 |
val refine = apply_text true; |
425 |
val refine_end = apply_text false; |
|
| 8234 | 426 |
|
427 |
end; |
|
428 |
||
| 5820 | 429 |
|
| 17359 | 430 |
(* refine via sub-proof *) |
431 |
||
432 |
local |
|
| 9469 | 433 |
|
| 11816 | 434 |
fun refine_tac rule = |
|
12703
af5fec8a418f
refine_tac: Tactic.norm_hhf_tac before trying rule;
wenzelm
parents:
12698
diff
changeset
|
435 |
Tactic.norm_hhf_tac THEN' Tactic.rtac rule THEN_ALL_NEW (SUBGOAL (fn (goal, i) => |
| 11816 | 436 |
if can Logic.dest_goal (Logic.strip_assums_concl goal) then |
437 |
Tactic.etac Drule.triv_goal i |
|
438 |
else all_tac)); |
|
439 |
||
| 17359 | 440 |
fun refine_goal print_rule inner raw_rule state = |
441 |
let val (outer, (_, {statement, using, goal, after_qed})) = find_goal state in
|
|
442 |
raw_rule |
|
443 |
|> ProofContext.export true inner outer |
|
444 |
|> Seq.maps (fn rule => |
|
445 |
(print_rule outer rule; |
|
446 |
goal |
|
447 |
|> FINDGOAL (refine_tac rule) |
|
448 |
|> Seq.map (fn goal' => map_goal I (K (statement, using, goal', after_qed)) state))) |
|
449 |
end; |
|
| 6932 | 450 |
|
| 17359 | 451 |
in |
| 12146 | 452 |
|
| 17359 | 453 |
fun refine_goals print_rule inner raw_rules = |
454 |
Seq.EVERY (map (refine_goal print_rule inner) raw_rules); |
|
455 |
||
456 |
end; |
|
| 6932 | 457 |
|
458 |
||
| 17359 | 459 |
(* conclude_goal *) |
| 6932 | 460 |
|
| 17359 | 461 |
fun conclude_goal state props goal = |
| 5820 | 462 |
let |
463 |
val ctxt = context_of state; |
|
464 |
fun err msg = raise STATE (msg, state); |
|
465 |
||
| 17359 | 466 |
val ngoals = Thm.nprems_of goal; |
467 |
val _ = conditional (ngoals > 0) (fn () => |
|
468 |
err (Pretty.string_of (Pretty.chunks |
|
469 |
(pretty_goals_local ctxt "" (true, ! show_main_goal) (! Display.goals_limit) goal @ |
|
| 12085 | 470 |
[Pretty.str (string_of_int ngoals ^ " unsolved goal(s)!")])))); |
| 5820 | 471 |
|
| 17359 | 472 |
val {hyps, thy, ...} = Thm.rep_thm goal;
|
473 |
val casms = List.concat (map #1 (ProofContext.assumptions_of ctxt)); |
|
474 |
val bad_hyps = fold (remove (fn (asm, hyp) => Thm.term_of asm aconv hyp)) casms hyps; |
|
| 12146 | 475 |
|
| 17359 | 476 |
val th = goal RS Drule.rev_triv_goal; |
477 |
val ths = Drule.conj_elim_precise (length props) th |
|
| 12146 | 478 |
handle THM _ => err "Main goal structure corrupted"; |
| 5820 | 479 |
in |
| 12146 | 480 |
conditional (not (null bad_hyps)) (fn () => err ("Additional hypotheses:\n" ^
|
481 |
cat_lines (map (ProofContext.string_of_term ctxt) bad_hyps))); |
|
| 17359 | 482 |
conditional (exists (not o Pattern.matches thy) (props ~~ map Thm.prop_of ths)) (fn () => |
| 12146 | 483 |
err ("Proved a different theorem: " ^ Pretty.string_of (ProofContext.pretty_thm ctxt th)));
|
484 |
ths |
|
| 5820 | 485 |
end; |
486 |
||
487 |
||
488 |
||
489 |
(*** structured proof commands ***) |
|
490 |
||
| 17112 | 491 |
(** context elements **) |
| 5820 | 492 |
|
| 17112 | 493 |
(* bindings *) |
| 5820 | 494 |
|
| 16813 | 495 |
local |
496 |
||
| 17359 | 497 |
fun gen_bind bind args state = |
| 5820 | 498 |
state |
499 |
|> assert_forward |
|
| 17359 | 500 |
|> map_context (bind args) |
501 |
|> put_facts NONE; |
|
| 5820 | 502 |
|
| 16813 | 503 |
in |
504 |
||
|
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
505 |
val match_bind = gen_bind (ProofContext.match_bind false); |
|
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
506 |
val match_bind_i = gen_bind (ProofContext.match_bind_i false); |
|
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
507 |
val let_bind = gen_bind (ProofContext.match_bind true); |
|
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
508 |
val let_bind_i = gen_bind (ProofContext.match_bind_i true); |
| 5820 | 509 |
|
| 16813 | 510 |
end; |
511 |
||
| 5820 | 512 |
|
| 17359 | 513 |
(* fix *) |
| 9196 | 514 |
|
| 12714 | 515 |
local |
516 |
||
| 17359 | 517 |
fun gen_fix fx args = |
518 |
assert_forward |
|
519 |
#> map_context (fx args) |
|
520 |
#> put_facts NONE; |
|
| 5820 | 521 |
|
| 16813 | 522 |
in |
523 |
||
| 5820 | 524 |
val fix = gen_fix ProofContext.fix; |
525 |
val fix_i = gen_fix ProofContext.fix_i; |
|
526 |
||
| 16813 | 527 |
end; |
528 |
||
| 5820 | 529 |
|
| 17359 | 530 |
(* assume etc. *) |
| 5820 | 531 |
|
| 16813 | 532 |
local |
533 |
||
| 17112 | 534 |
fun gen_assume asm prep_att exp args state = |
| 5820 | 535 |
state |
536 |
|> assert_forward |
|
| 17112 | 537 |
|> map_context_result (asm exp (Attrib.map_specs (prep_att (theory_of state)) args)) |
| 17359 | 538 |
|> these_factss [] |> #2; |
| 6932 | 539 |
|
| 16813 | 540 |
in |
541 |
||
| 17112 | 542 |
val assm = gen_assume ProofContext.assume Attrib.local_attribute; |
543 |
val assm_i = gen_assume ProofContext.assume_i (K I); |
|
544 |
val assume = assm ProofContext.export_assume; |
|
545 |
val assume_i = assm_i ProofContext.export_assume; |
|
546 |
val presume = assm ProofContext.export_presume; |
|
| 11917 | 547 |
val presume_i = assm_i ProofContext.export_presume; |
| 5820 | 548 |
|
| 16813 | 549 |
end; |
550 |
||
| 7271 | 551 |
|
| 17359 | 552 |
(* def *) |
| 11891 | 553 |
|
| 16813 | 554 |
local |
555 |
||
| 17112 | 556 |
fun gen_def fix prep_att prep_term prep_pats |
557 |
(raw_name, raw_atts) (x, (raw_rhs, raw_pats)) state = |
|
| 11891 | 558 |
let |
559 |
val _ = assert_forward state; |
|
| 17112 | 560 |
val thy = theory_of state; |
| 11891 | 561 |
val ctxt = context_of state; |
562 |
||
563 |
val rhs = prep_term ctxt raw_rhs; |
|
564 |
val T = Term.fastype_of rhs; |
|
565 |
val pats = prep_pats T (ProofContext.declare_term rhs ctxt) raw_pats; |
|
| 17112 | 566 |
val lhs = ProofContext.bind_skolem ctxt [x] (Free (x, T)); |
| 11891 | 567 |
|
| 17112 | 568 |
val name = if raw_name = "" then Thm.def_name x else raw_name; |
569 |
val atts = map (prep_att thy) raw_atts; |
|
| 11891 | 570 |
in |
571 |
state |
|
| 15531 | 572 |
|> fix [([x], NONE)] |
| 11891 | 573 |
|> match_bind_i [(pats, rhs)] |
| 11917 | 574 |
|> assm_i ProofContext.export_def [((name, atts), [(Logic.mk_equals (lhs, rhs), ([], []))])] |
| 11891 | 575 |
end; |
576 |
||
| 16813 | 577 |
in |
578 |
||
| 17112 | 579 |
val def = gen_def fix Attrib.local_attribute ProofContext.read_term ProofContext.read_term_pats; |
580 |
val def_i = gen_def fix_i (K I) ProofContext.cert_term ProofContext.cert_term_pats; |
|
| 11891 | 581 |
|
| 16813 | 582 |
end; |
583 |
||
| 11891 | 584 |
|
| 8374 | 585 |
|
| 17112 | 586 |
(** facts **) |
| 5820 | 587 |
|
| 17359 | 588 |
(* chain *) |
| 5820 | 589 |
|
| 17359 | 590 |
val chain = |
591 |
assert_forward |
|
592 |
#> tap the_facts |
|
593 |
#> enter_chain; |
|
| 5820 | 594 |
|
| 17359 | 595 |
fun chain_facts facts = |
596 |
put_facts (SOME facts) |
|
597 |
#> chain; |
|
| 5820 | 598 |
|
599 |
||
| 17359 | 600 |
(* note etc. *) |
| 17112 | 601 |
|
| 17166 | 602 |
fun no_binding args = map (pair ("", [])) args;
|
| 17112 | 603 |
|
604 |
local |
|
605 |
||
| 17359 | 606 |
fun gen_thmss note_ctxt more_facts opt_chain opt_result prep_atts args state = |
| 17112 | 607 |
state |
608 |
|> assert_forward |
|
609 |
|> map_context_result (note_ctxt (Attrib.map_facts (prep_atts (theory_of state)) args)) |
|
610 |
|> these_factss (more_facts state) |
|
| 17359 | 611 |
||> opt_chain |
612 |
|> opt_result; |
|
| 17112 | 613 |
|
614 |
in |
|
615 |
||
| 17359 | 616 |
val note_thmss = gen_thmss ProofContext.note_thmss (K []) I #2 Attrib.local_attribute; |
617 |
val note_thmss_i = gen_thmss ProofContext.note_thmss_i (K []) I #2 (K I); |
|
| 17112 | 618 |
|
619 |
val from_thmss = |
|
| 17359 | 620 |
gen_thmss ProofContext.note_thmss (K []) chain #2 Attrib.local_attribute o no_binding; |
621 |
val from_thmss_i = gen_thmss ProofContext.note_thmss_i (K []) chain #2 (K I) o no_binding; |
|
| 17112 | 622 |
|
623 |
val with_thmss = |
|
| 17359 | 624 |
gen_thmss ProofContext.note_thmss the_facts chain #2 Attrib.local_attribute o no_binding; |
625 |
val with_thmss_i = gen_thmss ProofContext.note_thmss_i the_facts chain #2 (K I) o no_binding; |
|
626 |
||
627 |
val local_results = gen_thmss ProofContext.note_thmss_i (K []) I I (K I); |
|
628 |
fun global_results kind = |
|
629 |
swap oo (PureThy.note_thmss_i (Drule.kind kind) o map (apsnd Thm.simple_fact)); |
|
630 |
||
631 |
fun get_thmss state srcs = the_facts (note_thmss [(("", []), srcs)] state);
|
|
632 |
fun simple_note_thms name thms = note_thmss_i [((name, []), [(thms, [])])]; |
|
| 17112 | 633 |
|
634 |
end; |
|
635 |
||
636 |
||
637 |
(* using *) |
|
638 |
||
639 |
local |
|
640 |
||
641 |
fun gen_using_thmss note prep_atts args state = |
|
642 |
state |
|
643 |
|> assert_backward |
|
| 17166 | 644 |
|> map_context_result (note (Attrib.map_facts (prep_atts (theory_of state)) (no_binding args))) |
| 17359 | 645 |
|-> (fn named_facts => map_goal I (fn (statement, using, goal, after_qed) => |
646 |
(statement, using @ List.concat (map snd named_facts), goal, after_qed))); |
|
| 17112 | 647 |
|
648 |
in |
|
649 |
||
650 |
val using_thmss = gen_using_thmss ProofContext.note_thmss Attrib.local_attribute; |
|
651 |
val using_thmss_i = gen_using_thmss ProofContext.note_thmss_i (K I); |
|
652 |
||
653 |
end; |
|
654 |
||
655 |
||
656 |
(* case *) |
|
657 |
||
658 |
local |
|
659 |
||
660 |
fun gen_invoke_case prep_att (name, xs, raw_atts) state = |
|
661 |
let |
|
662 |
val atts = map (prep_att (theory_of state)) raw_atts; |
|
| 17359 | 663 |
val ((lets, asms), state') = |
| 17112 | 664 |
map_context_result (ProofContext.apply_case (get_case state name xs)) state; |
665 |
val assumptions = asms |> map (fn (a, ts) => |
|
666 |
((NameSpace.qualified name a, atts), map (rpair ([], [])) ts)); |
|
667 |
in |
|
668 |
state' |
|
669 |
|> add_binds_i lets |
|
| 17359 | 670 |
|> map_context (ProofContext.qualified_names #> ProofContext.no_base_names) |
| 17112 | 671 |
|> assume_i assumptions |
672 |
|> map_context (ProofContext.restore_naming (context_of state)) |
|
673 |
|> `the_facts |-> simple_note_thms name |
|
674 |
end; |
|
675 |
||
676 |
in |
|
677 |
||
678 |
val invoke_case = gen_invoke_case Attrib.local_attribute; |
|
679 |
val invoke_case_i = gen_invoke_case (K I); |
|
680 |
||
681 |
end; |
|
682 |
||
683 |
||
684 |
||
| 17359 | 685 |
(** proof structure **) |
686 |
||
687 |
(* blocks *) |
|
688 |
||
689 |
val begin_block = |
|
690 |
assert_forward |
|
691 |
#> open_block |
|
692 |
#> put_goal NONE |
|
693 |
#> open_block; |
|
694 |
||
695 |
val next_block = |
|
696 |
assert_forward |
|
697 |
#> close_block |
|
698 |
#> open_block |
|
699 |
#> put_goal NONE |
|
700 |
#> put_facts NONE; |
|
701 |
||
702 |
fun end_block state = |
|
703 |
state |
|
704 |
|> assert_forward |
|
705 |
|> close_block |
|
706 |
|> assert_current_goal false |
|
707 |
|> close_block |
|
708 |
|> export_facts state; |
|
709 |
||
710 |
||
711 |
(* sub-proofs *) |
|
712 |
||
713 |
fun proof opt_text = |
|
714 |
assert_backward |
|
715 |
#> refine (if_none opt_text Method.default_text) |
|
716 |
#> Seq.map (using_facts [] #> enter_forward); |
|
717 |
||
718 |
fun end_proof bot txt = |
|
719 |
assert_forward |
|
720 |
#> assert_bottom bot |
|
721 |
#> close_block |
|
722 |
#> assert_current_goal true |
|
723 |
#> using_facts [] |
|
724 |
#> refine (Method.finish_text txt); |
|
725 |
||
726 |
||
727 |
(* unstructured refinement *) |
|
728 |
||
729 |
fun defer i = assert_no_chain #> refine (Method.Basic (K (Method.defer i))); |
|
730 |
fun prefer i = assert_no_chain #> refine (Method.Basic (K (Method.prefer i))); |
|
731 |
||
732 |
fun apply text = assert_backward #> refine text #> Seq.map (using_facts []); |
|
733 |
fun apply_end text = assert_forward #> refine_end text; |
|
734 |
||
735 |
||
736 |
||
| 17112 | 737 |
(** goals **) |
738 |
||
| 17359 | 739 |
(* goal names *) |
740 |
||
741 |
fun prep_names prep_att stmt = |
|
742 |
let |
|
743 |
val (names_attss, propp) = split_list (Attrib.map_specs prep_att stmt); |
|
744 |
val (names, attss) = split_list names_attss; |
|
745 |
in ((names, attss), propp) end; |
|
| 5820 | 746 |
|
| 17359 | 747 |
fun goal_names target name names = |
748 |
(case target of NONE => "" | SOME loc => " (in " ^ loc ^ ")") ^ |
|
749 |
(if name = "" then "" else " " ^ name) ^ |
|
750 |
(case filter_out (equal "") names of [] => "" |
|
751 |
| nms => " " ^ enclose "(" ")" (space_implode " and " (Library.take (7, nms) @
|
|
752 |
(if length nms > 7 then ["..."] else [])))); |
|
753 |
||
| 16813 | 754 |
|
| 17359 | 755 |
(* generic goals *) |
756 |
||
757 |
fun check_tvars props state = |
|
758 |
(case fold Term.add_tvars props [] of [] => () |
|
759 |
| tvars => raise STATE ("Goal statement contains illegal schematic type variable(s): " ^
|
|
760 |
commas (map (ProofContext.string_of_typ (context_of state) o TVar) tvars), state)); |
|
| 8991 | 761 |
|
| 17359 | 762 |
fun check_vars props state = |
763 |
(case fold Term.add_vars props [] of [] => () |
|
764 |
| vars => warning ("Goal statement contains unbound schematic variable(s): " ^
|
|
765 |
commas (map (ProofContext.string_of_term (context_of state) o Var) vars))); |
|
766 |
||
767 |
fun generic_goal prepp kind after_qed raw_propp state = |
|
| 5820 | 768 |
let |
| 17359 | 769 |
val thy = theory_of state; |
770 |
val chaining = can assert_chain state; |
|
771 |
||
772 |
val ((propss, after_ctxt), goal_state) = |
|
| 5936 | 773 |
state |
774 |
|> assert_forward_or_chain |
|
775 |
|> enter_forward |
|
| 17359 | 776 |
|> open_block |
| 12534 | 777 |
|> map_context_result (fn ctxt => prepp (ctxt, raw_propp)); |
| 15703 | 778 |
|
| 15570 | 779 |
val props = List.concat propss; |
| 17359 | 780 |
val goal = Drule.mk_triv_goal (Thm.cterm_of thy (Logic.mk_conjunction_list props)); |
| 10553 | 781 |
|
| 17359 | 782 |
val after_qed' = after_qed |>> (fn after_local => fn raw_results => fn results => |
783 |
map_context after_ctxt |
|
784 |
#> after_local raw_results results); |
|
| 5820 | 785 |
in |
| 17359 | 786 |
goal_state |
787 |
|> tap (check_tvars props) |
|
788 |
|> tap (check_vars props) |
|
789 |
|> put_goal (SOME (make_goal ((kind, propss), [], goal, after_qed'))) |
|
790 |
|> map_context (ProofContext.add_cases (AutoBind.cases thy props)) |
|
791 |
|> map_context (ProofContext.auto_bind_goal props) |
|
792 |
|> K chaining ? (`the_facts #-> using_facts) |
|
793 |
|> put_facts NONE |
|
794 |
|> open_block |
|
795 |
|> put_goal NONE |
|
| 5820 | 796 |
|> enter_backward |
797 |
end; |
|
798 |
||
| 17359 | 799 |
fun generic_qed state = |
| 12503 | 800 |
let |
| 17359 | 801 |
val (goal_ctxt, {statement = (_, stmt), using, goal, after_qed}) = current_goal state;
|
|
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
802 |
val outer_state = state |> close_block; |
|
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
803 |
val outer_ctxt = context_of outer_state; |
|
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
804 |
|
| 17359 | 805 |
val raw_props = List.concat stmt; |
806 |
val raw_results = conclude_goal state raw_props goal; |
|
807 |
val props = ProofContext.generalize goal_ctxt outer_ctxt raw_props; |
|
808 |
val results = |
|
809 |
Seq.map_list (ProofContext.export false goal_ctxt outer_ctxt) raw_results |
|
810 |
|> Seq.map (Library.unflat stmt); |
|
811 |
in |
|
812 |
outer_state |
|
813 |
|> map_context (ProofContext.auto_bind_facts props) |
|
814 |
|> pair (after_qed, ((goal_ctxt, raw_results), results)) |
|
815 |
end; |
|
816 |
||
| 9469 | 817 |
|
| 17359 | 818 |
(* local goals *) |
819 |
||
820 |
fun local_goal print_results prep_att prepp kind after_qed stmt state = |
|
821 |
let |
|
822 |
val ((names, attss), propp) = prep_names (prep_att (theory_of state)) stmt; |
|
823 |
||
824 |
fun after_qed' raw_results results = |
|
825 |
local_results ((names ~~ attss) ~~ map Thm.simple_fact results) |
|
826 |
#-> (fn res => tap (fn st => print_results (context_of st) ((kind, ""), res) : unit)) |
|
827 |
#> after_qed raw_results results; |
|
| 5820 | 828 |
in |
| 17359 | 829 |
state |
830 |
|> generic_goal prepp (kind ^ goal_names NONE "" names) (after_qed', K (K I)) propp |
|
831 |
|> warn_extra_tfrees state |
|
| 5820 | 832 |
end; |
833 |
||
| 17359 | 834 |
fun local_qed txt = |
835 |
end_proof false txt |
|
836 |
#> Seq.map generic_qed |
|
837 |
#> Seq.maps (uncurry (fn (after_qed, (raw_results, results)) => |
|
838 |
Seq.lifts (#1 after_qed raw_results) results)); |
|
| 5820 | 839 |
|
840 |
||
| 17359 | 841 |
(* global goals *) |
842 |
||
| 17437 | 843 |
fun auto_fix (ctxt, (propss, after_ctxt)) = |
844 |
(ctxt |> ProofContext.fix_frees (List.concat propss), (propss, after_ctxt)); |
|
845 |
||
| 17359 | 846 |
fun global_goal print_results prep_att prepp |
847 |
kind after_qed target (name, raw_atts) stmt ctxt = |
|
848 |
let |
|
849 |
val thy = ProofContext.theory_of ctxt; |
|
850 |
val thy_ctxt = ProofContext.init thy; |
|
851 |
||
852 |
val atts = map (prep_att thy) raw_atts; |
|
853 |
val ((names, attss), propp) = prep_names (prep_att thy) stmt; |
|
854 |
||
855 |
fun after_qed' raw_results results = |
|
856 |
(map o map) (ProofContext.export_standard ctxt thy_ctxt |
|
857 |
#> Drule.strip_shyps_warning) results |
|
858 |
|> (fn res => global_results kind ((names ~~ attss) ~~ res)) |
|
859 |
#-> (fn res' => |
|
860 |
(print_results thy_ctxt ((kind, name), res') : unit; |
|
861 |
#2 o global_results kind [((name, atts), List.concat (map #2 res'))])) |
|
862 |
#> after_qed raw_results results; |
|
863 |
||
| 17437 | 864 |
val prepp' = prepp #> auto_fix; |
| 17359 | 865 |
in |
866 |
init ctxt |
|
867 |
|> generic_goal prepp' (kind ^ goal_names target name names) |
|
868 |
(K (K Seq.single), after_qed') propp |
|
869 |
end; |
|
| 12065 | 870 |
|
| 17112 | 871 |
fun check_result msg state sq = |
872 |
(case Seq.pull sq of |
|
873 |
NONE => raise STATE (msg, state) |
|
874 |
| SOME res => Seq.cons res); |
|
875 |
||
| 17359 | 876 |
fun global_qeds txt = |
877 |
end_proof true txt |
|
878 |
#> Seq.map generic_qed |
|
879 |
#> Seq.maps (fn ((after_qed, (raw_results, results)), state) => |
|
880 |
Seq.lift (#2 after_qed raw_results) results (theory_of state) |
|
881 |
|> Seq.map (rpair (context_of state))) |
|
| 17112 | 882 |
|> Seq.DETERM; (*backtracking may destroy theory!*) |
883 |
||
| 17359 | 884 |
fun global_qed txt state = |
| 17112 | 885 |
state |
| 17359 | 886 |
|> global_qeds txt |
| 17112 | 887 |
|> check_result "Failed to finish proof" state |
888 |
|> Seq.hd; |
|
| 12959 | 889 |
|
| 5820 | 890 |
|
| 17359 | 891 |
(* concluding steps *) |
| 17112 | 892 |
|
| 17359 | 893 |
fun local_terminal_proof (text, opt_text) = |
894 |
proof (SOME text) #> Seq.maps (local_qed (opt_text, true)); |
|
| 17112 | 895 |
|
896 |
val local_default_proof = local_terminal_proof (Method.default_text, NONE); |
|
897 |
val local_immediate_proof = local_terminal_proof (Method.this_text, NONE); |
|
| 17359 | 898 |
val local_done_proof = proof (SOME Method.done_text) #> Seq.maps (local_qed (NONE, false)); |
899 |
fun local_skip_proof int = local_terminal_proof (Method.sorry_text int, NONE); |
|
| 17112 | 900 |
|
| 17359 | 901 |
fun global_term_proof immed (text, opt_text) state = |
| 5820 | 902 |
state |
| 17112 | 903 |
|> proof (SOME text) |
904 |
|> check_result "Terminal proof method failed" state |
|
| 17359 | 905 |
|> Seq.maps (global_qeds (opt_text, immed)) |
| 17112 | 906 |
|> check_result "Failed to finish proof (after successful terminal method)" state |
907 |
|> Seq.hd; |
|
908 |
||
909 |
val global_terminal_proof = global_term_proof true; |
|
910 |
val global_default_proof = global_terminal_proof (Method.default_text, NONE); |
|
911 |
val global_immediate_proof = global_terminal_proof (Method.this_text, NONE); |
|
912 |
val global_done_proof = global_term_proof false (Method.done_text, NONE); |
|
| 17359 | 913 |
fun global_skip_proof int = global_terminal_proof (Method.sorry_text int, NONE); |
| 5820 | 914 |
|
| 6896 | 915 |
|
| 17359 | 916 |
(* common goal statements *) |
917 |
||
918 |
local |
|
919 |
||
920 |
fun gen_have prep_att prepp after_qed stmt int = |
|
921 |
local_goal (ProofDisplay.print_results int) prep_att prepp "have" after_qed stmt; |
|
| 6896 | 922 |
|
| 17359 | 923 |
fun gen_show prep_att prepp after_qed stmt int state = |
924 |
let |
|
925 |
val testing = ref false; |
|
926 |
val rule = ref (NONE: thm option); |
|
927 |
fun fail_msg ctxt = |
|
928 |
"Local statement will fail to solve any pending goal" :: |
|
929 |
(case ! rule of NONE => [] | SOME th => [ProofDisplay.string_of_rule ctxt "Failed" th]) |
|
930 |
|> cat_lines; |
|
| 6896 | 931 |
|
| 17359 | 932 |
fun print_results ctxt res = |
933 |
if ! testing then () else ProofDisplay.print_results int ctxt res; |
|
934 |
fun print_rule ctxt th = |
|
935 |
if ! testing then rule := SOME th |
|
936 |
else if int then priority (ProofDisplay.string_of_rule ctxt "Successful" th) |
|
937 |
else (); |
|
938 |
val test_proof = |
|
939 |
(Seq.pull oo local_skip_proof) true |
|
940 |
|> setmp testing true |
|
941 |
|> setmp proofs 0 |
|
942 |
|> transform_error |
|
943 |
|> capture; |
|
| 6896 | 944 |
|
| 17359 | 945 |
fun after_qed' raw_results results = |
946 |
refine_goals print_rule (context_of state) (List.concat results) |
|
947 |
#> Seq.maps (after_qed raw_results results); |
|
948 |
in |
|
949 |
state |
|
950 |
|> local_goal print_results prep_att prepp "show" after_qed' stmt |
|
951 |
|> K int ? (fn goal_state => |
|
952 |
(case test_proof goal_state of |
|
953 |
Result (SOME _) => goal_state |
|
954 |
| Result NONE => raise STATE (fail_msg (context_of goal_state), goal_state) |
|
955 |
| Exn Interrupt => raise Interrupt |
|
956 |
| Exn exn => raise EXCEPTION (exn, fail_msg (context_of goal_state)))) |
|
957 |
end; |
|
958 |
||
959 |
fun gen_theorem prep_att prepp kind after_qed target a = |
|
960 |
global_goal ProofDisplay.present_results prep_att prepp kind after_qed target a; |
|
961 |
||
962 |
in |
|
963 |
||
964 |
val have = gen_have Attrib.local_attribute ProofContext.bind_propp; |
|
965 |
val have_i = gen_have (K I) ProofContext.bind_propp_i; |
|
966 |
val show = gen_show Attrib.local_attribute ProofContext.bind_propp; |
|
967 |
val show_i = gen_show (K I) ProofContext.bind_propp_i; |
|
968 |
val theorem = gen_theorem Attrib.global_attribute ProofContext.bind_propp_schematic; |
|
969 |
val theorem_i = gen_theorem (K I) ProofContext.bind_propp_schematic_i; |
|
| 6896 | 970 |
|
| 5820 | 971 |
end; |
| 17359 | 972 |
|
973 |
end; |