| author | wenzelm | 
| Mon, 16 Nov 2009 13:53:02 +0100 | |
| changeset 33712 | cffc97238102 | 
| parent 33339 | d41f77196338 | 
| child 35232 | f588e1169c8b | 
| permissions | -rw-r--r-- | 
| 21163 | 1  | 
(* Title: HOL/simpdata.ML  | 
2  | 
Author: Tobias Nipkow  | 
|
3  | 
Copyright 1991 University of Cambridge  | 
|
4  | 
||
5  | 
Instantiation of the generic simplifier for HOL.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
(** tools setup **)  | 
|
9  | 
||
10  | 
structure Quantifier1 = Quantifier1Fun  | 
|
11  | 
(struct  | 
|
12  | 
(*abstract syntax*)  | 
|
13  | 
  fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t)
 | 
|
14  | 
| dest_eq _ = NONE;  | 
|
15  | 
  fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t)
 | 
|
16  | 
| dest_conj _ = NONE;  | 
|
17  | 
  fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t)
 | 
|
18  | 
| dest_imp _ = NONE;  | 
|
19  | 
val conj = HOLogic.conj  | 
|
20  | 
val imp = HOLogic.imp  | 
|
21  | 
(*rules*)  | 
|
| 22147 | 22  | 
  val iff_reflection = @{thm eq_reflection}
 | 
23  | 
  val iffI = @{thm iffI}
 | 
|
24  | 
  val iff_trans = @{thm trans}
 | 
|
25  | 
  val conjI= @{thm conjI}
 | 
|
26  | 
  val conjE= @{thm conjE}
 | 
|
27  | 
  val impI = @{thm impI}
 | 
|
28  | 
  val mp   = @{thm mp}
 | 
|
29  | 
  val uncurry = @{thm uncurry}
 | 
|
30  | 
  val exI  = @{thm exI}
 | 
|
31  | 
  val exE  = @{thm exE}
 | 
|
32  | 
  val iff_allI = @{thm iff_allI}
 | 
|
33  | 
  val iff_exI = @{thm iff_exI}
 | 
|
34  | 
  val all_comm = @{thm all_comm}
 | 
|
35  | 
  val ex_comm = @{thm ex_comm}
 | 
|
| 21163 | 36  | 
end);  | 
37  | 
||
| 21551 | 38  | 
structure Simpdata =  | 
| 21163 | 39  | 
struct  | 
40  | 
||
| 22147 | 41  | 
fun mk_meta_eq r = r RS @{thm eq_reflection};
 | 
| 21163 | 42  | 
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r;  | 
43  | 
||
| 22147 | 44  | 
fun mk_eq th = case concl_of th  | 
| 21163 | 45  | 
(*expects Trueprop if not == *)  | 
| 21551 | 46  | 
  of Const ("==",_) $ _ $ _ => th
 | 
47  | 
   | _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq th
 | 
|
| 22147 | 48  | 
   | _ $ (Const ("Not", _) $ _) => th RS @{thm Eq_FalseI}
 | 
49  | 
   | _ => th RS @{thm Eq_TrueI}
 | 
|
| 21163 | 50  | 
|
| 22147 | 51  | 
fun mk_eq_True r =  | 
52  | 
  SOME (r RS @{thm meta_eq_to_obj_eq} RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE;
 | 
|
| 21163 | 53  | 
|
54  | 
(* Produce theorems of the form  | 
|
55  | 
(P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y)  | 
|
56  | 
*)  | 
|
| 22838 | 57  | 
|
| 22147 | 58  | 
fun lift_meta_eq_to_obj_eq i st =  | 
| 21163 | 59  | 
let  | 
60  | 
    fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
 | 
|
61  | 
| count_imp _ = 0;  | 
|
62  | 
val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1)))  | 
|
| 22147 | 63  | 
  in if j = 0 then @{thm meta_eq_to_obj_eq}
 | 
| 21163 | 64  | 
else  | 
65  | 
let  | 
|
66  | 
        val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
 | 
|
| 33339 | 67  | 
fun mk_simp_implies Q = fold_rev (fn R => fn S =>  | 
68  | 
          Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Ps Q
 | 
|
| 21163 | 69  | 
        val aT = TFree ("'a", HOLogic.typeS);
 | 
70  | 
        val x = Free ("x", aT);
 | 
|
71  | 
        val y = Free ("y", aT)
 | 
|
72  | 
in Goal.prove_global (Thm.theory_of_thm st) []  | 
|
73  | 
[mk_simp_implies (Logic.mk_equals (x, y))]  | 
|
74  | 
(mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y))))  | 
|
| 26711 | 75  | 
        (fn {prems, ...} => EVERY
 | 
| 22147 | 76  | 
         [rewrite_goals_tac @{thms simp_implies_def},
 | 
77  | 
          REPEAT (ares_tac (@{thm meta_eq_to_obj_eq} ::
 | 
|
78  | 
            map (rewrite_rule @{thms simp_implies_def}) prems) 1)])
 | 
|
| 21163 | 79  | 
end  | 
80  | 
end;  | 
|
81  | 
||
82  | 
(*Congruence rules for = (instead of ==)*)  | 
|
83  | 
fun mk_meta_cong rl = zero_var_indexes  | 
|
84  | 
(let val rl' = Seq.hd (TRYALL (fn i => fn st =>  | 
|
85  | 
rtac (lift_meta_eq_to_obj_eq i st) i st) rl)  | 
|
86  | 
in mk_meta_eq rl' handle THM _ =>  | 
|
87  | 
if can Logic.dest_equals (concl_of rl') then rl'  | 
|
88  | 
else error "Conclusion of congruence rules must be =-equality"  | 
|
89  | 
end);  | 
|
90  | 
||
91  | 
fun mk_atomize pairs =  | 
|
92  | 
let  | 
|
| 
21313
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
93  | 
fun atoms thm =  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
94  | 
let  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
95  | 
fun res th = map (fn rl => th RS rl); (*exception THM*)  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
96  | 
fun res_fixed rls =  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
97  | 
if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
98  | 
else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.thm_context thm) [thm];  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
99  | 
in  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
100  | 
case concl_of thm  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
101  | 
          of Const ("Trueprop", _) $ p => (case head_of p
 | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
102  | 
of Const (a, _) => (case AList.lookup (op =) pairs a  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
103  | 
of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm])  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
104  | 
| NONE => [thm])  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
105  | 
| _ => [thm])  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
106  | 
| _ => [thm]  | 
| 
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
107  | 
end;  | 
| 21163 | 108  | 
in atoms end;  | 
109  | 
||
110  | 
fun mksimps pairs =  | 
|
| 
21313
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
111  | 
map_filter (try mk_eq) o mk_atomize pairs o gen_all;  | 
| 21163 | 112  | 
|
| 22147 | 113  | 
fun unsafe_solver_tac prems =  | 
114  | 
  (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN'
 | 
|
115  | 
  FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), atac,
 | 
|
116  | 
    etac @{thm FalseE}];
 | 
|
117  | 
||
| 21163 | 118  | 
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac;  | 
119  | 
||
| 22838 | 120  | 
|
| 21163 | 121  | 
(*No premature instantiation of variables during simplification*)  | 
| 22147 | 122  | 
fun safe_solver_tac prems =  | 
123  | 
  (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN'
 | 
|
124  | 
  FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems),
 | 
|
125  | 
         eq_assume_tac, ematch_tac @{thms FalseE}];
 | 
|
126  | 
||
| 21163 | 127  | 
val safe_solver = mk_solver "HOL safe" safe_solver_tac;  | 
128  | 
||
| 
32177
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
129  | 
structure Splitter = Splitter  | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
130  | 
(  | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
131  | 
  val thy = @{theory}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
132  | 
val mk_eq = mk_eq  | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
133  | 
  val meta_eq_to_iff = @{thm meta_eq_to_obj_eq}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
134  | 
  val iffD = @{thm iffD2}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
135  | 
  val disjE = @{thm disjE}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
136  | 
  val conjE = @{thm conjE}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
137  | 
  val exE = @{thm exE}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
138  | 
  val contrapos = @{thm contrapos_nn}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
139  | 
  val contrapos2 = @{thm contrapos_pp}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
140  | 
  val notnotD = @{thm notnotD}
 | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
141  | 
);  | 
| 21163 | 142  | 
|
| 
32177
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
143  | 
val split_tac = Splitter.split_tac;  | 
| 21674 | 144  | 
val split_inside_tac = Splitter.split_inside_tac;  | 
145  | 
||
| 
32177
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
146  | 
val op addsplits = Splitter.addsplits;  | 
| 
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
147  | 
val op delsplits = Splitter.delsplits;  | 
| 21674 | 148  | 
|
149  | 
||
| 21163 | 150  | 
(* integration of simplifier with classical reasoner *)  | 
151  | 
||
152  | 
structure Clasimp = ClasimpFun  | 
|
153  | 
(structure Simplifier = Simplifier and Splitter = Splitter  | 
|
154  | 
and Classical = Classical and Blast = Blast  | 
|
| 22147 | 155  | 
  val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE});
 | 
| 21674 | 156  | 
open Clasimp;  | 
| 21163 | 157  | 
|
| 27338 | 158  | 
val _ = ML_Antiquote.value "clasimpset"  | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
159  | 
(Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())");  | 
| 22128 | 160  | 
|
| 21163 | 161  | 
val mksimps_pairs =  | 
| 22147 | 162  | 
  [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
 | 
163  | 
   ("All", [@{thm spec}]), ("True", []), ("False", []),
 | 
|
164  | 
   ("HOL.If", [@{thm if_bool_eq_conj} RS @{thm iffD1}])];
 | 
|
| 21163 | 165  | 
|
| 21674 | 166  | 
val HOL_basic_ss =  | 
| 32155 | 167  | 
  Simplifier.theory_context @{theory} empty_ss
 | 
| 21163 | 168  | 
setsubgoaler asm_simp_tac  | 
169  | 
setSSolver safe_solver  | 
|
170  | 
setSolver unsafe_solver  | 
|
171  | 
setmksimps (mksimps mksimps_pairs)  | 
|
172  | 
setmkeqTrue mk_eq_True  | 
|
173  | 
setmkcong mk_meta_cong;  | 
|
174  | 
||
| 21674 | 175  | 
fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews);  | 
| 21163 | 176  | 
|
177  | 
fun unfold_tac ths =  | 
|
| 21674 | 178  | 
let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths  | 
| 21163 | 179  | 
in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end;  | 
180  | 
||
181  | 
val defALL_regroup =  | 
|
| 32010 | 182  | 
  Simplifier.simproc @{theory}
 | 
| 21163 | 183  | 
"defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all;  | 
184  | 
||
185  | 
val defEX_regroup =  | 
|
| 32010 | 186  | 
  Simplifier.simproc @{theory}
 | 
| 21163 | 187  | 
"defined EX" ["EX x. P x"] Quantifier1.rearrange_ex;  | 
188  | 
||
189  | 
||
| 24035 | 190  | 
val simpset_simprocs = HOL_basic_ss addsimprocs [defALL_regroup, defEX_regroup]  | 
| 21163 | 191  | 
|
| 
21313
 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 
wenzelm 
parents: 
21163 
diff
changeset
 | 
192  | 
end;  | 
| 21551 | 193  | 
|
194  | 
structure Splitter = Simpdata.Splitter;  | 
|
195  | 
structure Clasimp = Simpdata.Clasimp;  |