author | haftmann |
Sun, 23 Jul 2006 07:19:36 +0200 | |
changeset 20183 | fd546b0c8a7c |
parent 20157 | 28638d2a6bc7 |
child 20228 | e0f9e8a6556b |
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 |
|
11 |
end; |
|
12 |
||
13 |
signature GOAL = |
|
14 |
sig |
|
15 |
include BASIC_GOAL |
|
16 |
val init: cterm -> thm |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
17 |
val protect: thm -> thm |
17980 | 18 |
val conclude: thm -> thm |
19 |
val finish: thm -> thm |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
20 |
val norm_hhf: thm -> thm |
18252 | 21 |
val norm_hhf_protect: thm -> thm |
18119 | 22 |
val compose_hhf: thm -> int -> thm -> thm Seq.seq |
23 |
val compose_hhf_tac: thm -> int -> tactic |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
24 |
val comp_hhf: thm -> thm -> thm |
20056 | 25 |
val prove_multi: Context.proof -> string list -> term list -> term list -> |
17980 | 26 |
(thm list -> tactic) -> thm list |
20056 | 27 |
val prove: Context.proof -> string list -> term list -> term -> (thm list -> tactic) -> thm |
28 |
val prove_global: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
29 |
val prove_raw: cterm list -> cterm -> (thm list -> tactic) -> thm |
19184 | 30 |
val extract: int -> int -> thm -> thm Seq.seq |
31 |
val retrofit: int -> int -> thm -> thm -> thm Seq.seq |
|
17980 | 32 |
end; |
33 |
||
34 |
structure Goal: GOAL = |
|
35 |
struct |
|
36 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
37 |
(** goals **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
38 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
39 |
(* |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
40 |
-------- (init) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
41 |
C ==> #C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
42 |
*) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
43 |
fun init ct = Drule.instantiate' [] [SOME ct] Drule.protectI; |
17980 | 44 |
|
45 |
(* |
|
18119 | 46 |
C |
47 |
--- (protect) |
|
48 |
#C |
|
17980 | 49 |
*) |
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
50 |
fun protect th = th COMP Drule.incr_indexes th Drule.protectI; |
17980 | 51 |
|
52 |
(* |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
53 |
A ==> ... ==> #C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
54 |
---------------- (conclude) |
17980 | 55 |
A ==> ... ==> C |
56 |
*) |
|
57 |
fun conclude th = |
|
18497 | 58 |
(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
|
59 |
(Drule.incr_indexes th Drule.protectD) of |
17980 | 60 |
SOME th' => th' |
61 |
| NONE => raise THM ("Failed to conclude goal", 0, [th])); |
|
62 |
||
63 |
(* |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
64 |
#C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
65 |
--- (finish) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
66 |
C |
17983 | 67 |
*) |
17980 | 68 |
fun finish th = |
69 |
(case Thm.nprems_of th of |
|
70 |
0 => conclude th |
|
71 |
| n => raise THM ("Proof failed.\n" ^ |
|
72 |
Pretty.string_of (Pretty.chunks (Display.pretty_goals n th)) ^ |
|
73 |
("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th])); |
|
74 |
||
75 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
76 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
77 |
(** results **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
78 |
|
18207 | 79 |
(* HHF normal form: !! before ==>, outermost !! generalized *) |
80 |
||
81 |
local |
|
17980 | 82 |
|
18207 | 83 |
fun gen_norm_hhf ss = |
84 |
(not o Drule.is_norm_hhf o Thm.prop_of) ? |
|
85 |
Drule.fconv_rule (MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE)) ss) |
|
86 |
#> Thm.adjust_maxidx_thm |
|
87 |
#> Drule.gen_all; |
|
88 |
||
89 |
val ss = |
|
18180 | 90 |
MetaSimplifier.theory_context ProtoPure.thy MetaSimplifier.empty_ss |
91 |
addsimps [Drule.norm_hhf_eq]; |
|
92 |
||
18207 | 93 |
in |
18180 | 94 |
|
18207 | 95 |
val norm_hhf = gen_norm_hhf ss; |
18252 | 96 |
val norm_hhf_protect = gen_norm_hhf (ss addeqcongs [Drule.protect_cong]); |
17980 | 97 |
|
18207 | 98 |
end; |
18180 | 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 |
(* composition of normal results *) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
102 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
103 |
fun compose_hhf tha i thb = |
18145 | 104 |
Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of thb i) tha, 0) i thb; |
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
105 |
|
18119 | 106 |
fun compose_hhf_tac th i = PRIMSEQ (compose_hhf th i); |
107 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
108 |
fun comp_hhf tha thb = |
19862 | 109 |
(case Seq.chop 2 (compose_hhf tha 1 thb) of |
18119 | 110 |
([th], _) => th |
111 |
| ([], _) => raise THM ("comp_hhf: no unifiers", 1, [tha, thb]) |
|
112 |
| _ => raise THM ("comp_hhf: multiple unifiers", 1, [tha, thb])); |
|
17986 | 113 |
|
114 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
115 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
116 |
(** tactical theorem proving **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
117 |
|
18119 | 118 |
(* prove_multi *) |
17986 | 119 |
|
20056 | 120 |
fun prove_multi ctxt xs asms props tac = |
17980 | 121 |
let |
20056 | 122 |
val thy = Context.theory_of_proof ctxt; |
123 |
val string_of_term = Sign.string_of_term thy; |
|
124 |
||
17980 | 125 |
val prop = Logic.mk_conjunction_list props; |
126 |
val statement = Logic.list_implies (asms, prop); |
|
127 |
||
18678 | 128 |
fun err msg = cat_error msg |
20056 | 129 |
("The error(s) above occurred for the goal statement:\n" ^ string_of_term statement); |
17980 | 130 |
fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm t) |
131 |
handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg; |
|
132 |
||
133 |
val _ = cert_safe statement; |
|
134 |
val _ = Term.no_dummy_patterns statement handle TERM (msg, _) => err msg; |
|
135 |
val casms = map cert_safe asms; |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
136 |
val prems = map (norm_hhf o Thm.assume) casms; |
17980 | 137 |
|
20056 | 138 |
val ctxt' = ctxt |
139 |
|> Variable.set_body false |
|
140 |
|> (snd o Variable.add_fixes xs) |
|
20157
28638d2a6bc7
prove: Variable.declare_internal (more efficient);
wenzelm
parents:
20056
diff
changeset
|
141 |
|> fold Variable.declare_internal (asms @ props); |
20056 | 142 |
|
19774
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
143 |
val res = |
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
144 |
(case SINGLE (tac prems) (init (cert_safe prop)) of |
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
145 |
NONE => err "Tactic failed." |
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
146 |
| SOME res => res); |
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
147 |
val [results] = |
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents:
19619
diff
changeset
|
148 |
Conjunction.elim_precise [length props] (finish res) handle THM (msg, _, _) => err msg; |
19862 | 149 |
val _ = Unify.matches_list thy (map (Thm.term_of o cert_safe) props) (map Thm.prop_of results) |
20056 | 150 |
orelse err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res)); |
17980 | 151 |
in |
20056 | 152 |
results |
153 |
|> map (Drule.implies_intr_list casms) |
|
154 |
|> Variable.export ctxt' ctxt |
|
155 |
|> map (norm_hhf #> Drule.zero_var_indexes) |
|
17980 | 156 |
end; |
157 |
||
158 |
||
18119 | 159 |
(* prove *) |
17980 | 160 |
|
20056 | 161 |
fun prove ctxt xs asms prop tac = hd (prove_multi ctxt xs asms [prop] tac); |
162 |
||
163 |
fun prove_global thy xs asms prop tac = |
|
164 |
Drule.standard (prove (Context.init_proof thy) xs asms prop tac); |
|
17980 | 165 |
|
166 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
167 |
(* prove_raw -- no checks, no normalization of result! *) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
168 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
169 |
fun prove_raw casms cprop tac = |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
170 |
(case SINGLE (tac (map (norm_hhf o Thm.assume) casms)) (init cprop) of |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
171 |
SOME th => Drule.implies_intr_list casms (finish th) |
18678 | 172 |
| NONE => error "Tactic failed."); |
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
173 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
174 |
|
17980 | 175 |
|
19184 | 176 |
(** local goal states **) |
18207 | 177 |
|
19184 | 178 |
fun extract i n st = |
179 |
(if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty |
|
180 |
else if n = 1 then Seq.single (Thm.cprem_of st i) |
|
19423 | 181 |
else Seq.single (foldr1 Conjunction.mk_conjunction (map (Thm.cprem_of st) (i upto i + n - 1)))) |
19184 | 182 |
|> Seq.map (Thm.adjust_maxidx #> init); |
17980 | 183 |
|
19221 | 184 |
fun retrofit i n st' st = |
185 |
(if n = 1 then st |
|
19423 | 186 |
else st |> Drule.rotate_prems (i - 1) |> Conjunction.uncurry n |> Drule.rotate_prems (1 - i)) |
19221 | 187 |
|> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i; |
18207 | 188 |
|
17980 | 189 |
fun SELECT_GOAL tac i st = |
19191 | 190 |
if Thm.nprems_of st = 1 andalso i = 1 then tac st |
19184 | 191 |
else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st; |
17980 | 192 |
|
18207 | 193 |
end; |
194 |
||
17980 | 195 |
structure BasicGoal: BASIC_GOAL = Goal; |
196 |
open BasicGoal; |