1 (* Title: FOL/simpdata.ML
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
4 Copyright 1994 University of Cambridge
6 Simplification data for FOL.
10 (* Elimination of True from asumptions: *)
12 bind_thm ("True_implies_equals", prove_goal IFOL.thy
13 "(True ==> PROP P) == PROP P"
14 (K [rtac equal_intr_rule 1, atac 2,
15 METAHYPS (fn prems => resolve_tac prems 1) 1,
19 (*** Rewrite rules ***)
24 (fn prems => [ (cut_facts_tac prems 1),
25 (IntPr.fast_tac 1) ]));
27 bind_thms ("conj_simps", map int_prove_fun
28 ["P & True <-> P", "True & P <-> P",
29 "P & False <-> False", "False & P <-> False",
30 "P & P <-> P", "P & P & Q <-> P & Q",
31 "P & ~P <-> False", "~P & P <-> False",
32 "(P & Q) & R <-> P & (Q & R)"]);
34 bind_thms ("disj_simps", map int_prove_fun
35 ["P | True <-> True", "True | P <-> True",
36 "P | False <-> P", "False | P <-> P",
37 "P | P <-> P", "P | P | Q <-> P | Q",
38 "(P | Q) | R <-> P | (Q | R)"]);
40 bind_thms ("not_simps", map int_prove_fun
41 ["~(P|Q) <-> ~P & ~Q",
42 "~ False <-> True", "~ True <-> False"]);
44 bind_thms ("imp_simps", map int_prove_fun
45 ["(P --> False) <-> ~P", "(P --> True) <-> True",
46 "(False --> P) <-> True", "(True --> P) <-> P",
47 "(P --> P) <-> True", "(P --> ~P) <-> ~P"]);
49 bind_thms ("iff_simps", map int_prove_fun
50 ["(True <-> P) <-> P", "(P <-> True) <-> P",
52 "(False <-> P) <-> ~P", "(P <-> False) <-> ~P"]);
54 (*The x=t versions are needed for the simplification procedures*)
55 bind_thms ("quant_simps", map int_prove_fun
57 "(ALL x. x=t --> P(x)) <-> P(t)",
58 "(ALL x. t=x --> P(x)) <-> P(t)",
60 "EX x. x=t", "EX x. t=x",
61 "(EX x. x=t & P(x)) <-> P(t)",
62 "(EX x. t=x & P(x)) <-> P(t)"]);
64 (*These are NOT supplied by default!*)
65 bind_thms ("distrib_simps", map int_prove_fun
66 ["P & (Q | R) <-> P&Q | P&R",
67 "(Q | R) & P <-> Q&P | R&P",
68 "(P | Q --> R) <-> (P --> R) & (Q --> R)"]);
70 (** Conversion into rewrite rules **)
72 bind_thm ("P_iff_F", int_prove_fun "~P ==> (P <-> False)");
73 bind_thm ("iff_reflection_F", P_iff_F RS iff_reflection);
75 bind_thm ("P_iff_T", int_prove_fun "P ==> (P <-> True)");
76 bind_thm ("iff_reflection_T", P_iff_T RS iff_reflection);
78 (*Make meta-equalities. The operator below is Trueprop*)
80 fun mk_meta_eq th = case concl_of th of
81 _ $ (Const("op =",_)$_$_) => th RS eq_reflection
82 | _ $ (Const("op <->",_)$_$_) => th RS iff_reflection
84 error("conclusion must be a =-equality or <->");;
86 fun mk_eq th = case concl_of th of
87 Const("==",_)$_$_ => th
88 | _ $ (Const("op =",_)$_$_) => mk_meta_eq th
89 | _ $ (Const("op <->",_)$_$_) => mk_meta_eq th
90 | _ $ (Const("Not",_)$_) => th RS iff_reflection_F
91 | _ => th RS iff_reflection_T;
93 (*Replace premises x=y, X<->Y by X==Y*)
96 (REPEAT_FIRST (resolve_tac [meta_eq_to_obj_eq, def_imp_iff]));
98 (*Congruence rules for = or <-> (instead of ==)*)
100 standard(mk_meta_eq (mk_meta_prems rl))
102 error("Premises and conclusion of congruence rules must use =-equality or <->");
105 [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
106 ("All", [spec]), ("True", []), ("False", [])];
108 (* ###FIXME: move to simplifier.ML
109 val mk_atomize: (string * thm list) list -> thm -> thm list
111 (* ###FIXME: move to simplifier.ML *)
112 fun mk_atomize pairs =
115 Const("Trueprop",_) $ p =>
118 (case AList.lookup (op =) pairs a of
119 SOME(rls) => List.concat (map atoms ([th] RL rls))
125 fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all);
127 (*** Classical laws ***)
131 prove_goal (the_context ()) s
132 (fn prems => [ (cut_facts_tac prems 1),
133 (Cla.fast_tac FOL_cs 1) ]));
135 (*Avoids duplication of subgoals after expand_if, when the true and false
136 cases boil down to the same thing.*)
137 bind_thm ("cases_simp", prove_fun "(P --> Q) & (~P --> Q) <-> Q");
140 (*** Miniscoping: pushing quantifiers in
141 We do NOT distribute of ALL over &, or dually that of EX over |
142 Baaz and Leitsch, On Skolemization and Proof Complexity (1994)
143 show that this step can increase proof length!
146 (*existential miniscoping*)
147 bind_thms ("int_ex_simps", map int_prove_fun
148 ["(EX x. P(x) & Q) <-> (EX x. P(x)) & Q",
149 "(EX x. P & Q(x)) <-> P & (EX x. Q(x))",
150 "(EX x. P(x) | Q) <-> (EX x. P(x)) | Q",
151 "(EX x. P | Q(x)) <-> P | (EX x. Q(x))"]);
154 bind_thms ("cla_ex_simps", map prove_fun
155 ["(EX x. P(x) --> Q) <-> (ALL x. P(x)) --> Q",
156 "(EX x. P --> Q(x)) <-> P --> (EX x. Q(x))"]);
158 bind_thms ("ex_simps", int_ex_simps @ cla_ex_simps);
160 (*universal miniscoping*)
161 bind_thms ("int_all_simps", map int_prove_fun
162 ["(ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q",
163 "(ALL x. P & Q(x)) <-> P & (ALL x. Q(x))",
164 "(ALL x. P(x) --> Q) <-> (EX x. P(x)) --> Q",
165 "(ALL x. P --> Q(x)) <-> P --> (ALL x. Q(x))"]);
168 bind_thms ("cla_all_simps", map prove_fun
169 ["(ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q",
170 "(ALL x. P | Q(x)) <-> P | (ALL x. Q(x))"]);
172 bind_thms ("all_simps", int_all_simps @ cla_all_simps);
175 (*** Named rewrite rules proved for IFOL ***)
177 fun int_prove nm thm = qed_goal nm IFOL.thy thm
178 (fn prems => [ (cut_facts_tac prems 1),
179 (IntPr.fast_tac 1) ]);
181 fun prove nm thm = qed_goal nm (the_context ()) thm (fn _ => [Blast_tac 1]);
183 int_prove "conj_commute" "P&Q <-> Q&P";
184 int_prove "conj_left_commute" "P&(Q&R) <-> Q&(P&R)";
185 bind_thms ("conj_comms", [conj_commute, conj_left_commute]);
187 int_prove "disj_commute" "P|Q <-> Q|P";
188 int_prove "disj_left_commute" "P|(Q|R) <-> Q|(P|R)";
189 bind_thms ("disj_comms", [disj_commute, disj_left_commute]);
191 int_prove "conj_disj_distribL" "P&(Q|R) <-> (P&Q | P&R)";
192 int_prove "conj_disj_distribR" "(P|Q)&R <-> (P&R | Q&R)";
194 int_prove "disj_conj_distribL" "P|(Q&R) <-> (P|Q) & (P|R)";
195 int_prove "disj_conj_distribR" "(P&Q)|R <-> (P|R) & (Q|R)";
197 int_prove "imp_conj_distrib" "(P --> (Q&R)) <-> (P-->Q) & (P-->R)";
198 int_prove "imp_conj" "((P&Q)-->R) <-> (P --> (Q --> R))";
199 int_prove "imp_disj" "(P|Q --> R) <-> (P-->R) & (Q-->R)";
201 prove "imp_disj1" "(P-->Q) | R <-> (P-->Q | R)";
202 prove "imp_disj2" "Q | (P-->R) <-> (P-->Q | R)";
204 int_prove "de_Morgan_disj" "(~(P | Q)) <-> (~P & ~Q)";
205 prove "de_Morgan_conj" "(~(P & Q)) <-> (~P | ~Q)";
207 prove "not_imp" "~(P --> Q) <-> (P & ~Q)";
208 prove "not_iff" "~(P <-> Q) <-> (P <-> ~Q)";
210 prove "not_all" "(~ (ALL x. P(x))) <-> (EX x.~P(x))";
211 prove "imp_all" "((ALL x. P(x)) --> Q) <-> (EX x. P(x) --> Q)";
212 int_prove "not_ex" "(~ (EX x. P(x))) <-> (ALL x.~P(x))";
213 int_prove "imp_ex" "((EX x. P(x)) --> Q) <-> (ALL x. P(x) --> Q)";
215 int_prove "ex_disj_distrib"
216 "(EX x. P(x) | Q(x)) <-> ((EX x. P(x)) | (EX x. Q(x)))";
217 int_prove "all_conj_distrib"
218 "(ALL x. P(x) & Q(x)) <-> ((ALL x. P(x)) & (ALL x. Q(x)))";
222 val uncurry = prove_goal (the_context()) "P --> Q --> R ==> P & Q --> R"
223 (fn prems => [cut_facts_tac prems 1, Blast_tac 1]);
225 val iff_allI = allI RS
226 prove_goal (the_context()) "ALL x. P(x) <-> Q(x) ==> (ALL x. P(x)) <-> (ALL x. Q(x))"
227 (fn prems => [cut_facts_tac prems 1, Blast_tac 1])
228 val iff_exI = allI RS
229 prove_goal (the_context()) "ALL x. P(x) <-> Q(x) ==> (EX x. P(x)) <-> (EX x. Q(x))"
230 (fn prems => [cut_facts_tac prems 1, Blast_tac 1])
232 val all_comm = prove_goal (the_context()) "(ALL x y. P(x,y)) <-> (ALL y x. P(x,y))"
233 (fn _ => [Blast_tac 1])
234 val ex_comm = prove_goal (the_context()) "(EX x y. P(x,y)) <-> (EX y x. P(x,y))"
235 (fn _ => [Blast_tac 1])
238 (** make simplification procedures for quantifier elimination **)
239 structure Quantifier1 = Quantifier1Fun(
242 fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
244 fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
245 | dest_conj _ = NONE;
246 fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
248 val conj = FOLogic.conj
249 val imp = FOLogic.imp
251 val iff_reflection = iff_reflection
253 val iff_trans = iff_trans
258 val uncurry = uncurry
261 val iff_allI = iff_allI
262 val iff_exI = iff_exI
263 val all_comm = all_comm
264 val ex_comm = ex_comm
270 Simplifier.simproc (the_context ())
271 "defined EX" ["EX x. P(x)"] Quantifier1.rearrange_ex;
274 Simplifier.simproc (the_context ())
275 "defined ALL" ["ALL x. P(x)"] Quantifier1.rearrange_all;
278 (*** Case splitting ***)
280 bind_thm ("meta_eq_to_iff", prove_goal IFOL.thy "x==y ==> x<->y"
281 (fn [prem] => [rewtac prem, rtac iffI 1, atac 1, atac 1]));
283 structure SplitterData =
285 structure Simplifier = Simplifier
287 val meta_eq_to_iff = meta_eq_to_iff
292 val contrapos = contrapos
293 val contrapos2 = contrapos2
294 val notnotD = notnotD
297 structure Splitter = SplitterFun(SplitterData);
299 val split_tac = Splitter.split_tac;
300 val split_inside_tac = Splitter.split_inside_tac;
301 val split_asm_tac = Splitter.split_asm_tac;
302 val op addsplits = Splitter.addsplits;
303 val op delsplits = Splitter.delsplits;
304 val Addsplits = Splitter.Addsplits;
305 val Delsplits = Splitter.Delsplits;
308 (*** Standard simpsets ***)
310 structure Induction = InductionFun(struct val spec=IFOL.spec end);
315 bind_thms ("meta_simps",
316 [triv_forall_equality, (* prunes params *)
317 True_implies_equals]); (* prune asms `True' *)
319 bind_thms ("IFOL_simps",
320 [refl RS P_iff_T] @ conj_simps @ disj_simps @ not_simps @
321 imp_simps @ iff_simps @ quant_simps);
323 bind_thm ("notFalseI", int_prove_fun "~False");
324 bind_thms ("triv_rls", [TrueI,refl,reflexive_thm,iff_refl,notFalseI]);
326 fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls@prems),
328 (*No premature instantiation of variables during simplification*)
329 fun safe_solver prems = FIRST'[match_tac (triv_rls@prems),
330 eq_assume_tac, ematch_tac [FalseE]];
332 (*No simprules, but basic infastructure for simplification*)
334 Simplifier.theory_context (the_context ()) empty_ss
335 setsubgoaler asm_simp_tac
336 setSSolver (mk_solver "FOL safe" safe_solver)
337 setSolver (mk_solver "FOL unsafe" unsafe_solver)
338 setmksimps (mksimps mksimps_pairs)
339 setmkcong mk_meta_cong;
342 let val ss0 = Simplifier.clear_ss FOL_basic_ss addsimps ths
343 in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end;
346 (*intuitionistic simprules only*)
347 val IFOL_ss = FOL_basic_ss
348 addsimps (meta_simps @ IFOL_simps @ int_ex_simps @ int_all_simps)
349 addsimprocs [defALL_regroup, defEX_regroup]
352 bind_thms ("cla_simps",
353 [de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2,
354 not_imp, not_all, not_ex, cases_simp] @
356 ["~(P&Q) <-> ~P | ~Q",
358 "~ ~ P <-> P", "(~P --> P) <-> P",
359 "(~P <-> ~Q) <-> (P<->Q)"]);
361 (*classical simprules too*)
362 val FOL_ss = IFOL_ss addsimps (cla_simps @ cla_ex_simps @ cla_all_simps);
364 val simpsetup = [fn thy => (change_simpset_of thy (fn _ => FOL_ss); thy)];
367 (*** integration of simplifier with classical reasoner ***)
369 structure Clasimp = ClasimpFun
370 (structure Simplifier = Simplifier and Splitter = Splitter
371 and Classical = Cla and Blast = Blast
372 val iffD1 = iffD1 val iffD2 = iffD2 val notE = notE);
375 val FOL_css = (FOL_cs, FOL_ss);