1 (* Title: HOL/simpdata.ML
4 Copyright 1991 University of Cambridge
6 Instantiation of the generic simplifier for HOL.
9 (* legacy ML bindings - FIXME get rid of this *)
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";
21 val simp_impliesI = thm "simp_impliesI";
22 val simp_implies_cong = thm "simp_implies_cong";
23 val simp_implies_def = thm "simp_implies_def";
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"
33 (*** make simplification procedures for quantifier elimination ***)
35 structure Quantifier1 = Quantifier1Fun
38 fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
40 fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
42 fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
44 val conj = HOLogic.conj
47 val iff_reflection = HOL.eq_reflection
49 val iff_trans = HOL.trans
57 val iff_allI = iff_allI
59 val all_comm = all_comm
66 Simplifier.simproc (the_context ())
67 "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex;
70 Simplifier.simproc (the_context ())
71 "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all;
74 (* simproc for proving "(y = x) == False" from premise "~(x = y)" *)
76 val use_neq_simproc = ref true;
79 val neq_to_EQ_False = thm "not_sym" RS Eq_FalseI;
80 fun neq_prover sg ss (eq $ lhs $ rhs) =
82 fun test thm = (case #prop (rep_thm thm) of
83 _ $ (Not $ (eq' $ l' $ r')) =>
84 Not = HOLogic.Not andalso eq' = eq andalso
85 r' aconv lhs andalso l' aconv rhs
87 in if !use_neq_simproc then case find_first test (prems_of_ss ss)
89 | SOME thm => SOME (thm RS neq_to_EQ_False)
94 val neq_simproc = Simplifier.simproc (the_context ())
95 "neq_simproc" ["x = y"] neq_prover;
100 (* Simproc for Let *)
102 val use_let_simproc = ref true;
105 val Let_folded = thm "Let_folded";
106 val Let_unfold = thm "Let_unfold";
107 val (f_Let_unfold,x_Let_unfold) =
108 let val [(_$(f$x)$_)] = prems_of Let_unfold
109 in (cterm_of (the_context ()) f,cterm_of (the_context ()) x) end
110 val (f_Let_folded,x_Let_folded) =
111 let val [(_$(f$x)$_)] = prems_of Let_folded
112 in (cterm_of (the_context ()) f, cterm_of (the_context ()) x) end;
114 let val [(_$_$(g$_))] = prems_of Let_folded in cterm_of (the_context ()) g end;
118 Simplifier.simproc (the_context ()) "let_simp" ["Let x f"]
119 (fn sg => fn ss => fn t =>
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)
123 (case t' of (Const ("Let",_)$x$f) => (* x and f are already in normal form *)
124 if not (!use_let_simproc) then NONE
125 else if is_Free x orelse is_Bound x orelse is_Const x
126 then SOME (thm "Let_def")
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);
139 val rl = cterm_instantiate [(f_Let_unfold,cf),(x_Let_unfold,cx)] Let_unfold;
140 in SOME (rl OF [fx_g]) end
141 else if Term.betapply (f,x) aconv g then NONE (*avoid identity conversion*)
143 val abs_g'= Abs (n,xT,g');
145 val g_g'x = symmetric (beta_conversion false (cterm_of sg g'x));
146 val rl = cterm_instantiate
147 [(f_Let_folded,cterm_of sg f),(x_Let_folded,cx),
148 (g_Let_folded,cterm_of sg abs_g')]
150 in SOME (rl OF [transitive fx_g g_g'x])
158 (*** Case splitting ***)
160 (*Make meta-equalities. The operator below is Trueprop*)
162 fun mk_meta_eq r = r RS HOL.eq_reflection;
163 fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r;
165 fun mk_eq th = case concl_of th of
166 Const("==",_)$_$_ => th
167 | _$(Const("op =",_)$_$_) => mk_meta_eq th
168 | _$(Const("Not",_)$_) => th RS Eq_FalseI
169 | _ => th RS Eq_TrueI;
170 (* Expects Trueprop(.) if not == *)
173 SOME (r RS HOL.meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => NONE;
175 (* Produce theorems of the form
176 (P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y)
178 fun lift_meta_eq_to_obj_eq i st =
180 val {sign, ...} = rep_thm st;
181 fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
183 val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1)))
184 in if j = 0 then HOL.meta_eq_to_obj_eq
187 val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
188 fun mk_simp_implies Q = foldr (fn (R, S) =>
189 Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps
190 val aT = TFree ("'a", HOLogic.typeS);
191 val x = Free ("x", aT);
192 val y = Free ("y", aT)
193 in Goal.prove_global (Thm.theory_of_thm st) []
194 [mk_simp_implies (Logic.mk_equals (x, y))]
195 (mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y))))
197 [rewrite_goals_tac [simp_implies_def],
198 REPEAT (ares_tac (HOL.meta_eq_to_obj_eq :: map (rewrite_rule [simp_implies_def]) prems) 1)])
202 (*Congruence rules for = (instead of ==)*)
203 fun mk_meta_cong rl = zero_var_indexes
204 (let val rl' = Seq.hd (TRYALL (fn i => fn st =>
205 rtac (lift_meta_eq_to_obj_eq i st) i st) rl)
206 in mk_meta_eq rl' handle THM _ =>
207 if can Logic.dest_equals (concl_of rl') then rl'
208 else error "Conclusion of congruence rules must be =-equality"
211 structure SplitterData =
213 structure Simplifier = Simplifier
215 val meta_eq_to_iff = HOL.meta_eq_to_obj_eq
217 val disjE = HOL.disjE
218 val conjE = HOL.conjE
220 val contrapos = HOL.contrapos_nn
221 val contrapos2 = HOL.contrapos_pp
222 val notnotD = HOL.notnotD
225 structure Splitter = SplitterFun(SplitterData);
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;
230 val op addsplits = Splitter.addsplits;
231 val op delsplits = Splitter.delsplits;
232 val Addsplits = Splitter.Addsplits;
233 val Delsplits = Splitter.Delsplits;
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])];
241 val mk_atomize: (string * thm list) list -> thm -> thm list
242 looks too specific to move it somewhere else
244 fun mk_atomize pairs =
247 Const("Trueprop",_) $ p =>
250 (case AList.lookup (op =) pairs a of
251 SOME(rls) => List.concat (map atoms ([th] RL rls))
258 (List.mapPartial (try mk_eq) o mk_atomize pairs o gen_all);
260 fun unsafe_solver_tac prems =
261 (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN'
262 FIRST'[resolve_tac(reflexive_thm :: HOL.TrueI :: HOL.refl :: prems), atac, etac HOL.FalseE];
263 val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac;
265 (*No premature instantiation of variables during simplification*)
266 fun safe_solver_tac prems =
267 (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN'
268 FIRST'[match_tac(reflexive_thm :: HOL.TrueI :: HOL.refl :: prems),
269 eq_assume_tac, ematch_tac [HOL.FalseE]];
270 val safe_solver = mk_solver "HOL safe" safe_solver_tac;
273 Simplifier.theory_context (the_context ()) empty_ss
274 setsubgoaler asm_simp_tac
275 setSSolver safe_solver
276 setSolver unsafe_solver
277 setmksimps (mksimps mksimps_pairs)
278 setmkeqTrue mk_eq_True
279 setmkcong mk_meta_cong;
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;
285 (*In general it seems wrong to add distributive laws by default: they
286 might cause exponential blow-up. But imp_disjL has been in for a while
287 and cannot be removed without affecting existing proofs. Moreover,
288 rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the
289 grounds that it allows simplification of R in the two cases.*)
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";
307 HOL_basic_ss addsimps
308 ([triv_forall_equality, (* prunes params *)
309 True_implies_equals, (* prune asms `True' *)
310 if_True, if_False, if_cancel, if_eq_cancel,
311 imp_disjL, conj_assoc, disj_assoc,
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]
315 @ ex_simps @ all_simps @ simp_thms)
316 addsimprocs [defALL_regroup,defEX_regroup,neq_simproc,let_simproc]
317 addcongs [imp_cong, simp_implies_cong]
318 addsplits [thm "split_if"];
322 fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews);
324 (* default simpset *)
326 (fn thy => (change_simpset_of thy (fn _ => HOL_ss); thy));
329 (*** integration of simplifier with classical reasoner ***)
331 structure Clasimp = ClasimpFun
332 (structure Simplifier = Simplifier and Splitter = Splitter
333 and Classical = Classical and Blast = Blast
334 val iffD1 = HOL.iffD1 val iffD2 = HOL.iffD2 val notE = HOL.notE);
337 val HOL_css = (HOL_cs, HOL_ss);
341 (*** A general refutation procedure ***)
346 tests if a term is at all relevant to the refutation proof;
347 if not, then it can be discarded. Can improve performance,
348 esp. if disjunctions can be discarded (no case distinction needed!).
350 prep_tac: int -> tactic
351 A preparation tactic to be applied to the goal once all relevant premises
352 have been moved to the conclusion.
354 ref_tac: int -> tactic
355 the actual refutation tactic. Should be able to deal with goals
356 [| A1; ...; An |] ==> False
357 where the Ai are atomic, i.e. no top-level &, | or EX
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,
365 thm "not_all", thm "not_ex", thm "not_not"];
366 fun prem_nnf_tac i st =
367 full_simp_tac (Simplifier.theory_context (Thm.theory_of_thm st) nnf_simpset) i st;
369 fun refute_tac test prep_tac ref_tac =
370 let val refute_prems_tac =
372 (eresolve_tac [HOL.conjE, HOL.exE] 1 ORELSE
373 filter_prems_tac test 1 ORELSE
374 etac HOL.disjE 1) THEN
375 ((etac HOL.notE 1 THEN eq_assume_tac 1) ORELSE
377 in EVERY'[TRY o filter_prems_tac test,
378 REPEAT_DETERM o etac HOL.rev_mp, prep_tac, rtac HOL.ccontr, prem_nnf_tac,
379 SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)]