author | urbanc |
Thu, 22 Mar 2007 14:26:05 +0100 | |
changeset 22500 | 8436bfd21bf3 |
parent 22147 | f4ed4d940d44 |
child 22838 | 466599ecf610 |
permissions | -rw-r--r-- |
21163 | 1 |
(* Title: HOL/simpdata.ML |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1991 University of Cambridge |
|
5 |
||
6 |
Instantiation of the generic simplifier for HOL. |
|
7 |
*) |
|
8 |
||
9 |
(** tools setup **) |
|
10 |
||
11 |
structure Quantifier1 = Quantifier1Fun |
|
12 |
(struct |
|
13 |
(*abstract syntax*) |
|
14 |
fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t) |
|
15 |
| dest_eq _ = NONE; |
|
16 |
fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t) |
|
17 |
| dest_conj _ = NONE; |
|
18 |
fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t) |
|
19 |
| dest_imp _ = NONE; |
|
20 |
val conj = HOLogic.conj |
|
21 |
val imp = HOLogic.imp |
|
22 |
(*rules*) |
|
22147 | 23 |
val iff_reflection = @{thm eq_reflection} |
24 |
val iffI = @{thm iffI} |
|
25 |
val iff_trans = @{thm trans} |
|
26 |
val conjI= @{thm conjI} |
|
27 |
val conjE= @{thm conjE} |
|
28 |
val impI = @{thm impI} |
|
29 |
val mp = @{thm mp} |
|
30 |
val uncurry = @{thm uncurry} |
|
31 |
val exI = @{thm exI} |
|
32 |
val exE = @{thm exE} |
|
33 |
val iff_allI = @{thm iff_allI} |
|
34 |
val iff_exI = @{thm iff_exI} |
|
35 |
val all_comm = @{thm all_comm} |
|
36 |
val ex_comm = @{thm ex_comm} |
|
21163 | 37 |
end); |
38 |
||
21551 | 39 |
structure Simpdata = |
21163 | 40 |
struct |
41 |
||
22147 | 42 |
fun mk_meta_eq r = r RS @{thm eq_reflection}; |
21163 | 43 |
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; |
44 |
||
22147 | 45 |
fun mk_eq th = case concl_of th |
21163 | 46 |
(*expects Trueprop if not == *) |
21551 | 47 |
of Const ("==",_) $ _ $ _ => th |
48 |
| _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq th |
|
22147 | 49 |
| _ $ (Const ("Not", _) $ _) => th RS @{thm Eq_FalseI} |
50 |
| _ => th RS @{thm Eq_TrueI} |
|
21163 | 51 |
|
22147 | 52 |
fun mk_eq_True r = |
53 |
SOME (r RS @{thm meta_eq_to_obj_eq} RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE; |
|
21163 | 54 |
|
55 |
(* Produce theorems of the form |
|
56 |
(P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y) |
|
57 |
*) |
|
22147 | 58 |
fun lift_meta_eq_to_obj_eq i st = |
21163 | 59 |
let |
60 |
fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q |
|
61 |
| count_imp _ = 0; |
|
62 |
val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) |
|
22147 | 63 |
in if j = 0 then @{thm meta_eq_to_obj_eq} |
21163 | 64 |
else |
65 |
let |
|
66 |
val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); |
|
67 |
fun mk_simp_implies Q = foldr (fn (R, S) => |
|
68 |
Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps |
|
69 |
val aT = TFree ("'a", HOLogic.typeS); |
|
70 |
val x = Free ("x", aT); |
|
71 |
val y = Free ("y", aT) |
|
72 |
in Goal.prove_global (Thm.theory_of_thm st) [] |
|
73 |
[mk_simp_implies (Logic.mk_equals (x, y))] |
|
74 |
(mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y)))) |
|
75 |
(fn prems => EVERY |
|
22147 | 76 |
[rewrite_goals_tac @{thms simp_implies_def}, |
77 |
REPEAT (ares_tac (@{thm meta_eq_to_obj_eq} :: |
|
78 |
map (rewrite_rule @{thms simp_implies_def}) prems) 1)]) |
|
21163 | 79 |
end |
80 |
end; |
|
81 |
||
82 |
(*Congruence rules for = (instead of ==)*) |
|
83 |
fun mk_meta_cong rl = zero_var_indexes |
|
84 |
(let val rl' = Seq.hd (TRYALL (fn i => fn st => |
|
85 |
rtac (lift_meta_eq_to_obj_eq i st) i st) rl) |
|
86 |
in mk_meta_eq rl' handle THM _ => |
|
87 |
if can Logic.dest_equals (concl_of rl') then rl' |
|
88 |
else error "Conclusion of congruence rules must be =-equality" |
|
89 |
end); |
|
90 |
||
91 |
fun mk_atomize pairs = |
|
92 |
let |
|
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
93 |
fun atoms thm = |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
94 |
let |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
95 |
fun res th = map (fn rl => th RS rl); (*exception THM*) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
96 |
fun res_fixed rls = |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
97 |
if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
98 |
else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.thm_context thm) [thm]; |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
99 |
in |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
100 |
case concl_of thm |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
101 |
of Const ("Trueprop", _) $ p => (case head_of p |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
102 |
of Const (a, _) => (case AList.lookup (op =) pairs a |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
103 |
of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm]) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
104 |
| NONE => [thm]) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
105 |
| _ => [thm]) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
106 |
| _ => [thm] |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
107 |
end; |
21163 | 108 |
in atoms end; |
109 |
||
110 |
fun mksimps pairs = |
|
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
111 |
map_filter (try mk_eq) o mk_atomize pairs o gen_all; |
21163 | 112 |
|
22147 | 113 |
fun unsafe_solver_tac prems = |
114 |
(fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' |
|
115 |
FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), atac, |
|
116 |
etac @{thm FalseE}]; |
|
117 |
||
21163 | 118 |
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; |
119 |
||
120 |
(*No premature instantiation of variables during simplification*) |
|
22147 | 121 |
fun safe_solver_tac prems = |
122 |
(fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' |
|
123 |
FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), |
|
124 |
eq_assume_tac, ematch_tac @{thms FalseE}]; |
|
125 |
||
21163 | 126 |
val safe_solver = mk_solver "HOL safe" safe_solver_tac; |
127 |
||
128 |
structure SplitterData = |
|
129 |
struct |
|
130 |
structure Simplifier = Simplifier |
|
21551 | 131 |
val mk_eq = mk_eq |
22147 | 132 |
val meta_eq_to_iff = @{thm meta_eq_to_obj_eq} |
133 |
val iffD = @{thm iffD2} |
|
134 |
val disjE = @{thm disjE} |
|
135 |
val conjE = @{thm conjE} |
|
136 |
val exE = @{thm exE} |
|
137 |
val contrapos = @{thm contrapos_nn} |
|
138 |
val contrapos2 = @{thm contrapos_pp} |
|
139 |
val notnotD = @{thm notnotD} |
|
21163 | 140 |
end; |
141 |
||
142 |
structure Splitter = SplitterFun(SplitterData); |
|
143 |
||
21674 | 144 |
val split_tac = Splitter.split_tac; |
145 |
val split_inside_tac = Splitter.split_inside_tac; |
|
146 |
||
147 |
val op addsplits = Splitter.addsplits; |
|
148 |
val op delsplits = Splitter.delsplits; |
|
149 |
val Addsplits = Splitter.Addsplits; |
|
150 |
val Delsplits = Splitter.Delsplits; |
|
151 |
||
152 |
||
21163 | 153 |
(* integration of simplifier with classical reasoner *) |
154 |
||
155 |
structure Clasimp = ClasimpFun |
|
156 |
(structure Simplifier = Simplifier and Splitter = Splitter |
|
157 |
and Classical = Classical and Blast = Blast |
|
22147 | 158 |
val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE}); |
21674 | 159 |
open Clasimp; |
21163 | 160 |
|
22128 | 161 |
val _ = ML_Context.value_antiq "clasimpset" |
162 |
(Scan.succeed ("clasimpset", "Clasimp.local_clasimpset_of (ML_Context.the_local_context ())")); |
|
163 |
||
21163 | 164 |
val mksimps_pairs = |
22147 | 165 |
[("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]), |
166 |
("All", [@{thm spec}]), ("True", []), ("False", []), |
|
167 |
("HOL.If", [@{thm if_bool_eq_conj} RS @{thm iffD1}])]; |
|
21163 | 168 |
|
21674 | 169 |
val HOL_basic_ss = |
22147 | 170 |
Simplifier.theory_context @{theory} empty_ss |
21163 | 171 |
setsubgoaler asm_simp_tac |
172 |
setSSolver safe_solver |
|
173 |
setSolver unsafe_solver |
|
174 |
setmksimps (mksimps mksimps_pairs) |
|
175 |
setmkeqTrue mk_eq_True |
|
176 |
setmkcong mk_meta_cong; |
|
177 |
||
21674 | 178 |
fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); |
21163 | 179 |
|
180 |
fun unfold_tac ths = |
|
21674 | 181 |
let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths |
21163 | 182 |
in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; |
183 |
||
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
184 |
|
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
185 |
|
21163 | 186 |
(** simprocs **) |
187 |
||
188 |
(* simproc for proving "(y = x) == False" from premise "~(x = y)" *) |
|
189 |
||
190 |
val use_neq_simproc = ref true; |
|
191 |
||
192 |
local |
|
22147 | 193 |
val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI}; |
21163 | 194 |
fun neq_prover sg ss (eq $ lhs $ rhs) = |
195 |
let |
|
196 |
fun test thm = (case #prop (rep_thm thm) of |
|
197 |
_ $ (Not $ (eq' $ l' $ r')) => |
|
198 |
Not = HOLogic.Not andalso eq' = eq andalso |
|
199 |
r' aconv lhs andalso l' aconv rhs |
|
200 |
| _ => false) |
|
201 |
in if !use_neq_simproc then case find_first test (prems_of_ss ss) |
|
202 |
of NONE => NONE |
|
203 |
| SOME thm => SOME (thm RS neq_to_EQ_False) |
|
204 |
else NONE |
|
205 |
end |
|
206 |
in |
|
207 |
||
22147 | 208 |
val neq_simproc = Simplifier.simproc @{theory} "neq_simproc" ["x = y"] neq_prover; |
21163 | 209 |
|
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
210 |
end; |
21163 | 211 |
|
212 |
||
213 |
(* simproc for Let *) |
|
214 |
||
215 |
val use_let_simproc = ref true; |
|
216 |
||
217 |
local |
|
218 |
val (f_Let_unfold, x_Let_unfold) = |
|
22147 | 219 |
let val [(_$(f$x)$_)] = prems_of @{thm Let_unfold} |
220 |
in (cterm_of @{theory} f, cterm_of @{theory} x) end |
|
21163 | 221 |
val (f_Let_folded, x_Let_folded) = |
22147 | 222 |
let val [(_$(f$x)$_)] = prems_of @{thm Let_folded} |
223 |
in (cterm_of @{theory} f, cterm_of @{theory} x) end; |
|
21163 | 224 |
val g_Let_folded = |
22147 | 225 |
let val [(_$_$(g$_))] = prems_of @{thm Let_folded} in cterm_of @{theory} g end; |
21163 | 226 |
in |
227 |
||
228 |
val let_simproc = |
|
22147 | 229 |
Simplifier.simproc @{theory} "let_simp" ["Let x f"] |
21163 | 230 |
(fn sg => fn ss => fn t => |
231 |
let val ctxt = Simplifier.the_context ss; |
|
232 |
val ([t'], ctxt') = Variable.import_terms false [t] ctxt; |
|
233 |
in Option.map (hd o Variable.export ctxt' ctxt o single) |
|
234 |
(case t' of (Const ("Let",_)$x$f) => (* x and f are already in normal form *) |
|
235 |
if not (!use_let_simproc) then NONE |
|
236 |
else if is_Free x orelse is_Bound x orelse is_Const x |
|
22147 | 237 |
then SOME @{thm Let_def} |
21163 | 238 |
else |
239 |
let |
|
240 |
val n = case f of (Abs (x,_,_)) => x | _ => "x"; |
|
241 |
val cx = cterm_of sg x; |
|
242 |
val {T=xT,...} = rep_cterm cx; |
|
243 |
val cf = cterm_of sg f; |
|
244 |
val fx_g = Simplifier.rewrite ss (Thm.capply cf cx); |
|
245 |
val (_$_$g) = prop_of fx_g; |
|
246 |
val g' = abstract_over (x,g); |
|
247 |
in (if (g aconv g') |
|
248 |
then |
|
249 |
let |
|
22147 | 250 |
val rl = |
251 |
cterm_instantiate [(f_Let_unfold,cf),(x_Let_unfold,cx)] @{thm Let_unfold}; |
|
21163 | 252 |
in SOME (rl OF [fx_g]) end |
253 |
else if Term.betapply (f,x) aconv g then NONE (*avoid identity conversion*) |
|
254 |
else let |
|
255 |
val abs_g'= Abs (n,xT,g'); |
|
256 |
val g'x = abs_g'$x; |
|
257 |
val g_g'x = symmetric (beta_conversion false (cterm_of sg g'x)); |
|
258 |
val rl = cterm_instantiate |
|
259 |
[(f_Let_folded,cterm_of sg f),(x_Let_folded,cx), |
|
260 |
(g_Let_folded,cterm_of sg abs_g')] |
|
22147 | 261 |
@{thm Let_folded}; |
21163 | 262 |
in SOME (rl OF [transitive fx_g g_g'x]) |
263 |
end) |
|
264 |
end |
|
265 |
| _ => NONE) |
|
266 |
end) |
|
267 |
||
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
268 |
end; |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
269 |
|
21163 | 270 |
|
271 |
(* generic refutation procedure *) |
|
272 |
||
273 |
(* parameters: |
|
274 |
||
275 |
test: term -> bool |
|
276 |
tests if a term is at all relevant to the refutation proof; |
|
277 |
if not, then it can be discarded. Can improve performance, |
|
278 |
esp. if disjunctions can be discarded (no case distinction needed!). |
|
279 |
||
280 |
prep_tac: int -> tactic |
|
281 |
A preparation tactic to be applied to the goal once all relevant premises |
|
282 |
have been moved to the conclusion. |
|
283 |
||
284 |
ref_tac: int -> tactic |
|
285 |
the actual refutation tactic. Should be able to deal with goals |
|
286 |
[| A1; ...; An |] ==> False |
|
287 |
where the Ai are atomic, i.e. no top-level &, | or EX |
|
288 |
*) |
|
289 |
||
290 |
local |
|
291 |
val nnf_simpset = |
|
292 |
empty_ss setmkeqTrue mk_eq_True |
|
293 |
setmksimps (mksimps mksimps_pairs) |
|
22147 | 294 |
addsimps [@{thm imp_conv_disj}, @{thm iff_conv_conj_imp}, @{thm de_Morgan_disj}, |
295 |
@{thm de_Morgan_conj}, @{thm not_all}, @{thm not_ex}, @{thm not_not}]; |
|
21163 | 296 |
fun prem_nnf_tac i st = |
297 |
full_simp_tac (Simplifier.theory_context (Thm.theory_of_thm st) nnf_simpset) i st; |
|
298 |
in |
|
299 |
fun refute_tac test prep_tac ref_tac = |
|
300 |
let val refute_prems_tac = |
|
301 |
REPEAT_DETERM |
|
22147 | 302 |
(eresolve_tac [@{thm conjE}, @{thm exE}] 1 ORELSE |
21163 | 303 |
filter_prems_tac test 1 ORELSE |
22147 | 304 |
etac @{thm disjE} 1) THEN |
305 |
((etac @{thm notE} 1 THEN eq_assume_tac 1) ORELSE |
|
21163 | 306 |
ref_tac 1); |
307 |
in EVERY'[TRY o filter_prems_tac test, |
|
22147 | 308 |
REPEAT_DETERM o etac @{thm rev_mp}, prep_tac, rtac @{thm ccontr}, prem_nnf_tac, |
21163 | 309 |
SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] |
310 |
end; |
|
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
311 |
end; |
21163 | 312 |
|
313 |
val defALL_regroup = |
|
22147 | 314 |
Simplifier.simproc @{theory} |
21163 | 315 |
"defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all; |
316 |
||
317 |
val defEX_regroup = |
|
22147 | 318 |
Simplifier.simproc @{theory} |
21163 | 319 |
"defined EX" ["EX x. P x"] Quantifier1.rearrange_ex; |
320 |
||
321 |
||
21674 | 322 |
val simpset_simprocs = HOL_basic_ss |
21163 | 323 |
addsimprocs [defALL_regroup, defEX_regroup, neq_simproc, let_simproc] |
324 |
||
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
325 |
end; |
21551 | 326 |
|
327 |
structure Splitter = Simpdata.Splitter; |
|
328 |
structure Clasimp = Simpdata.Clasimp; |