1 (* Title: HOL/simpdata.ML
4 Copyright 1991 University of Cambridge
6 Instantiation of the generic simplifier for HOL.
9 (* legacy ML bindings *)
11 val Eq_FalseI = thm "Eq_FalseI";
12 val Eq_TrueI = thm "Eq_TrueI";
13 val all_conj_distrib = thm "all_conj_distrib";
14 val all_simps = thms "all_simps";
15 val cases_simp = thm "cases_simp";
16 val conj_assoc = thm "conj_assoc";
17 val conj_comms = thms "conj_comms";
18 val conj_commute = thm "conj_commute";
19 val conj_cong = thm "conj_cong";
20 val conj_disj_distribL = thm "conj_disj_distribL";
21 val conj_disj_distribR = thm "conj_disj_distribR";
22 val conj_left_commute = thm "conj_left_commute";
23 val de_Morgan_conj = thm "de_Morgan_conj";
24 val de_Morgan_disj = thm "de_Morgan_disj";
25 val disj_assoc = thm "disj_assoc";
26 val disj_comms = thms "disj_comms";
27 val disj_commute = thm "disj_commute";
28 val disj_cong = thm "disj_cong";
29 val disj_conj_distribL = thm "disj_conj_distribL";
30 val disj_conj_distribR = thm "disj_conj_distribR";
31 val disj_left_commute = thm "disj_left_commute";
32 val disj_not1 = thm "disj_not1";
33 val disj_not2 = thm "disj_not2";
34 val eq_ac = thms "eq_ac";
35 val eq_assoc = thm "eq_assoc";
36 val eq_commute = thm "eq_commute";
37 val eq_left_commute = thm "eq_left_commute";
38 val eq_sym_conv = thm "eq_sym_conv";
39 val eta_contract_eq = thm "eta_contract_eq";
40 val ex_disj_distrib = thm "ex_disj_distrib";
41 val ex_simps = thms "ex_simps";
42 val if_False = thm "if_False";
43 val if_P = thm "if_P";
44 val if_True = thm "if_True";
45 val if_bool_eq_conj = thm "if_bool_eq_conj";
46 val if_bool_eq_disj = thm "if_bool_eq_disj";
47 val if_cancel = thm "if_cancel";
48 val if_def2 = thm "if_def2";
49 val if_eq_cancel = thm "if_eq_cancel";
50 val if_not_P = thm "if_not_P";
51 val if_splits = thms "if_splits";
52 val iff_conv_conj_imp = thm "iff_conv_conj_imp";
53 val imp_all = thm "imp_all";
54 val imp_cong = thm "imp_cong";
55 val imp_conjL = thm "imp_conjL";
56 val imp_conjR = thm "imp_conjR";
57 val imp_conv_disj = thm "imp_conv_disj";
58 val imp_disj1 = thm "imp_disj1";
59 val imp_disj2 = thm "imp_disj2";
60 val imp_disjL = thm "imp_disjL";
61 val imp_disj_not1 = thm "imp_disj_not1";
62 val imp_disj_not2 = thm "imp_disj_not2";
63 val imp_ex = thm "imp_ex";
64 val meta_eq_to_obj_eq = thm "meta_eq_to_obj_eq";
65 val neq_commute = thm "neq_commute";
66 val not_all = thm "not_all";
67 val not_ex = thm "not_ex";
68 val not_iff = thm "not_iff";
69 val not_imp = thm "not_imp";
70 val not_not = thm "not_not";
71 val rev_conj_cong = thm "rev_conj_cong";
72 val simp_impliesE = thm "simp_impliesI";
73 val simp_impliesI = thm "simp_impliesI";
74 val simp_implies_cong = thm "simp_implies_cong";
75 val simp_implies_def = thm "simp_implies_def";
76 val simp_thms = thms "simp_thms";
77 val split_if = thm "split_if";
78 val split_if_asm = thm "split_if_asm";
79 val atomize_not = thm"atomize_not";
82 val uncurry = prove_goal (the_context()) "P --> Q --> R ==> P & Q --> R"
83 (fn prems => [cut_facts_tac prems 1, Blast_tac 1]);
85 val iff_allI = allI RS
86 prove_goal (the_context()) "!x. P x = Q x ==> (!x. P x) = (!x. Q x)"
87 (fn prems => [cut_facts_tac prems 1, Blast_tac 1])
89 prove_goal (the_context()) "!x. P x = Q x ==> (? x. P x) = (? x. Q x)"
90 (fn prems => [cut_facts_tac prems 1, Blast_tac 1])
92 val all_comm = prove_goal (the_context()) "(!x y. P x y) = (!y x. P x y)"
93 (fn _ => [Blast_tac 1])
94 val ex_comm = prove_goal (the_context()) "(? x y. P x y) = (? y x. P x y)"
95 (fn _ => [Blast_tac 1])
98 (*** make simplification procedures for quantifier elimination ***)
100 structure Quantifier1 = Quantifier1Fun
103 fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
105 fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
106 | dest_conj _ = NONE;
107 fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
109 val conj = HOLogic.conj
110 val imp = HOLogic.imp
112 val iff_reflection = eq_reflection
114 val iff_trans = trans
119 val uncurry = uncurry
122 val iff_allI = iff_allI
123 val iff_exI = iff_exI
124 val all_comm = all_comm
125 val ex_comm = ex_comm
131 Simplifier.simproc (Theory.sign_of (the_context ()))
132 "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex;
135 Simplifier.simproc (Theory.sign_of (the_context ()))
136 "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all;
138 (*** Simproc for Let ***)
140 val use_let_simproc = ref true;
143 val Let_folded = thm "Let_folded";
144 val Let_unfold = thm "Let_unfold";
147 let val [(_$(f$_)$_)] = prems_of Let_unfold in cterm_of (sign_of (the_context ())) f end
149 let val [(_$(f$_)$_)] = prems_of Let_folded in cterm_of (sign_of (the_context ())) f end;
151 let val [(_$_$(g$_))] = prems_of Let_folded in cterm_of (sign_of (the_context ())) g end;
154 Simplifier.simproc (Theory.sign_of (the_context ())) "let_simp" ["Let x f"]
155 (fn sg => fn ss => fn t =>
156 (case t of (Const ("Let",_)$x$f) => (* x and f are already in normal form *)
157 if not (!use_let_simproc) then NONE
158 else if is_Free x orelse is_Bound x orelse is_Const x
162 val n = case f of (Abs (x,_,_)) => x | _ => "x";
163 val cx = cterm_of sg x;
164 val {T=xT,...} = rep_cterm cx;
165 val cf = cterm_of sg f;
166 val fx_g = Simplifier.rewrite ss (Thm.capply cf cx);
167 val (_$_$g) = prop_of fx_g;
168 val g' = abstract_over (x,g);
172 val rl = cterm_instantiate [(f_Let_unfold,cf)] Let_unfold;
173 in SOME (rl OF [fx_g]) end
174 else if betapply (f,x) aconv g then NONE (* avoid identity conversion *)
176 val abs_g'= Abs (n,xT,g');
178 val g_g'x = symmetric (beta_conversion false (cterm_of sg g'x));
179 val rl = cterm_instantiate
180 [(f_Let_folded,cterm_of sg f),
181 (g_Let_folded,cterm_of sg abs_g')]
183 in SOME (rl OF [transitive fx_g g_g'x]) end)
188 (*** Case splitting ***)
190 (*Make meta-equalities. The operator below is Trueprop*)
192 fun mk_meta_eq r = r RS eq_reflection;
193 fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r;
195 fun mk_eq th = case concl_of th of
196 Const("==",_)$_$_ => th
197 | _$(Const("op =",_)$_$_) => mk_meta_eq th
198 | _$(Const("Not",_)$_) => th RS Eq_FalseI
199 | _ => th RS Eq_TrueI;
200 (* Expects Trueprop(.) if not == *)
203 SOME (r RS meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => NONE;
205 (* Produce theorems of the form
206 (P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y)
208 fun lift_meta_eq_to_obj_eq i st =
210 val {sign, ...} = rep_thm st;
211 fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
213 val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1)))
214 in if j = 0 then meta_eq_to_obj_eq
217 val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
218 fun mk_simp_implies Q = foldr (fn (R, S) =>
219 Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps
220 val aT = TFree ("'a", HOLogic.typeS);
221 val x = Free ("x", aT);
222 val y = Free ("y", aT)
223 in prove_goalw_cterm [simp_implies_def]
224 (cterm_of sign (Logic.mk_implies
225 (mk_simp_implies (Logic.mk_equals (x, y)),
226 mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y))))))
227 (fn hyps => [REPEAT (ares_tac (meta_eq_to_obj_eq :: hyps) 1)])
231 (*Congruence rules for = (instead of ==)*)
232 fun mk_meta_cong rl = zero_var_indexes
233 (let val rl' = Seq.hd (TRYALL (fn i => fn st =>
234 rtac (lift_meta_eq_to_obj_eq i st) i st) rl)
235 in mk_meta_eq rl' handle THM _ =>
236 if Logic.is_equals (concl_of rl') then rl'
237 else error "Conclusion of congruence rules must be =-equality"
240 (* Elimination of True from asumptions: *)
242 local fun rd s = read_cterm (sign_of (the_context())) (s, propT);
243 in val True_implies_equals = standard' (equal_intr
244 (implies_intr_hyps (implies_elim (assume (rd "True ==> PROP P")) TrueI))
245 (implies_intr_hyps (implies_intr (rd "True") (assume (rd "PROP P")))));
249 structure SplitterData =
251 structure Simplifier = Simplifier
253 val meta_eq_to_iff = meta_eq_to_obj_eq
258 val contrapos = contrapos_nn
259 val contrapos2 = contrapos_pp
260 val notnotD = notnotD
263 structure Splitter = SplitterFun(SplitterData);
265 val split_tac = Splitter.split_tac;
266 val split_inside_tac = Splitter.split_inside_tac;
267 val split_asm_tac = Splitter.split_asm_tac;
268 val op addsplits = Splitter.addsplits;
269 val op delsplits = Splitter.delsplits;
270 val Addsplits = Splitter.Addsplits;
271 val Delsplits = Splitter.Delsplits;
274 [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
275 ("All", [spec]), ("True", []), ("False", []),
276 ("HOL.If", [if_bool_eq_conj RS iffD1])];
279 val mk_atomize: (string * thm list) list -> thm -> thm list
280 looks too specific to move it somewhere else
282 fun mk_atomize pairs =
285 Const("Trueprop",_) $ p =>
288 (case AList.lookup (op =) pairs a of
289 SOME(rls) => List.concat (map atoms ([th] RL rls))
296 (List.mapPartial (try mk_eq) o mk_atomize pairs o gen_all);
298 fun unsafe_solver_tac prems =
299 (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN'
300 FIRST'[resolve_tac(reflexive_thm::TrueI::refl::prems), atac, etac FalseE];
301 val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac;
303 (*No premature instantiation of variables during simplification*)
304 fun safe_solver_tac prems =
305 (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN'
306 FIRST'[match_tac(reflexive_thm::TrueI::refl::prems),
307 eq_assume_tac, ematch_tac [FalseE]];
308 val safe_solver = mk_solver "HOL safe" safe_solver_tac;
311 empty_ss setsubgoaler asm_simp_tac
312 setSSolver safe_solver
313 setSolver unsafe_solver
314 setmksimps (mksimps mksimps_pairs)
315 setmkeqTrue mk_eq_True
316 setmkcong mk_meta_cong;
318 fun unfold_tac ss ths =
319 ALLGOALS (full_simp_tac
320 (Simplifier.inherit_bounds ss (Simplifier.clear_ss HOL_basic_ss) addsimps ths));
322 (*In general it seems wrong to add distributive laws by default: they
323 might cause exponential blow-up. But imp_disjL has been in for a while
324 and cannot be removed without affecting existing proofs. Moreover,
325 rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the
326 grounds that it allows simplification of R in the two cases.*)
329 HOL_basic_ss addsimps
330 ([triv_forall_equality, (* prunes params *)
331 True_implies_equals, (* prune asms `True' *)
332 eta_contract_eq, (* prunes eta-expansions *)
333 if_True, if_False, if_cancel, if_eq_cancel,
334 imp_disjL, conj_assoc, disj_assoc,
335 de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp,
336 disj_not1, not_all, not_ex, cases_simp,
337 thm "the_eq_trivial", the_sym_eq_trivial]
338 @ ex_simps @ all_simps @ simp_thms)
339 addsimprocs [defALL_regroup,defEX_regroup,let_simproc]
340 addcongs [imp_cong, simp_implies_cong]
341 addsplits [split_if];
343 fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews);
346 (*Simplifies x assuming c and y assuming ~c*)
347 val prems = Goalw [if_def]
348 "[| b=c; c ==> x=u; ~c ==> y=v |] ==> \
349 \ (if b then x else y) = (if c then u else v)";
350 by (asm_simp_tac (HOL_ss addsimps prems) 1);
353 (*Prevents simplification of x and y: faster and allows the execution
354 of functional programs. NOW THE DEFAULT.*)
355 Goal "b=c ==> (if b then x else y) = (if c then x else y)";
356 by (etac arg_cong 1);
359 (*Prevents simplification of t: much faster*)
360 Goal "a = b ==> (let x=a in t(x)) = (let x=b in t(x))";
361 by (etac arg_cong 1);
364 (*To tidy up the result of a simproc. Only the RHS will be simplified.*)
365 Goal "u = u' ==> (t==u) == (t==u')";
366 by (asm_simp_tac HOL_ss 1);
369 Goal "f(if c then x else y) = (if c then f x else f y)";
370 by (simp_tac (HOL_ss setloop (split_tac [split_if])) 1);
373 (*For expand_case_tac*)
374 val prems = Goal "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)";
376 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems)));
379 (*Used in Auth proofs. Typically P contains Vars that become instantiated
380 during unification.*)
381 fun expand_case_tac P i =
382 res_inst_tac [("P",P)] expand_case i THEN
386 (*This lemma restricts the effect of the rewrite rule u=v to the left-hand
387 side of an equality. Used in {Integ,Real}/simproc.ML*)
388 Goal "x=y ==> (x=z) = (y=z)";
389 by (asm_simp_tac HOL_ss 1);
390 qed "restrict_to_left";
392 (* default simpset *)
394 [fn thy => (simpset_ref_of thy := HOL_ss addcongs [if_weak_cong]; thy)];
397 (*** integration of simplifier with classical reasoner ***)
399 structure Clasimp = ClasimpFun
400 (structure Simplifier = Simplifier and Splitter = Splitter
401 and Classical = Classical and Blast = Blast
402 val iffD1 = iffD1 val iffD2 = iffD2 val notE = notE
403 val cla_make_elim = cla_make_elim);
406 val HOL_css = (HOL_cs, HOL_ss);
410 (*** A general refutation procedure ***)
415 tests if a term is at all relevant to the refutation proof;
416 if not, then it can be discarded. Can improve performance,
417 esp. if disjunctions can be discarded (no case distinction needed!).
419 prep_tac: int -> tactic
420 A preparation tactic to be applied to the goal once all relevant premises
421 have been moved to the conclusion.
423 ref_tac: int -> tactic
424 the actual refutation tactic. Should be able to deal with goals
425 [| A1; ...; An |] ==> False
426 where the Ai are atomic, i.e. no top-level &, | or EX
431 [imp_conv_disj,iff_conv_conj_imp,de_Morgan_disj,de_Morgan_conj,
432 not_all,not_ex,not_not];
434 empty_ss setmkeqTrue mk_eq_True
435 setmksimps (mksimps mksimps_pairs)
437 val prem_nnf_tac = full_simp_tac nnf_simpset
439 fun refute_tac test prep_tac ref_tac =
440 let val refute_prems_tac =
442 (eresolve_tac [conjE, exE] 1 ORELSE
443 filter_prems_tac test 1 ORELSE
445 ((etac notE 1 THEN eq_assume_tac 1) ORELSE
447 in EVERY'[TRY o filter_prems_tac test,
448 REPEAT_DETERM o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac,
449 SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)]