1 (* Title: FOL/simpdata.ML
2 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
3 Copyright 1994 University of Cambridge
5 Simplification data for FOL.
8 (*Make meta-equalities. The operator below is Trueprop*)
10 fun mk_meta_eq th = case concl_of th of
11 _ $ (Const("op =",_)$_$_) => th RS @{thm eq_reflection}
12 | _ $ (Const("op <->",_)$_$_) => th RS @{thm iff_reflection}
14 error("conclusion must be a =-equality or <->");;
16 fun mk_eq th = case concl_of th of
17 Const("==",_)$_$_ => th
18 | _ $ (Const("op =",_)$_$_) => mk_meta_eq th
19 | _ $ (Const("op <->",_)$_$_) => mk_meta_eq th
20 | _ $ (Const("Not",_)$_) => th RS @{thm iff_reflection_F}
21 | _ => th RS @{thm iff_reflection_T};
23 (*Replace premises x=y, X<->Y by X==Y*)
26 (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
28 (*Congruence rules for = or <-> (instead of ==)*)
30 standard(mk_meta_eq (mk_meta_prems rl))
32 error("Premises and conclusion of congruence rules must use =-equality or <->");
35 [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
36 ("All", [@{thm spec}]), ("True", []), ("False", [])];
38 (* ###FIXME: move to simplifier.ML
39 val mk_atomize: (string * thm list) list -> thm -> thm list
41 (* ###FIXME: move to simplifier.ML *)
42 fun mk_atomize pairs =
45 Const("Trueprop",_) $ p =>
48 (case AList.lookup (op =) pairs a of
49 SOME(rls) => List.concat (map atoms ([th] RL rls))
55 fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all);
58 (** make simplification procedures for quantifier elimination **)
59 structure Quantifier1 = Quantifier1Fun(
62 fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
64 fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
66 fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
68 val conj = FOLogic.conj
71 val iff_reflection = @{thm iff_reflection}
72 val iffI = @{thm iffI}
73 val iff_trans = @{thm iff_trans}
74 val conjI= @{thm conjI}
75 val conjE= @{thm conjE}
76 val impI = @{thm impI}
78 val uncurry = @{thm uncurry}
81 val iff_allI = @{thm iff_allI}
82 val iff_exI = @{thm iff_exI}
83 val all_comm = @{thm all_comm}
84 val ex_comm = @{thm ex_comm}
88 Simplifier.simproc @{theory}
89 "defined EX" ["EX x. P(x)"] Quantifier1.rearrange_ex;
92 Simplifier.simproc @{theory}
93 "defined ALL" ["ALL x. P(x)"] Quantifier1.rearrange_all;
96 (*** Case splitting ***)
98 structure Splitter = Splitter
102 val meta_eq_to_iff = @{thm meta_eq_to_iff}
103 val iffD = @{thm iffD2}
104 val disjE = @{thm disjE}
105 val conjE = @{thm conjE}
107 val contrapos = @{thm contrapos}
108 val contrapos2 = @{thm contrapos2}
109 val notnotD = @{thm notnotD}
112 val split_tac = Splitter.split_tac;
113 val split_inside_tac = Splitter.split_inside_tac;
114 val split_asm_tac = Splitter.split_asm_tac;
115 val op addsplits = Splitter.addsplits;
116 val op delsplits = Splitter.delsplits;
119 (*** Standard simpsets ***)
121 val triv_rls = [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl}, @{thm notFalseI}];
123 fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls @ prems),
124 atac, etac @{thm FalseE}];
125 (*No premature instantiation of variables during simplification*)
126 fun safe_solver prems = FIRST'[match_tac (triv_rls @ prems),
127 eq_assume_tac, ematch_tac [@{thm FalseE}]];
129 (*No simprules, but basic infastructure for simplification*)
131 Simplifier.theory_context @{theory} empty_ss
132 setsubgoaler asm_simp_tac
133 setSSolver (mk_solver "FOL safe" safe_solver)
134 setSolver (mk_solver "FOL unsafe" unsafe_solver)
135 setmksimps (mksimps mksimps_pairs)
136 setmkcong mk_meta_cong;
139 let val ss0 = Simplifier.clear_ss FOL_basic_ss addsimps ths
140 in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end;
143 (*intuitionistic simprules only*)
146 addsimps (@{thms meta_simps} @ @{thms IFOL_simps} @ @{thms int_ex_simps} @ @{thms int_all_simps})
147 addsimprocs [defALL_regroup, defEX_regroup]
148 addcongs [@{thm imp_cong}];
150 (*classical simprules too*)
151 val FOL_ss = IFOL_ss addsimps (@{thms cla_simps} @ @{thms cla_ex_simps} @ @{thms cla_all_simps});
153 val simpsetup = Simplifier.map_simpset (K FOL_ss);
156 (*** integration of simplifier with classical reasoner ***)
158 structure Clasimp = ClasimpFun
159 (structure Simplifier = Simplifier and Splitter = Splitter
160 and Classical = Cla and Blast = Blast
161 val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE});
164 ML_Antiquote.value "clasimpset"
165 (Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())");
167 val FOL_css = (FOL_cs, FOL_ss);