author | wenzelm |
Sun, 04 Nov 2001 20:58:01 +0100 | |
changeset 12045 | bfaa23b19f47 |
parent 12015 | 68b2a53161e6 |
child 12055 | a9c44895cc8c |
permissions | -rw-r--r-- |
5820 | 1 |
(* Title: Pure/Isar/proof.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
8807 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
5820 | 5 |
|
6 |
Proof states and methods. |
|
7 |
*) |
|
8 |
||
8152 | 9 |
signature BASIC_PROOF = |
10 |
sig |
|
8374 | 11 |
val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq |
12 |
val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq |
|
8152 | 13 |
end; |
14 |
||
5820 | 15 |
signature PROOF = |
16 |
sig |
|
8152 | 17 |
include BASIC_PROOF |
5820 | 18 |
type context |
19 |
type state |
|
20 |
exception STATE of string * state |
|
6871 | 21 |
val check_result: string -> state -> 'a Seq.seq -> 'a Seq.seq |
12045 | 22 |
val init_state: theory -> state |
5820 | 23 |
val context_of: state -> context |
24 |
val theory_of: state -> theory |
|
25 |
val sign_of: state -> Sign.sg |
|
7924 | 26 |
val warn_extra_tfrees: state -> state -> state |
7605 | 27 |
val reset_thms: string -> state -> state |
6091 | 28 |
val the_facts: state -> thm list |
9469 | 29 |
val the_fact: state -> thm |
11079
a378479996f7
val get_goal: state -> context * (thm list * thm);
wenzelm
parents:
10938
diff
changeset
|
30 |
val get_goal: state -> context * (thm list * thm) |
6091 | 31 |
val goal_facts: (state -> thm list) -> state -> state |
5820 | 32 |
val use_facts: state -> state |
33 |
val reset_facts: state -> state |
|
9469 | 34 |
val is_chain: state -> bool |
6891 | 35 |
val assert_forward: state -> state |
9469 | 36 |
val assert_forward_or_chain: state -> state |
5820 | 37 |
val assert_backward: state -> state |
8206 | 38 |
val assert_no_chain: state -> state |
5820 | 39 |
val enter_forward: state -> state |
6982 | 40 |
val show_hyps: bool ref |
41 |
val pretty_thm: thm -> Pretty.T |
|
7412 | 42 |
val pretty_thms: thm list -> Pretty.T |
6529 | 43 |
val verbose: bool ref |
7201 | 44 |
val print_state: int -> state -> unit |
10360 | 45 |
val pretty_goals: bool -> state -> Pretty.T list |
6776 | 46 |
val level: state -> int |
5820 | 47 |
type method |
6848 | 48 |
val method: (thm list -> tactic) -> method |
8374 | 49 |
val method_cases: (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> method |
5820 | 50 |
val refine: (context -> method) -> state -> state Seq.seq |
8234 | 51 |
val refine_end: (context -> method) -> state -> state Seq.seq |
5936 | 52 |
val match_bind: (string list * string) list -> state -> state |
53 |
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
|
54 |
val let_bind: (string list * string) list -> state -> state |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
55 |
val let_bind_i: (term list * term) list -> state -> state |
6876 | 56 |
val simple_have_thms: string -> thm list -> state -> state |
9196 | 57 |
val have_thmss: ((string * context attribute list) * |
58 |
(thm list * context attribute list) list) list -> state -> state |
|
59 |
val with_thmss: ((string * context attribute list) * |
|
60 |
(thm list * context attribute list) list) list -> state -> state |
|
7412 | 61 |
val fix: (string list * string option) list -> state -> state |
7665 | 62 |
val fix_i: (string list * typ option) list -> state -> state |
9469 | 63 |
val assm: ProofContext.exporter |
10464 | 64 |
-> ((string * context attribute list) * (string * (string list * string list)) list) list |
7271 | 65 |
-> state -> state |
9469 | 66 |
val assm_i: ProofContext.exporter |
10464 | 67 |
-> ((string * context attribute list) * (term * (term list * term list)) list) list |
6932 | 68 |
-> state -> state |
10464 | 69 |
val assume: |
70 |
((string * context attribute list) * (string * (string list * string list)) list) list |
|
7271 | 71 |
-> state -> state |
10464 | 72 |
val assume_i: ((string * context attribute list) * (term * (term list * term list)) list) list |
6932 | 73 |
-> state -> state |
10464 | 74 |
val presume: |
75 |
((string * context attribute list) * (string * (string list * string list)) list) list |
|
6932 | 76 |
-> state -> state |
10464 | 77 |
val presume_i: ((string * context attribute list) * (term * (term list * term list)) list) list |
6932 | 78 |
-> state -> state |
11891 | 79 |
val def: string -> context attribute list -> string * (string * string list) -> state -> state |
80 |
val def_i: string -> context attribute list -> string * (term * term list) -> state -> state |
|
11793
5f0ab6f5c280
support impromptu terminology of cases parameters;
wenzelm
parents:
11742
diff
changeset
|
81 |
val invoke_case: string * string option list * context attribute list -> state -> state |
12045 | 82 |
val theorem: string -> xstring option -> context attribute Locale.element list -> bstring |
83 |
-> theory attribute list -> string * (string list * string list) -> theory -> state |
|
84 |
val theorem_i: string -> string option -> context attribute Locale.element_i list -> bstring |
|
85 |
-> theory attribute list -> term * (term list * term list) -> theory -> state |
|
5820 | 86 |
val chain: state -> state |
6091 | 87 |
val from_facts: thm list -> state -> state |
10936 | 88 |
val show: (bool -> state -> state) -> (state -> state Seq.seq) -> string |
89 |
-> context attribute list -> string * (string list * string list) -> bool -> state -> state |
|
90 |
val show_i: (bool -> state -> state) -> (state -> state Seq.seq) -> string |
|
91 |
-> context attribute list -> term * (term list * term list) -> bool -> state -> state |
|
7176 | 92 |
val have: (state -> state Seq.seq) -> string -> context attribute list |
93 |
-> string * (string list * string list) -> state -> state |
|
94 |
val have_i: (state -> state Seq.seq) -> string -> context attribute list |
|
95 |
-> term * (term list * term list) -> state -> state |
|
6404 | 96 |
val at_bottom: state -> bool |
6982 | 97 |
val local_qed: (state -> state Seq.seq) |
98 |
-> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) -> state -> state Seq.seq |
|
6950 | 99 |
val global_qed: (state -> state Seq.seq) -> state |
100 |
-> (theory * {kind: string, name: string, thm: thm}) Seq.seq |
|
6896 | 101 |
val begin_block: state -> state |
6932 | 102 |
val end_block: state -> state Seq.seq |
6896 | 103 |
val next_block: state -> state |
5820 | 104 |
end; |
105 |
||
8152 | 106 |
signature PRIVATE_PROOF = |
5820 | 107 |
sig |
108 |
include PROOF |
|
109 |
val put_data: Object.kind -> ('a -> Object.T) -> 'a -> state -> state |
|
110 |
end; |
|
111 |
||
8152 | 112 |
structure Proof: PRIVATE_PROOF = |
5820 | 113 |
struct |
114 |
||
115 |
||
116 |
(** proof state **) |
|
117 |
||
118 |
type context = ProofContext.context; |
|
119 |
||
120 |
||
121 |
(* type goal *) |
|
122 |
||
123 |
datatype kind = |
|
12010 | 124 |
Theorem of string * string option * theory attribute list | (*theorem with kind and locale*) |
11742 | 125 |
Show of context attribute list | (*intermediate result, solving subgoal*) |
11985 | 126 |
Have of context attribute list; (*intermediate result*) |
5820 | 127 |
|
12015 | 128 |
fun kind_name _ (Theorem (s, None, _)) = s |
129 |
| kind_name sg (Theorem (s, Some loc, _)) = s ^ " (in " ^ Locale.cond_extern sg loc ^ ")" |
|
130 |
| kind_name _ (Show _) = "show" |
|
131 |
| kind_name _ (Have _) = "have"; |
|
5820 | 132 |
|
133 |
type goal = |
|
6932 | 134 |
(kind * (*result kind*) |
135 |
string * (*result name*) |
|
136 |
term) * (*result statement*) |
|
137 |
(thm list * (*use facts*) |
|
138 |
thm); (*goal: subgoals ==> statement*) |
|
5820 | 139 |
|
140 |
||
141 |
(* type mode *) |
|
142 |
||
143 |
datatype mode = Forward | ForwardChain | Backward; |
|
7201 | 144 |
val mode_name = (fn Forward => "state" | ForwardChain => "chain" | Backward => "prove"); |
5820 | 145 |
|
146 |
||
147 |
(* datatype state *) |
|
148 |
||
7176 | 149 |
datatype node = |
150 |
Node of |
|
151 |
{context: context, |
|
152 |
facts: thm list option, |
|
153 |
mode: mode, |
|
154 |
goal: (goal * (state -> state Seq.seq)) option} |
|
155 |
and state = |
|
5820 | 156 |
State of |
157 |
node * (*current*) |
|
158 |
node list; (*parents wrt. block structure*) |
|
159 |
||
7176 | 160 |
fun make_node (context, facts, mode, goal) = |
161 |
Node {context = context, facts = facts, mode = mode, goal = goal}; |
|
162 |
||
163 |
||
5820 | 164 |
exception STATE of string * state; |
165 |
||
166 |
fun err_malformed name state = |
|
167 |
raise STATE (name ^ ": internal error -- malformed proof state", state); |
|
168 |
||
6871 | 169 |
fun check_result msg state sq = |
170 |
(case Seq.pull sq of |
|
171 |
None => raise STATE (msg, state) |
|
172 |
| Some s_sq => Seq.cons s_sq); |
|
173 |
||
5820 | 174 |
|
7176 | 175 |
fun map_current f (State (Node {context, facts, mode, goal}, nodes)) = |
5820 | 176 |
State (make_node (f (context, facts, mode, goal)), nodes); |
177 |
||
12045 | 178 |
fun init_state thy = |
179 |
State (make_node (ProofContext.init thy, None, Forward, None), []); |
|
180 |
||
5820 | 181 |
|
182 |
||
183 |
(** basic proof state operations **) |
|
184 |
||
185 |
(* context *) |
|
186 |
||
7176 | 187 |
fun context_of (State (Node {context, ...}, _)) = context; |
5820 | 188 |
val theory_of = ProofContext.theory_of o context_of; |
189 |
val sign_of = ProofContext.sign_of o context_of; |
|
190 |
||
191 |
fun map_context f = map_current (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal)); |
|
192 |
||
7176 | 193 |
fun map_context_result f (state as State (Node {context, facts, mode, goal}, nodes)) = |
5820 | 194 |
let val (context', result) = f context |
195 |
in (State (make_node (context', facts, mode, goal), nodes), result) end; |
|
196 |
||
197 |
||
198 |
fun put_data kind f = map_context o ProofContext.put_data kind f; |
|
7924 | 199 |
val warn_extra_tfrees = map_context o ProofContext.warn_extra_tfrees o context_of; |
10809 | 200 |
val add_binds_i = map_context o ProofContext.add_binds_i; |
6790 | 201 |
val auto_bind_goal = map_context o ProofContext.auto_bind_goal; |
6798 | 202 |
val auto_bind_facts = map_context oo ProofContext.auto_bind_facts; |
6091 | 203 |
val put_thms = map_context o ProofContext.put_thms; |
204 |
val put_thmss = map_context o ProofContext.put_thmss; |
|
7605 | 205 |
val reset_thms = map_context o ProofContext.reset_thms; |
6932 | 206 |
val assumptions = ProofContext.assumptions o context_of; |
8374 | 207 |
val get_case = ProofContext.get_case o context_of; |
5820 | 208 |
|
209 |
||
210 |
(* facts *) |
|
211 |
||
7176 | 212 |
fun the_facts (State (Node {facts = Some facts, ...}, _)) = facts |
5820 | 213 |
| the_facts state = raise STATE ("No current facts available", state); |
214 |
||
9469 | 215 |
fun the_fact state = |
216 |
(case the_facts state of |
|
217 |
[thm] => thm |
|
218 |
| _ => raise STATE ("Single fact expected", state)); |
|
219 |
||
6848 | 220 |
fun assert_facts state = (the_facts state; state); |
7176 | 221 |
fun get_facts (State (Node {facts, ...}, _)) = facts; |
6848 | 222 |
|
7605 | 223 |
|
224 |
val thisN = "this"; |
|
225 |
||
5820 | 226 |
fun put_facts facts state = |
227 |
state |
|
228 |
|> map_current (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) |
|
7605 | 229 |
|> (case facts of None => reset_thms thisN | Some ths => put_thms (thisN, ths)); |
5820 | 230 |
|
231 |
val reset_facts = put_facts None; |
|
232 |
||
9196 | 233 |
fun these_factss (state, named_factss) = |
5820 | 234 |
state |
9196 | 235 |
|> put_thmss named_factss |
236 |
|> put_facts (Some (flat (map #2 named_factss))); |
|
5820 | 237 |
|
238 |
||
239 |
(* goal *) |
|
240 |
||
10553 | 241 |
local |
242 |
fun find i (state as State (Node {goal = Some goal, ...}, _)) = (context_of state, (i, goal)) |
|
243 |
| find i (State (Node {goal = None, ...}, node :: nodes)) = find (i + 1) (State (node, nodes)) |
|
244 |
| find _ (state as State (_, [])) = err_malformed "find_goal" state; |
|
245 |
in val find_goal = find 0 end; |
|
7176 | 246 |
|
11079
a378479996f7
val get_goal: state -> context * (thm list * thm);
wenzelm
parents:
10938
diff
changeset
|
247 |
fun get_goal state = |
a378479996f7
val get_goal: state -> context * (thm list * thm);
wenzelm
parents:
10938
diff
changeset
|
248 |
let val (ctxt, (_, ((_, x), _))) = find_goal state |
a378479996f7
val get_goal: state -> context * (thm list * thm);
wenzelm
parents:
10938
diff
changeset
|
249 |
in (ctxt, x) end; |
a378479996f7
val get_goal: state -> context * (thm list * thm);
wenzelm
parents:
10938
diff
changeset
|
250 |
|
5820 | 251 |
fun put_goal goal = map_current (fn (ctxt, facts, mode, _) => (ctxt, facts, mode, goal)); |
252 |
||
8374 | 253 |
fun map_goal f g (State (Node {context, facts, mode, goal = Some goal}, nodes)) = |
254 |
State (make_node (f context, facts, mode, Some (g goal)), nodes) |
|
255 |
| map_goal f g (State (nd, node :: nodes)) = |
|
256 |
let val State (node', nodes') = map_goal f g (State (node, nodes)) |
|
257 |
in map_context f (State (nd, node' :: nodes')) end |
|
258 |
| map_goal _ _ state = state; |
|
5820 | 259 |
|
260 |
fun goal_facts get state = |
|
261 |
state |
|
8374 | 262 |
|> map_goal I (fn ((result, (_, thm)), f) => ((result, (get state, thm)), f)); |
5820 | 263 |
|
264 |
fun use_facts state = |
|
265 |
state |
|
266 |
|> goal_facts the_facts |
|
267 |
|> reset_facts; |
|
268 |
||
269 |
||
270 |
(* mode *) |
|
271 |
||
7176 | 272 |
fun get_mode (State (Node {mode, ...}, _)) = mode; |
5820 | 273 |
fun put_mode mode = map_current (fn (ctxt, facts, _, goal) => (ctxt, facts, mode, goal)); |
274 |
||
275 |
val enter_forward = put_mode Forward; |
|
276 |
val enter_forward_chain = put_mode ForwardChain; |
|
277 |
val enter_backward = put_mode Backward; |
|
278 |
||
279 |
fun assert_mode pred state = |
|
280 |
let val mode = get_mode state in |
|
281 |
if pred mode then state |
|
12010 | 282 |
else raise STATE ("Illegal application of proof command in " |
283 |
^ quote (mode_name mode) ^ " mode", state) |
|
5820 | 284 |
end; |
285 |
||
286 |
fun is_chain state = get_mode state = ForwardChain; |
|
287 |
val assert_forward = assert_mode (equal Forward); |
|
288 |
val assert_forward_or_chain = assert_mode (equal Forward orf equal ForwardChain); |
|
289 |
val assert_backward = assert_mode (equal Backward); |
|
8206 | 290 |
val assert_no_chain = assert_mode (not_equal ForwardChain); |
5820 | 291 |
|
292 |
||
293 |
(* blocks *) |
|
294 |
||
6776 | 295 |
fun level (State (_, nodes)) = length nodes; |
296 |
||
5820 | 297 |
fun open_block (State (node, nodes)) = State (node, node :: nodes); |
298 |
||
299 |
fun new_block state = |
|
300 |
state |
|
301 |
|> open_block |
|
302 |
|> put_goal None; |
|
303 |
||
7487
c0f9b956a3e7
close_block: removed ProofContext.transfer_used_names;
wenzelm
parents:
7478
diff
changeset
|
304 |
fun close_block (state as State (_, node :: nodes)) = State (node, nodes) |
5820 | 305 |
| close_block state = raise STATE ("Unbalanced block parentheses", state); |
306 |
||
307 |
||
308 |
||
309 |
(** print_state **) |
|
310 |
||
6982 | 311 |
val show_hyps = ProofContext.show_hyps; |
312 |
val pretty_thm = ProofContext.pretty_thm; |
|
313 |
||
7412 | 314 |
fun pretty_thms [th] = pretty_thm th |
7580 | 315 |
| pretty_thms ths = Pretty.blk (0, Pretty.fbreaks (map pretty_thm ths)); |
7412 | 316 |
|
317 |
||
6529 | 318 |
val verbose = ProofContext.verbose; |
319 |
||
8462 | 320 |
fun pretty_facts _ None = [] |
321 |
| pretty_facts s (Some ths) = |
|
322 |
[Pretty.big_list (s ^ "this:") (map pretty_thm ths), Pretty.str ""]; |
|
6756 | 323 |
|
7201 | 324 |
fun print_state nr (state as State (Node {context, facts, mode, goal = _}, nodes)) = |
5820 | 325 |
let |
11891 | 326 |
val ref (_, _, begin_goal) = Display.current_goals_markers; |
5820 | 327 |
|
328 |
fun levels_up 0 = "" |
|
7575 | 329 |
| levels_up 1 = ", 1 level up" |
330 |
| levels_up i = ", " ^ string_of_int i ^ " levels up"; |
|
5820 | 331 |
|
11556 | 332 |
fun subgoals 0 = "" |
333 |
| subgoals 1 = ", 1 subgoal" |
|
334 |
| subgoals n = ", " ^ string_of_int n ^ " subgoals"; |
|
335 |
||
8462 | 336 |
fun pretty_goal (_, (i, (((kind, name, _), (goal_facts, thm)), _))) = |
337 |
pretty_facts "using " |
|
338 |
(if mode <> Backward orelse null goal_facts then None else Some goal_facts) @ |
|
12015 | 339 |
[Pretty.str ("goal (" ^ kind_name (sign_of state) kind ^ |
11556 | 340 |
(if name = "" then "" else " " ^ name) ^ |
341 |
levels_up (i div 2) ^ subgoals (Thm.nprems_of thm) ^ "):")] @ |
|
11891 | 342 |
Display.pretty_goals_marker begin_goal (! goals_limit) thm; |
6848 | 343 |
|
8462 | 344 |
val ctxt_prts = |
345 |
if ! verbose orelse mode = Forward then ProofContext.pretty_context context |
|
346 |
else if mode = Backward then ProofContext.pretty_prems context |
|
7575 | 347 |
else []; |
8462 | 348 |
|
349 |
val prts = |
|
8582 | 350 |
[Pretty.str ("proof (" ^ mode_name mode ^ "): step " ^ string_of_int nr ^ |
12010 | 351 |
(if ! verbose then ", depth " ^ string_of_int (length nodes div 2 - 1) |
8561 | 352 |
else "")), Pretty.str ""] @ |
8462 | 353 |
(if null ctxt_prts then [] else ctxt_prts @ [Pretty.str ""]) @ |
354 |
(if ! verbose orelse mode = Forward then |
|
10553 | 355 |
(pretty_facts "" facts @ pretty_goal (find_goal state)) |
8462 | 356 |
else if mode = ForwardChain then pretty_facts "picking " facts |
10553 | 357 |
else pretty_goal (find_goal state)) |
8462 | 358 |
in Pretty.writeln (Pretty.chunks prts) end; |
5820 | 359 |
|
10360 | 360 |
fun pretty_goals main_goal state = |
10553 | 361 |
let val (_, (_, ((_, (_, thm)), _))) = find_goal state |
11891 | 362 |
in Display.pretty_sub_goals main_goal (! goals_limit) thm end; |
10320 | 363 |
|
5820 | 364 |
|
365 |
||
366 |
(** proof steps **) |
|
367 |
||
368 |
(* datatype method *) |
|
369 |
||
8374 | 370 |
datatype method = |
371 |
Method of thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq; |
|
372 |
||
373 |
fun method tac = Method (fn facts => fn st => Seq.map (rpair []) (tac facts st)); |
|
374 |
val method_cases = Method; |
|
5820 | 375 |
|
376 |
||
377 |
(* refine goal *) |
|
378 |
||
8234 | 379 |
local |
380 |
||
5820 | 381 |
fun check_sign sg state = |
382 |
if Sign.subsig (sg, sign_of state) then state |
|
383 |
else raise STATE ("Bad signature of result: " ^ Sign.str_of_sg sg, state); |
|
384 |
||
8234 | 385 |
fun gen_refine current_context meth_fun state = |
6848 | 386 |
let |
10553 | 387 |
val (goal_ctxt, (_, ((result, (facts, thm)), f))) = find_goal state; |
8234 | 388 |
val Method meth = meth_fun (if current_context then context_of state else goal_ctxt); |
5820 | 389 |
|
8374 | 390 |
fun refn (thm', cases) = |
6848 | 391 |
state |
392 |
|> check_sign (Thm.sign_of_thm thm') |
|
8374 | 393 |
|> map_goal (ProofContext.add_cases cases) (K ((result, (facts, thm')), f)); |
6848 | 394 |
in Seq.map refn (meth facts thm) end; |
5820 | 395 |
|
8234 | 396 |
in |
397 |
||
398 |
val refine = gen_refine true; |
|
399 |
val refine_end = gen_refine false; |
|
400 |
||
401 |
end; |
|
402 |
||
5820 | 403 |
|
11917 | 404 |
(* goal addressing *) |
6932 | 405 |
|
8152 | 406 |
fun FINDGOAL tac st = |
8374 | 407 |
let fun find (i, n) = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1, n)) |
408 |
in find (1, Thm.nprems_of st) st end; |
|
8152 | 409 |
|
8166 | 410 |
fun HEADGOAL tac = tac 1; |
411 |
||
9469 | 412 |
|
413 |
(* export results *) |
|
414 |
||
11816 | 415 |
fun refine_tac rule = |
416 |
Tactic.rtac rule THEN_ALL_NEW (SUBGOAL (fn (goal, i) => |
|
417 |
if can Logic.dest_goal (Logic.strip_assums_concl goal) then |
|
418 |
Tactic.etac Drule.triv_goal i |
|
419 |
else all_tac)); |
|
420 |
||
6982 | 421 |
fun export_goal print_rule raw_rule inner state = |
6932 | 422 |
let |
10553 | 423 |
val (outer, (_, ((result, (facts, thm)), f))) = find_goal state; |
12010 | 424 |
val exp_tac = ProofContext.export true inner outer; |
9469 | 425 |
fun exp_rule rule = |
426 |
let |
|
427 |
val _ = print_rule rule; |
|
11816 | 428 |
val thmq = FINDGOAL (refine_tac rule) thm; |
9469 | 429 |
in Seq.map (fn th => map_goal I (K ((result, (facts, th)), f)) state) thmq end; |
430 |
in raw_rule |> exp_tac |> (Seq.flat o Seq.map exp_rule) end; |
|
6932 | 431 |
|
432 |
fun transfer_facts inner_state state = |
|
433 |
(case get_facts inner_state of |
|
434 |
None => Seq.single (reset_facts state) |
|
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
435 |
| Some thms => |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
436 |
let |
12010 | 437 |
val exp_tac = ProofContext.export false (context_of inner_state) (context_of state); |
11816 | 438 |
val thmqs = map exp_tac thms; |
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
439 |
in Seq.map (fn ths => put_facts (Some ths) state) (Seq.commute thmqs) end); |
6932 | 440 |
|
441 |
||
442 |
(* prepare result *) |
|
443 |
||
444 |
fun prep_result state t raw_thm = |
|
5820 | 445 |
let |
446 |
val ctxt = context_of state; |
|
447 |
fun err msg = raise STATE (msg, state); |
|
448 |
||
449 |
val ngoals = Thm.nprems_of raw_thm; |
|
450 |
val _ = |
|
451 |
if ngoals = 0 then () |
|
11891 | 452 |
else (Display.print_goals ngoals raw_thm; err (string_of_int ngoals ^ " unsolved goal(s)!")); |
5820 | 453 |
|
6932 | 454 |
val thm = raw_thm RS Drule.rev_triv_goal; |
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
455 |
val {hyps, prop, sign, ...} = Thm.rep_thm thm; |
5820 | 456 |
val tsig = Sign.tsig_of sign; |
6932 | 457 |
|
9469 | 458 |
val casms = flat (map #1 (assumptions state)); |
11816 | 459 |
val bad_hyps = Library.gen_rems Term.aconv (hyps, map Thm.term_of casms); |
5820 | 460 |
in |
6932 | 461 |
if not (null bad_hyps) then |
462 |
err ("Additional hypotheses:\n" ^ cat_lines (map (Sign.string_of_term sign) bad_hyps)) |
|
10553 | 463 |
else if not (Pattern.matches (Sign.tsig_of (sign_of state)) (t, prop)) then |
7605 | 464 |
err ("Proved a different theorem: " ^ Sign.string_of_term sign prop) |
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
465 |
else thm |
5820 | 466 |
end; |
467 |
||
468 |
||
469 |
||
470 |
(*** structured proof commands ***) |
|
471 |
||
472 |
(** context **) |
|
473 |
||
474 |
(* bind *) |
|
475 |
||
476 |
fun gen_bind f x state = |
|
477 |
state |
|
478 |
|> assert_forward |
|
479 |
|> map_context (f x) |
|
480 |
|> reset_facts; |
|
481 |
||
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
482 |
val match_bind = gen_bind (ProofContext.match_bind false); |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
483 |
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
|
484 |
val let_bind = gen_bind (ProofContext.match_bind true); |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
485 |
val let_bind_i = gen_bind (ProofContext.match_bind_i true); |
5820 | 486 |
|
487 |
||
9196 | 488 |
(* have_thmss etc. *) |
5820 | 489 |
|
9196 | 490 |
fun have_thmss args state = |
5820 | 491 |
state |
492 |
|> assert_forward |
|
9196 | 493 |
|> map_context_result (ProofContext.have_thmss args) |
494 |
|> these_factss; |
|
5820 | 495 |
|
9196 | 496 |
fun simple_have_thms name thms = have_thmss [((name, []), [(thms, [])])]; |
497 |
||
498 |
fun with_thmss args state = |
|
499 |
let val state' = state |> have_thmss args |
|
500 |
in state' |> simple_have_thms "" (the_facts state' @ the_facts state) end; |
|
6876 | 501 |
|
5820 | 502 |
|
503 |
(* fix *) |
|
504 |
||
505 |
fun gen_fix f xs state = |
|
506 |
state |
|
507 |
|> assert_forward |
|
508 |
|> map_context (f xs) |
|
509 |
|> reset_facts; |
|
510 |
||
511 |
val fix = gen_fix ProofContext.fix; |
|
512 |
val fix_i = gen_fix ProofContext.fix_i; |
|
513 |
||
514 |
||
11891 | 515 |
(* assume and presume *) |
5820 | 516 |
|
9469 | 517 |
fun gen_assume f exp args state = |
5820 | 518 |
state |
519 |
|> assert_forward |
|
9469 | 520 |
|> map_context_result (f exp args) |
11924 | 521 |
|> these_factss; |
6932 | 522 |
|
7271 | 523 |
val assm = gen_assume ProofContext.assume; |
524 |
val assm_i = gen_assume ProofContext.assume_i; |
|
11917 | 525 |
val assume = assm ProofContext.export_assume; |
526 |
val assume_i = assm_i ProofContext.export_assume; |
|
527 |
val presume = assm ProofContext.export_presume; |
|
528 |
val presume_i = assm_i ProofContext.export_presume; |
|
5820 | 529 |
|
7271 | 530 |
|
5820 | 531 |
|
11891 | 532 |
(** local definitions **) |
533 |
||
534 |
fun gen_def fix prep_term prep_pats raw_name atts (x, (raw_rhs, raw_pats)) state = |
|
535 |
let |
|
536 |
val _ = assert_forward state; |
|
537 |
val ctxt = context_of state; |
|
538 |
||
539 |
(*rhs*) |
|
540 |
val name = if raw_name = "" then Thm.def_name x else raw_name; |
|
541 |
val rhs = prep_term ctxt raw_rhs; |
|
542 |
val T = Term.fastype_of rhs; |
|
543 |
val pats = prep_pats T (ProofContext.declare_term rhs ctxt) raw_pats; |
|
544 |
||
545 |
(*lhs*) |
|
546 |
val lhs = ProofContext.bind_skolem ctxt [x] (Free (x, T)); |
|
547 |
in |
|
548 |
state |
|
549 |
|> fix [([x], None)] |
|
550 |
|> match_bind_i [(pats, rhs)] |
|
11917 | 551 |
|> assm_i ProofContext.export_def [((name, atts), [(Logic.mk_equals (lhs, rhs), ([], []))])] |
11891 | 552 |
end; |
553 |
||
554 |
val def = gen_def fix ProofContext.read_term ProofContext.read_term_pats; |
|
555 |
val def_i = gen_def fix_i ProofContext.cert_term ProofContext.cert_term_pats; |
|
556 |
||
557 |
||
8374 | 558 |
(* invoke_case *) |
559 |
||
11793
5f0ab6f5c280
support impromptu terminology of cases parameters;
wenzelm
parents:
11742
diff
changeset
|
560 |
fun invoke_case (name, xs, atts) state = |
9292 | 561 |
let |
10830 | 562 |
val (state', (lets, asms)) = |
11793
5f0ab6f5c280
support impromptu terminology of cases parameters;
wenzelm
parents:
11742
diff
changeset
|
563 |
map_context_result (ProofContext.apply_case (get_case state name xs)) state; |
9292 | 564 |
in |
10830 | 565 |
state' |
10809 | 566 |
|> add_binds_i lets |
10830 | 567 |
|> assume_i [((name, atts), map (rpair ([], [])) asms)] |
8374 | 568 |
end; |
569 |
||
570 |
||
5820 | 571 |
|
572 |
(** goals **) |
|
573 |
||
574 |
(* forward chaining *) |
|
575 |
||
576 |
fun chain state = |
|
577 |
state |
|
578 |
|> assert_forward |
|
6848 | 579 |
|> assert_facts |
5820 | 580 |
|> enter_forward_chain; |
581 |
||
582 |
fun from_facts facts state = |
|
583 |
state |
|
584 |
|> put_facts (Some facts) |
|
585 |
|> chain; |
|
586 |
||
587 |
||
588 |
(* setup goals *) |
|
589 |
||
11549 | 590 |
val rule_contextN = "rule_context"; |
8991 | 591 |
|
12010 | 592 |
fun setup_goal opt_block prepp autofix kind after_qed name raw_propp state = |
5820 | 593 |
let |
10464 | 594 |
val (state', ([[prop]], gen_binds)) = |
5936 | 595 |
state |
596 |
|> assert_forward_or_chain |
|
597 |
|> enter_forward |
|
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
598 |
|> opt_block |
10464 | 599 |
|> map_context_result (fn ct => prepp (ct, [[raw_propp]])); |
6932 | 600 |
val cprop = Thm.cterm_of (sign_of state') prop; |
7556 | 601 |
val goal = Drule.mk_triv_goal cprop; |
10553 | 602 |
|
603 |
val tvars = Term.term_tvars prop; |
|
604 |
val vars = Term.term_vars prop; |
|
5820 | 605 |
in |
10553 | 606 |
if null tvars then () |
607 |
else raise STATE ("Goal statement contains illegal schematic type variable(s): " ^ |
|
608 |
commas (map (Syntax.string_of_vname o #1) tvars), state); |
|
609 |
if null vars then () |
|
610 |
else warning ("Goal statement contains unbound schematic variable(s): " ^ |
|
611 |
commas (map (Sign.string_of_term (sign_of state)) vars)); |
|
5936 | 612 |
state' |
12010 | 613 |
|> map_context (autofix prop) |
614 |
|> put_goal (Some (((kind, name, prop), ([], goal)), after_qed o map_context gen_binds)) |
|
11549 | 615 |
|> map_context (ProofContext.add_cases (RuleCases.make true goal [rule_contextN])) |
6790 | 616 |
|> auto_bind_goal prop |
5820 | 617 |
|> (if is_chain state then use_facts else reset_facts) |
618 |
|> new_block |
|
619 |
|> enter_backward |
|
620 |
end; |
|
621 |
||
622 |
(*global goals*) |
|
12045 | 623 |
fun global_goal prepp act_locale act_elems kind locale elems name atts x thy = |
624 |
thy |> init_state |
|
625 |
|> open_block |> (case locale of Some loc => map_context (act_locale loc) | None => I) |
|
626 |
|> open_block |> map_context (act_elems elems) |
|
627 |
|> setup_goal I prepp (ProofContext.fix_frees o single) (Theorem (kind, locale, atts)) |
|
628 |
Seq.single name x; |
|
5820 | 629 |
|
12045 | 630 |
val theorem = global_goal ProofContext.bind_propp_schematic |
631 |
Locale.activate_locale Locale.activate_elements; |
|
632 |
||
633 |
val theorem_i = global_goal ProofContext.bind_propp_schematic_i |
|
634 |
Locale.activate_locale_i Locale.activate_elements_i; |
|
5820 | 635 |
|
636 |
||
637 |
(*local goals*) |
|
10938 | 638 |
fun local_goal' prepp kind (check: bool -> state -> state) |
639 |
f name atts args int state = |
|
7928 | 640 |
state |
12010 | 641 |
|> setup_goal open_block prepp (K I) (kind atts) f name args |
10936 | 642 |
|> warn_extra_tfrees state |
643 |
|> check int; |
|
5820 | 644 |
|
10936 | 645 |
fun local_goal prepp kind f name atts args = |
646 |
local_goal' prepp kind (K I) f name atts args false; |
|
647 |
||
648 |
val show = local_goal' ProofContext.bind_propp Show; |
|
649 |
val show_i = local_goal' ProofContext.bind_propp_i Show; |
|
650 |
val have = local_goal ProofContext.bind_propp Have; |
|
10380 | 651 |
val have_i = local_goal ProofContext.bind_propp_i Have; |
5820 | 652 |
|
653 |
||
654 |
||
655 |
(** conclusions **) |
|
656 |
||
657 |
(* current goal *) |
|
658 |
||
7176 | 659 |
fun current_goal (State (Node {context, goal = Some goal, ...}, _)) = (context, goal) |
5820 | 660 |
| current_goal state = raise STATE ("No current goal!", state); |
661 |
||
7176 | 662 |
fun assert_current_goal true (state as State (Node {goal = None, ...}, _)) = |
6404 | 663 |
raise STATE ("No goal in this block!", state) |
7176 | 664 |
| assert_current_goal false (state as State (Node {goal = Some _, ...}, _)) = |
6404 | 665 |
raise STATE ("Goal present in this block!", state) |
666 |
| assert_current_goal _ state = state; |
|
5820 | 667 |
|
12010 | 668 |
fun assert_bottom b (state as State (_, nodes)) = |
669 |
let val bot = (length nodes <= 2) in |
|
670 |
if b andalso not bot then raise STATE ("Not at bottom of proof!", state) |
|
671 |
else if not b andalso bot then raise STATE ("Already at bottom of proof!", state) |
|
672 |
else state |
|
673 |
end; |
|
5820 | 674 |
|
6404 | 675 |
val at_bottom = can (assert_bottom true o close_block); |
676 |
||
6932 | 677 |
fun end_proof bot state = |
5820 | 678 |
state |
679 |
|> assert_forward |
|
680 |
|> close_block |
|
681 |
|> assert_bottom bot |
|
7011 | 682 |
|> assert_current_goal true |
683 |
|> goal_facts (K []); |
|
5820 | 684 |
|
685 |
||
6404 | 686 |
(* local_qed *) |
5820 | 687 |
|
6982 | 688 |
fun finish_local (print_result, print_rule) state = |
5820 | 689 |
let |
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
690 |
val (goal_ctxt, (((kind, name, t), (_, raw_thm)), after_qed)) = current_goal state; |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
691 |
val outer_state = state |> close_block; |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
692 |
val outer_ctxt = context_of outer_state; |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
693 |
|
6932 | 694 |
val result = prep_result state t raw_thm; |
12010 | 695 |
val resultq = ProofContext.export false goal_ctxt outer_ctxt result; |
9469 | 696 |
|
5820 | 697 |
val (atts, opt_solve) = |
698 |
(case kind of |
|
10380 | 699 |
Show atts => (atts, export_goal print_rule result goal_ctxt) |
700 |
| Have atts => (atts, Seq.single) |
|
6932 | 701 |
| _ => err_malformed "finish_local" state); |
5820 | 702 |
in |
12015 | 703 |
print_result {kind = kind_name (sign_of state) kind, name = name, thm = result}; |
8617
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
704 |
outer_state |
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8582
diff
changeset
|
705 |
|> auto_bind_facts name [ProofContext.generalize goal_ctxt outer_ctxt t] |
9469 | 706 |
|> (fn st => Seq.map (fn res => |
707 |
have_thmss [((name, atts), [Thm.no_attributes [res]])] st) resultq) |
|
708 |
|> (Seq.flat o Seq.map opt_solve) |
|
7176 | 709 |
|> (Seq.flat o Seq.map after_qed) |
5820 | 710 |
end; |
711 |
||
6982 | 712 |
fun local_qed finalize print state = |
6404 | 713 |
state |
6932 | 714 |
|> end_proof false |
6871 | 715 |
|> finalize |
6982 | 716 |
|> (Seq.flat o Seq.map (finish_local print)); |
5820 | 717 |
|
718 |
||
6404 | 719 |
(* global_qed *) |
5820 | 720 |
|
6950 | 721 |
fun finish_global state = |
5820 | 722 |
let |
12010 | 723 |
val (goal_ctxt, (((kind, name, t), (_, raw_thm)), _)) = current_goal state; |
724 |
val full_name = if name = "" then "" else Sign.full_name (sign_of state) name; |
|
725 |
val locale_context = context_of (state |> close_block); (* FIXME unused *) |
|
726 |
val theory_context = context_of (state |> close_block |> close_block); |
|
727 |
||
728 |
val result = prep_result state t raw_thm; |
|
729 |
val resultq = |
|
730 |
ProofContext.export false goal_ctxt theory_context result |
|
731 |
|> Seq.map Drule.local_standard; |
|
5820 | 732 |
|
12010 | 733 |
val (atts, opt_locale) = |
734 |
(case kind of Theorem (s, loc, atts) => (atts @ [Drule.kind s], loc) |
|
6932 | 735 |
| _ => err_malformed "finish_global" state); |
12010 | 736 |
val thy = theory_of state; |
737 |
in |
|
738 |
resultq |> Seq.map (fn res => |
|
739 |
let val (thy', res') = PureThy.store_thm ((name, res), atts) thy |
|
12015 | 740 |
in (thy', {kind = kind_name (sign_of state) kind, name = name, thm = res'}) end) |
12010 | 741 |
end; |
5820 | 742 |
|
10553 | 743 |
(*note: clients should inspect first result only, as backtracking may destroy theory*) |
6950 | 744 |
fun global_qed finalize state = |
5820 | 745 |
state |
6932 | 746 |
|> end_proof true |
6871 | 747 |
|> finalize |
12010 | 748 |
|> Seq.map finish_global |
749 |
|> Seq.flat; |
|
5820 | 750 |
|
751 |
||
6896 | 752 |
|
753 |
(** blocks **) |
|
754 |
||
755 |
(* begin_block *) |
|
756 |
||
757 |
fun begin_block state = |
|
758 |
state |
|
759 |
|> assert_forward |
|
760 |
|> new_block |
|
761 |
|> open_block; |
|
762 |
||
763 |
||
764 |
(* end_block *) |
|
765 |
||
766 |
fun end_block state = |
|
767 |
state |
|
768 |
|> assert_forward |
|
769 |
|> close_block |
|
770 |
|> assert_current_goal false |
|
771 |
|> close_block |
|
6932 | 772 |
|> transfer_facts state; |
6896 | 773 |
|
774 |
||
775 |
(* next_block *) |
|
776 |
||
777 |
fun next_block state = |
|
778 |
state |
|
779 |
|> assert_forward |
|
780 |
|> close_block |
|
8718 | 781 |
|> new_block |
782 |
|> reset_facts; |
|
6896 | 783 |
|
784 |
||
5820 | 785 |
end; |
8152 | 786 |
|
787 |
structure BasicProof: BASIC_PROOF = Proof; |
|
788 |
open BasicProof; |