| author | wenzelm |
| Wed, 11 Oct 2006 00:27:38 +0200 | |
| changeset 20963 | a7fd8f05a2be |
| parent 20944 | 34b2c1bb7178 |
| child 20973 | 0b8e436ed071 |
| permissions | -rw-r--r-- |
| 1465 | 1 |
(* Title: HOL/simpdata.ML |
| 923 | 2 |
ID: $Id$ |
| 1465 | 3 |
Author: Tobias Nipkow |
| 923 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
| 5304 | 6 |
Instantiation of the generic simplifier for HOL. |
| 923 | 7 |
*) |
8 |
||
| 20944 | 9 |
(* legacy ML bindings - FIXME get rid of this *) |
| 3904 | 10 |
|
| 12281 | 11 |
val Eq_FalseI = thm "Eq_FalseI"; |
12 |
val Eq_TrueI = thm "Eq_TrueI"; |
|
13 |
val de_Morgan_conj = thm "de_Morgan_conj"; |
|
14 |
val de_Morgan_disj = thm "de_Morgan_disj"; |
|
15 |
val iff_conv_conj_imp = thm "iff_conv_conj_imp"; |
|
16 |
val imp_cong = thm "imp_cong"; |
|
17 |
val imp_conv_disj = thm "imp_conv_disj"; |
|
18 |
val imp_disj1 = thm "imp_disj1"; |
|
19 |
val imp_disj2 = thm "imp_disj2"; |
|
20 |
val imp_disjL = thm "imp_disjL"; |
|
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
21 |
val simp_impliesI = thm "simp_impliesI"; |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
22 |
val simp_implies_cong = thm "simp_implies_cong"; |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
23 |
val simp_implies_def = thm "simp_implies_def"; |
| 2134 | 24 |
|
| 11232 | 25 |
local |
| 20944 | 26 |
val uncurry = thm "uncurry" |
27 |
val iff_allI = thm "iff_allI" |
|
28 |
val iff_exI = thm "iff_exI" |
|
29 |
val all_comm = thm "all_comm" |
|
30 |
val ex_comm = thm "ex_comm" |
|
| 11232 | 31 |
in |
| 4351 | 32 |
|
33 |
(*** make simplification procedures for quantifier elimination ***) |
|
34 |
||
| 9851 | 35 |
structure Quantifier1 = Quantifier1Fun |
36 |
(struct |
|
| 4351 | 37 |
(*abstract syntax*) |
| 15531 | 38 |
fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
|
39 |
| dest_eq _ = NONE; |
|
40 |
fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
|
|
41 |
| dest_conj _ = NONE; |
|
42 |
fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
|
|
43 |
| dest_imp _ = NONE; |
|
| 4351 | 44 |
val conj = HOLogic.conj |
45 |
val imp = HOLogic.imp |
|
46 |
(*rules*) |
|
| 20944 | 47 |
val iff_reflection = HOL.eq_reflection |
48 |
val iffI = HOL.iffI |
|
49 |
val iff_trans = HOL.trans |
|
50 |
val conjI= HOL.conjI |
|
51 |
val conjE= HOL.conjE |
|
52 |
val impI = HOL.impI |
|
53 |
val mp = HOL.mp |
|
| 11232 | 54 |
val uncurry = uncurry |
| 20944 | 55 |
val exI = HOL.exI |
56 |
val exE = HOL.exE |
|
| 11232 | 57 |
val iff_allI = iff_allI |
|
12524
66eb843b1d35
mods due to mor powerful simprocs for 1-point rules (quantifier1).
nipkow
parents:
12475
diff
changeset
|
58 |
val iff_exI = iff_exI |
|
66eb843b1d35
mods due to mor powerful simprocs for 1-point rules (quantifier1).
nipkow
parents:
12475
diff
changeset
|
59 |
val all_comm = all_comm |
|
66eb843b1d35
mods due to mor powerful simprocs for 1-point rules (quantifier1).
nipkow
parents:
12475
diff
changeset
|
60 |
val ex_comm = ex_comm |
| 4351 | 61 |
end); |
62 |
||
| 11232 | 63 |
end; |
64 |
||
| 13462 | 65 |
val defEX_regroup = |
| 20944 | 66 |
Simplifier.simproc (the_context ()) |
| 13462 | 67 |
"defined EX" ["EX x. P x"] Quantifier1.rearrange_ex; |
68 |
||
69 |
val defALL_regroup = |
|
| 20944 | 70 |
Simplifier.simproc (the_context ()) |
| 13462 | 71 |
"defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all; |
| 3913 | 72 |
|
| 17778 | 73 |
|
| 20944 | 74 |
(* simproc for proving "(y = x) == False" from premise "~(x = y)" *) |
| 17778 | 75 |
|
76 |
val use_neq_simproc = ref true; |
|
77 |
||
78 |
local |
|
| 20944 | 79 |
val neq_to_EQ_False = thm "not_sym" RS Eq_FalseI; |
80 |
fun neq_prover sg ss (eq $ lhs $ rhs) = |
|
81 |
let |
|
82 |
fun test thm = (case #prop (rep_thm thm) of |
|
| 17778 | 83 |
_ $ (Not $ (eq' $ l' $ r')) => |
84 |
Not = HOLogic.Not andalso eq' = eq andalso |
|
85 |
r' aconv lhs andalso l' aconv rhs |
|
86 |
| _ => false) |
|
| 20944 | 87 |
in if !use_neq_simproc then case find_first test (prems_of_ss ss) |
88 |
of NONE => NONE |
|
89 |
| SOME thm => SOME (thm RS neq_to_EQ_False) |
|
90 |
else NONE |
|
91 |
end |
|
| 17778 | 92 |
in |
93 |
||
| 20944 | 94 |
val neq_simproc = Simplifier.simproc (the_context ()) |
95 |
"neq_simproc" ["x = y"] neq_prover; |
|
| 17778 | 96 |
|
97 |
end; |
|
98 |
||
99 |
||
| 20944 | 100 |
(* Simproc for Let *) |
| 15423 | 101 |
|
102 |
val use_let_simproc = ref true; |
|
103 |
||
104 |
local |
|
| 20944 | 105 |
val Let_folded = thm "Let_folded"; |
106 |
val Let_unfold = thm "Let_unfold"; |
|
107 |
val (f_Let_unfold,x_Let_unfold) = |
|
| 20070 | 108 |
let val [(_$(f$x)$_)] = prems_of Let_unfold |
| 20944 | 109 |
in (cterm_of (the_context ()) f,cterm_of (the_context ()) x) end |
110 |
val (f_Let_folded,x_Let_folded) = |
|
| 20070 | 111 |
let val [(_$(f$x)$_)] = prems_of Let_folded |
| 20944 | 112 |
in (cterm_of (the_context ()) f, cterm_of (the_context ()) x) end; |
113 |
val g_Let_folded = |
|
114 |
let val [(_$_$(g$_))] = prems_of Let_folded in cterm_of (the_context ()) g end; |
|
| 15423 | 115 |
in |
| 20944 | 116 |
|
| 15423 | 117 |
val let_simproc = |
| 20944 | 118 |
Simplifier.simproc (the_context ()) "let_simp" ["Let x f"] |
| 15423 | 119 |
(fn sg => fn ss => fn t => |
| 20014 | 120 |
let val ctxt = Simplifier.the_context ss; |
121 |
val ([t'],ctxt') = Variable.import_terms false [t] ctxt; |
|
122 |
in Option.map (hd o Variable.export ctxt' ctxt o single) |
|
| 20070 | 123 |
(case t' of (Const ("Let",_)$x$f) => (* x and f are already in normal form *)
|
| 15531 | 124 |
if not (!use_let_simproc) then NONE |
| 20070 | 125 |
else if is_Free x orelse is_Bound x orelse is_Const x |
| 20944 | 126 |
then SOME (thm "Let_def") |
| 15423 | 127 |
else |
128 |
let |
|
129 |
val n = case f of (Abs (x,_,_)) => x | _ => "x"; |
|
130 |
val cx = cterm_of sg x; |
|
131 |
val {T=xT,...} = rep_cterm cx;
|
|
132 |
val cf = cterm_of sg f; |
|
133 |
val fx_g = Simplifier.rewrite ss (Thm.capply cf cx); |
|
134 |
val (_$_$g) = prop_of fx_g; |
|
135 |
val g' = abstract_over (x,g); |
|
| 20070 | 136 |
in (if (g aconv g') |
| 15423 | 137 |
then |
138 |
let |
|
| 20014 | 139 |
val rl = cterm_instantiate [(f_Let_unfold,cf),(x_Let_unfold,cx)] Let_unfold; |
| 20070 | 140 |
in SOME (rl OF [fx_g]) end |
| 18176 | 141 |
else if Term.betapply (f,x) aconv g then NONE (*avoid identity conversion*) |
| 20070 | 142 |
else let |
| 15423 | 143 |
val abs_g'= Abs (n,xT,g'); |
144 |
val g'x = abs_g'$x; |
|
145 |
val g_g'x = symmetric (beta_conversion false (cterm_of sg g'x)); |
|
146 |
val rl = cterm_instantiate |
|
| 20014 | 147 |
[(f_Let_folded,cterm_of sg f),(x_Let_folded,cx), |
| 15423 | 148 |
(g_Let_folded,cterm_of sg abs_g')] |
| 20070 | 149 |
Let_folded; |
150 |
in SOME (rl OF [transitive fx_g g_g'x]) |
|
| 20014 | 151 |
end) |
| 15423 | 152 |
end |
| 20014 | 153 |
| _ => NONE) |
| 20070 | 154 |
end) |
| 20944 | 155 |
|
| 15423 | 156 |
end |
| 4351 | 157 |
|
158 |
(*** Case splitting ***) |
|
| 3913 | 159 |
|
| 12278 | 160 |
(*Make meta-equalities. The operator below is Trueprop*) |
161 |
||
| 20944 | 162 |
fun mk_meta_eq r = r RS HOL.eq_reflection; |
| 12278 | 163 |
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; |
164 |
||
165 |
fun mk_eq th = case concl_of th of |
|
166 |
Const("==",_)$_$_ => th
|
|
167 |
| _$(Const("op =",_)$_$_) => mk_meta_eq th
|
|
|
13600
9702c8636a7b
Removed nRS again because extract_rews now takes care of preserving names.
berghofe
parents:
13568
diff
changeset
|
168 |
| _$(Const("Not",_)$_) => th RS Eq_FalseI
|
|
9702c8636a7b
Removed nRS again because extract_rews now takes care of preserving names.
berghofe
parents:
13568
diff
changeset
|
169 |
| _ => th RS Eq_TrueI; |
|
13568
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
170 |
(* Expects Trueprop(.) if not == *) |
| 12278 | 171 |
|
172 |
fun mk_eq_True r = |
|
| 20944 | 173 |
SOME (r RS HOL.meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => NONE; |
| 12278 | 174 |
|
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
175 |
(* Produce theorems of the form |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
176 |
(P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y) |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
177 |
*) |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
178 |
fun lift_meta_eq_to_obj_eq i st = |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
179 |
let |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
180 |
val {sign, ...} = rep_thm st;
|
| 17197 | 181 |
fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
|
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
182 |
| count_imp _ = 0; |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
183 |
val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) |
| 20944 | 184 |
in if j = 0 then HOL.meta_eq_to_obj_eq |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
185 |
else |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
186 |
let |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
187 |
val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
|
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
188 |
fun mk_simp_implies Q = foldr (fn (R, S) => |
| 17197 | 189 |
Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps
|
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
190 |
val aT = TFree ("'a", HOLogic.typeS);
|
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
191 |
val x = Free ("x", aT);
|
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
192 |
val y = Free ("y", aT)
|
| 20046 | 193 |
in Goal.prove_global (Thm.theory_of_thm st) [] |
| 17985 | 194 |
[mk_simp_implies (Logic.mk_equals (x, y))] |
195 |
(mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y)))) |
|
196 |
(fn prems => EVERY |
|
197 |
[rewrite_goals_tac [simp_implies_def], |
|
| 20944 | 198 |
REPEAT (ares_tac (HOL.meta_eq_to_obj_eq :: map (rewrite_rule [simp_implies_def]) prems) 1)]) |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
199 |
end |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
200 |
end; |
| 17985 | 201 |
|
| 12278 | 202 |
(*Congruence rules for = (instead of ==)*) |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
203 |
fun mk_meta_cong rl = zero_var_indexes |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
204 |
(let val rl' = Seq.hd (TRYALL (fn i => fn st => |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
205 |
rtac (lift_meta_eq_to_obj_eq i st) i st) rl) |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
206 |
in mk_meta_eq rl' handle THM _ => |
| 20872 | 207 |
if can Logic.dest_equals (concl_of rl') then rl' |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
208 |
else error "Conclusion of congruence rules must be =-equality" |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
209 |
end); |
| 12278 | 210 |
|
| 5304 | 211 |
structure SplitterData = |
| 20944 | 212 |
struct |
| 5304 | 213 |
structure Simplifier = Simplifier |
| 5552 | 214 |
val mk_eq = mk_eq |
| 20944 | 215 |
val meta_eq_to_iff = HOL.meta_eq_to_obj_eq |
216 |
val iffD = HOL.iffD2 |
|
217 |
val disjE = HOL.disjE |
|
218 |
val conjE = HOL.conjE |
|
219 |
val exE = HOL.exE |
|
220 |
val contrapos = HOL.contrapos_nn |
|
221 |
val contrapos2 = HOL.contrapos_pp |
|
222 |
val notnotD = HOL.notnotD |
|
223 |
end; |
|
| 4681 | 224 |
|
| 5304 | 225 |
structure Splitter = SplitterFun(SplitterData); |
| 2263 | 226 |
|
| 5304 | 227 |
val split_tac = Splitter.split_tac; |
228 |
val split_inside_tac = Splitter.split_inside_tac; |
|
229 |
val split_asm_tac = Splitter.split_asm_tac; |
|
| 5307 | 230 |
val op addsplits = Splitter.addsplits; |
231 |
val op delsplits = Splitter.delsplits; |
|
| 5304 | 232 |
val Addsplits = Splitter.Addsplits; |
233 |
val Delsplits = Splitter.Delsplits; |
|
|
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
234 |
|
| 2134 | 235 |
val mksimps_pairs = |
| 20944 | 236 |
[("op -->", [HOL.mp]), ("op &", [thm "conjunct1", thm "conjunct2"]),
|
237 |
("All", [HOL.spec]), ("True", []), ("False", []),
|
|
238 |
("HOL.If", [thm "if_bool_eq_conj" RS HOL.iffD1])];
|
|
| 1758 | 239 |
|
|
13568
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
240 |
(* |
| 5304 | 241 |
val mk_atomize: (string * thm list) list -> thm -> thm list |
|
13568
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
242 |
looks too specific to move it somewhere else |
| 5304 | 243 |
*) |
244 |
fun mk_atomize pairs = |
|
245 |
let fun atoms th = |
|
246 |
(case concl_of th of |
|
247 |
Const("Trueprop",_) $ p =>
|
|
248 |
(case head_of p of |
|
249 |
Const(a,_) => |
|
| 17325 | 250 |
(case AList.lookup (op =) pairs a of |
| 15570 | 251 |
SOME(rls) => List.concat (map atoms ([th] RL rls)) |
| 15531 | 252 |
| NONE => [th]) |
| 5304 | 253 |
| _ => [th]) |
254 |
| _ => [th]) |
|
255 |
in atoms end; |
|
256 |
||
|
11624
8a45c7abef04
mksimps and mk_eq_True no longer raise THM exception.
berghofe
parents:
11534
diff
changeset
|
257 |
fun mksimps pairs = |
| 15570 | 258 |
(List.mapPartial (try mk_eq) o mk_atomize pairs o gen_all); |
| 5304 | 259 |
|
| 7570 | 260 |
fun unsafe_solver_tac prems = |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
261 |
(fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN' |
| 20944 | 262 |
FIRST'[resolve_tac(reflexive_thm :: HOL.TrueI :: HOL.refl :: prems), atac, etac HOL.FalseE]; |
| 7570 | 263 |
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; |
264 |
||
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
265 |
(*No premature instantiation of variables during simplification*) |
| 7570 | 266 |
fun safe_solver_tac prems = |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
267 |
(fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN' |
| 20944 | 268 |
FIRST'[match_tac(reflexive_thm :: HOL.TrueI :: HOL.refl :: prems), |
269 |
eq_assume_tac, ematch_tac [HOL.FalseE]]; |
|
| 7570 | 270 |
val safe_solver = mk_solver "HOL safe" safe_solver_tac; |
|
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
271 |
|
| 9713 | 272 |
val HOL_basic_ss = |
| 17892 | 273 |
Simplifier.theory_context (the_context ()) empty_ss |
274 |
setsubgoaler asm_simp_tac |
|
| 9713 | 275 |
setSSolver safe_solver |
276 |
setSolver unsafe_solver |
|
277 |
setmksimps (mksimps mksimps_pairs) |
|
278 |
setmkeqTrue mk_eq_True |
|
279 |
setmkcong mk_meta_cong; |
|
|
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
280 |
|
| 18324 | 281 |
fun unfold_tac ths = |
282 |
let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths |
|
283 |
in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; |
|
| 17003 | 284 |
|
|
13568
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
285 |
(*In general it seems wrong to add distributive laws by default: they |
|
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
286 |
might cause exponential blow-up. But imp_disjL has been in for a while |
|
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
287 |
and cannot be removed without affecting existing proofs. Moreover, |
|
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
288 |
rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the |
|
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
289 |
grounds that it allows simplification of R in the two cases.*) |
|
6b12df05f358
preserve names of rewrite rules when transforming them
nipkow
parents:
13462
diff
changeset
|
290 |
|
| 20944 | 291 |
local |
292 |
val ex_simps = thms "ex_simps"; |
|
293 |
val all_simps = thms "all_simps"; |
|
294 |
val simp_thms = thms "simp_thms"; |
|
295 |
val cases_simp = thm "cases_simp"; |
|
296 |
val conj_assoc = thm "conj_assoc"; |
|
297 |
val if_False = thm "if_False"; |
|
298 |
val if_True = thm "if_True"; |
|
299 |
val disj_assoc = thm "disj_assoc"; |
|
300 |
val disj_not1 = thm "disj_not1"; |
|
301 |
val if_cancel = thm "if_cancel"; |
|
302 |
val if_eq_cancel = thm "if_eq_cancel"; |
|
303 |
val True_implies_equals = thm "True_implies_equals"; |
|
304 |
in |
|
305 |
||
| 9713 | 306 |
val HOL_ss = |
307 |
HOL_basic_ss addsimps |
|
|
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
308 |
([triv_forall_equality, (* prunes params *) |
| 3654 | 309 |
True_implies_equals, (* prune asms `True' *) |
|
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
310 |
if_True, if_False, if_cancel, if_eq_cancel, |
| 5304 | 311 |
imp_disjL, conj_assoc, disj_assoc, |
| 20944 | 312 |
de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, thm "not_imp", |
313 |
disj_not1, thm "not_all", thm "not_ex", cases_simp, |
|
314 |
thm "the_eq_trivial", HOL.the_sym_eq_trivial] |
|
|
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
315 |
@ ex_simps @ all_simps @ simp_thms) |
| 17778 | 316 |
addsimprocs [defALL_regroup,defEX_regroup,neq_simproc,let_simproc] |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
317 |
addcongs [imp_cong, simp_implies_cong] |
| 20944 | 318 |
addsplits [thm "split_if"]; |
319 |
||
320 |
end; |
|
| 2082 | 321 |
|
| 11034 | 322 |
fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); |
323 |
||
| 7357 | 324 |
(* default simpset *) |
| 9713 | 325 |
val simpsetup = |
| 20944 | 326 |
(fn thy => (change_simpset_of thy (fn _ => HOL_ss); thy)); |
|
3615
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
327 |
|
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4651
diff
changeset
|
328 |
|
| 5219 | 329 |
(*** integration of simplifier with classical reasoner ***) |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
330 |
|
| 5219 | 331 |
structure Clasimp = ClasimpFun |
| 8473 | 332 |
(structure Simplifier = Simplifier and Splitter = Splitter |
| 9851 | 333 |
and Classical = Classical and Blast = Blast |
| 20944 | 334 |
val iffD1 = HOL.iffD1 val iffD2 = HOL.iffD2 val notE = HOL.notE); |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4651
diff
changeset
|
335 |
open Clasimp; |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
336 |
|
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
337 |
val HOL_css = (HOL_cs, HOL_ss); |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
338 |
|
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
339 |
|
|
8641
978db2870862
change_global/local_css move to Provers/clasimp.ML;
wenzelm
parents:
8473
diff
changeset
|
340 |
|
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
341 |
(*** A general refutation procedure ***) |
| 9713 | 342 |
|
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
343 |
(* Parameters: |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
344 |
|
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
345 |
test: term -> bool |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
346 |
tests if a term is at all relevant to the refutation proof; |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
347 |
if not, then it can be discarded. Can improve performance, |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
348 |
esp. if disjunctions can be discarded (no case distinction needed!). |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
349 |
|
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
350 |
prep_tac: int -> tactic |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
351 |
A preparation tactic to be applied to the goal once all relevant premises |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
352 |
have been moved to the conclusion. |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
353 |
|
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
354 |
ref_tac: int -> tactic |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
355 |
the actual refutation tactic. Should be able to deal with goals |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
356 |
[| A1; ...; An |] ==> False |
| 9876 | 357 |
where the Ai are atomic, i.e. no top-level &, | or EX |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
358 |
*) |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
359 |
|
| 15184 | 360 |
local |
361 |
val nnf_simpset = |
|
| 17892 | 362 |
empty_ss setmkeqTrue mk_eq_True |
363 |
setmksimps (mksimps mksimps_pairs) |
|
364 |
addsimps [imp_conv_disj,iff_conv_conj_imp,de_Morgan_disj,de_Morgan_conj, |
|
| 20944 | 365 |
thm "not_all", thm "not_ex", thm "not_not"]; |
| 17892 | 366 |
fun prem_nnf_tac i st = |
367 |
full_simp_tac (Simplifier.theory_context (Thm.theory_of_thm st) nnf_simpset) i st; |
|
| 15184 | 368 |
in |
369 |
fun refute_tac test prep_tac ref_tac = |
|
370 |
let val refute_prems_tac = |
|
| 12475 | 371 |
REPEAT_DETERM |
| 20944 | 372 |
(eresolve_tac [HOL.conjE, HOL.exE] 1 ORELSE |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
373 |
filter_prems_tac test 1 ORELSE |
| 20944 | 374 |
etac HOL.disjE 1) THEN |
375 |
((etac HOL.notE 1 THEN eq_assume_tac 1) ORELSE |
|
|
11200
f43fa07536c0
arith_tac now copes with propositional reasoning as well.
nipkow
parents:
11162
diff
changeset
|
376 |
ref_tac 1); |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
377 |
in EVERY'[TRY o filter_prems_tac test, |
| 20944 | 378 |
REPEAT_DETERM o etac HOL.rev_mp, prep_tac, rtac HOL.ccontr, prem_nnf_tac, |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
379 |
SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
380 |
end; |
| 17003 | 381 |
end; |