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