author | wenzelm |
Sat, 12 Apr 2008 17:00:35 +0200 | |
changeset 26626 | c6231d64d264 |
parent 25301 | 24e027f55f45 |
child 26628 | 63306cb94313 |
permissions | -rw-r--r-- |
17980 | 1 |
(* Title: Pure/goal.ML |
2 |
ID: $Id$ |
|
3 |
Author: Makarius and Lawrence C Paulson |
|
4 |
||
18139 | 5 |
Goals in tactical theorem proving. |
17980 | 6 |
*) |
7 |
||
8 |
signature BASIC_GOAL = |
|
9 |
sig |
|
10 |
val SELECT_GOAL: tactic -> int -> tactic |
|
23418 | 11 |
val CONJUNCTS: tactic -> int -> tactic |
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
12 |
val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic |
17980 | 13 |
end; |
14 |
||
15 |
signature GOAL = |
|
16 |
sig |
|
17 |
include BASIC_GOAL |
|
18 |
val init: cterm -> thm |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
19 |
val protect: thm -> thm |
17980 | 20 |
val conclude: thm -> thm |
21 |
val finish: thm -> thm |
|
21604
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
22 |
val norm_result: thm -> thm |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
23 |
val close_result: thm -> thm |
23356 | 24 |
val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm |
20290 | 25 |
val prove_multi: Proof.context -> string list -> term list -> term list -> |
26 |
({prems: thm list, context: Proof.context} -> tactic) -> thm list |
|
27 |
val prove: Proof.context -> string list -> term list -> term -> |
|
28 |
({prems: thm list, context: Proof.context} -> tactic) -> thm |
|
20056 | 29 |
val prove_global: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm |
19184 | 30 |
val extract: int -> int -> thm -> thm Seq.seq |
31 |
val retrofit: int -> int -> thm -> thm -> thm Seq.seq |
|
23418 | 32 |
val conjunction_tac: int -> tactic |
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
33 |
val precise_conjunction_tac: int -> int -> tactic |
23418 | 34 |
val recover_conjunction_tac: tactic |
21687 | 35 |
val norm_hhf_tac: int -> tactic |
36 |
val compose_hhf_tac: thm -> int -> tactic |
|
23237 | 37 |
val assume_rule_tac: Proof.context -> int -> tactic |
17980 | 38 |
end; |
39 |
||
40 |
structure Goal: GOAL = |
|
41 |
struct |
|
42 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
43 |
(** goals **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
44 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
45 |
(* |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
46 |
-------- (init) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
47 |
C ==> #C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
48 |
*) |
20290 | 49 |
val init = |
22902
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents:
21687
diff
changeset
|
50 |
let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI)) |
20290 | 51 |
in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end; |
17980 | 52 |
|
53 |
(* |
|
18119 | 54 |
C |
55 |
--- (protect) |
|
56 |
#C |
|
17980 | 57 |
*) |
21579 | 58 |
fun protect th = th COMP_INCR Drule.protectI; |
17980 | 59 |
|
60 |
(* |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
61 |
A ==> ... ==> #C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
62 |
---------------- (conclude) |
17980 | 63 |
A ==> ... ==> C |
64 |
*) |
|
65 |
fun conclude th = |
|
18497 | 66 |
(case SINGLE (Thm.compose_no_flatten false (th, Thm.nprems_of th) 1) |
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
67 |
(Drule.incr_indexes th Drule.protectD) of |
17980 | 68 |
SOME th' => th' |
69 |
| NONE => raise THM ("Failed to conclude goal", 0, [th])); |
|
70 |
||
71 |
(* |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
72 |
#C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
73 |
--- (finish) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
74 |
C |
17983 | 75 |
*) |
17980 | 76 |
fun finish th = |
77 |
(case Thm.nprems_of th of |
|
78 |
0 => conclude th |
|
79 |
| n => raise THM ("Proof failed.\n" ^ |
|
80 |
Pretty.string_of (Pretty.chunks (Display.pretty_goals n th)) ^ |
|
81 |
("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th])); |
|
82 |
||
83 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
84 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
85 |
(** results **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
86 |
|
21604
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
87 |
(* normal form *) |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
88 |
|
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
89 |
val norm_result = |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
90 |
Drule.flexflex_unique |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
91 |
#> MetaSimplifier.norm_hhf_protect |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
92 |
#> Thm.strip_shyps |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
93 |
#> Drule.zero_var_indexes; |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
94 |
|
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
95 |
val close_result = |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
96 |
Thm.compress |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
97 |
#> Drule.close_derivation; |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
98 |
|
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
99 |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
100 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
101 |
(** tactical theorem proving **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
102 |
|
23356 | 103 |
(* prove_internal -- minimal checks, no normalization of result! *) |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
104 |
|
23356 | 105 |
fun prove_internal casms cprop tac = |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
106 |
(case SINGLE (tac (map Assumption.assume casms)) (init cprop) of |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
107 |
SOME th => Drule.implies_intr_list casms (finish th) |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
108 |
| NONE => error "Tactic failed."); |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
109 |
|
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
110 |
|
18119 | 111 |
(* prove_multi *) |
17986 | 112 |
|
20056 | 113 |
fun prove_multi ctxt xs asms props tac = |
17980 | 114 |
let |
21516 | 115 |
val thy = ProofContext.theory_of ctxt; |
20056 | 116 |
val string_of_term = Sign.string_of_term thy; |
117 |
||
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
118 |
fun err msg = cat_error msg |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
119 |
("The error(s) above occurred for the goal statement:\n" ^ |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
120 |
string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props))); |
17980 | 121 |
|
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
122 |
fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t)) |
17980 | 123 |
handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg; |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
124 |
val casms = map cert_safe asms; |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
125 |
val cprops = map cert_safe props; |
17980 | 126 |
|
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
127 |
val (prems, ctxt') = ctxt |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
128 |
|> Variable.add_fixes_direct xs |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
129 |
|> fold Variable.declare_internal (asms @ props) |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
130 |
|> Assumption.add_assumes casms; |
17980 | 131 |
|
23418 | 132 |
val goal = init (Conjunction.mk_conjunction_balanced cprops); |
19774
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
133 |
val res = |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
134 |
(case SINGLE (tac {prems = prems, context = ctxt'}) goal of |
19774
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
135 |
NONE => err "Tactic failed." |
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
136 |
| SOME res => res); |
23418 | 137 |
val results = Conjunction.elim_balanced (length props) (finish res) |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
138 |
handle THM (msg, _, _) => err msg; |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
139 |
val _ = Unify.matches_list thy (map Thm.term_of cprops) (map Thm.prop_of results) |
20056 | 140 |
orelse err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res)); |
17980 | 141 |
in |
20056 | 142 |
results |
20290 | 143 |
|> map (Assumption.export false ctxt' ctxt) |
20056 | 144 |
|> Variable.export ctxt' ctxt |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
145 |
|> map Drule.zero_var_indexes |
17980 | 146 |
end; |
147 |
||
148 |
||
18119 | 149 |
(* prove *) |
17980 | 150 |
|
20056 | 151 |
fun prove ctxt xs asms prop tac = hd (prove_multi ctxt xs asms [prop] tac); |
152 |
||
153 |
fun prove_global thy xs asms prop tac = |
|
21516 | 154 |
Drule.standard (prove (ProofContext.init thy) xs asms prop (fn {prems, ...} => tac prems)); |
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
155 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
156 |
|
17980 | 157 |
|
21687 | 158 |
(** goal structure **) |
159 |
||
160 |
(* nested goals *) |
|
18207 | 161 |
|
19184 | 162 |
fun extract i n st = |
163 |
(if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty |
|
164 |
else if n = 1 then Seq.single (Thm.cprem_of st i) |
|
23418 | 165 |
else |
166 |
Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1)))) |
|
20260 | 167 |
|> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init); |
17980 | 168 |
|
19221 | 169 |
fun retrofit i n st' st = |
170 |
(if n = 1 then st |
|
23418 | 171 |
else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n)) |
19221 | 172 |
|> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i; |
18207 | 173 |
|
17980 | 174 |
fun SELECT_GOAL tac i st = |
19191 | 175 |
if Thm.nprems_of st = 1 andalso i = 1 then tac st |
19184 | 176 |
else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st; |
17980 | 177 |
|
21687 | 178 |
|
179 |
(* multiple goals *) |
|
180 |
||
23418 | 181 |
fun precise_conjunction_tac 0 i = eq_assume_tac i |
182 |
| precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i |
|
183 |
| precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n)); |
|
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
184 |
|
23418 | 185 |
val adhoc_conjunction_tac = REPEAT_ALL_NEW |
186 |
(SUBGOAL (fn (goal, i) => |
|
187 |
if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i |
|
188 |
else no_tac)); |
|
21687 | 189 |
|
23418 | 190 |
val conjunction_tac = SUBGOAL (fn (goal, i) => |
191 |
precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE |
|
192 |
TRY (adhoc_conjunction_tac i)); |
|
21687 | 193 |
|
23418 | 194 |
val recover_conjunction_tac = PRIMITIVE (fn th => |
195 |
Conjunction.uncurry_balanced (Thm.nprems_of th) th); |
|
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
196 |
|
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
197 |
fun PRECISE_CONJUNCTS n tac = |
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
198 |
SELECT_GOAL (precise_conjunction_tac n 1 |
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
199 |
THEN tac |
23418 | 200 |
THEN recover_conjunction_tac); |
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
201 |
|
21687 | 202 |
fun CONJUNCTS tac = |
203 |
SELECT_GOAL (conjunction_tac 1 |
|
204 |
THEN tac |
|
23418 | 205 |
THEN recover_conjunction_tac); |
21687 | 206 |
|
207 |
||
208 |
(* hhf normal form *) |
|
209 |
||
210 |
val norm_hhf_tac = |
|
211 |
rtac Drule.asm_rl (*cheap approximation -- thanks to builtin Logic.flatten_params*) |
|
212 |
THEN' SUBGOAL (fn (t, i) => |
|
213 |
if Drule.is_norm_hhf t then all_tac |
|
23536
60a1672e298e
moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
wenzelm
parents:
23418
diff
changeset
|
214 |
else MetaSimplifier.rewrite_goal_tac [Drule.norm_hhf_eq] i); |
21687 | 215 |
|
25301 | 216 |
fun compose_hhf_tac th i st = |
217 |
PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st; |
|
21687 | 218 |
|
23237 | 219 |
|
220 |
(* non-atomic goal assumptions *) |
|
221 |
||
23356 | 222 |
fun non_atomic (Const ("==>", _) $ _ $ _) = true |
223 |
| non_atomic (Const ("all", _) $ _) = true |
|
224 |
| non_atomic _ = false; |
|
225 |
||
23237 | 226 |
fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) => |
227 |
let |
|
228 |
val ((_, goal'), ctxt') = Variable.focus goal ctxt; |
|
229 |
val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; |
|
23356 | 230 |
val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal''); |
23237 | 231 |
val tacs = Rs |> map (fn R => |
232 |
Tactic.etac (MetaSimplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac); |
|
233 |
in fold_rev (curry op APPEND') tacs (K no_tac) i end); |
|
234 |
||
18207 | 235 |
end; |
236 |
||
17980 | 237 |
structure BasicGoal: BASIC_GOAL = Goal; |
238 |
open BasicGoal; |