author | haftmann |
Fri, 15 Feb 2013 08:31:31 +0100 | |
changeset 51143 | 0a2371e7ced3 |
parent 45625 | 750c5a47400b |
child 51717 | 9e7d1c139569 |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Tools/simpdata.ML |
21163 | 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 |
||
42458 | 10 |
structure Quantifier1 = Quantifier1 |
11 |
( |
|
21163 | 12 |
(*abstract syntax*) |
42460 | 13 |
fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t) |
21163 | 14 |
| dest_eq _ = NONE; |
42460 | 15 |
fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t) |
21163 | 16 |
| dest_conj _ = NONE; |
42460 | 17 |
fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t) |
21163 | 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} |
|
42458 | 36 |
); |
21163 | 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 == *) |
35364 | 46 |
of Const (@{const_name "=="},_) $ _ $ _ => th |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
47 |
| _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => mk_meta_eq th |
38558 | 48 |
| _ $ (Const (@{const_name Not}, _) $ _) => th RS @{thm Eq_FalseI} |
22147 | 49 |
| _ => th RS @{thm Eq_TrueI} |
21163 | 50 |
|
36543
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
wenzelm
parents:
35364
diff
changeset
|
51 |
fun mk_eq_True (_: simpset) r = |
22147 | 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 |
35364 | 60 |
fun count_imp (Const (@{const_name HOL.simp_implies}, _) $ P $ Q) = 1 + count_imp Q |
21163 | 61 |
| count_imp _ = 0; |
45622 | 62 |
val j = count_imp (Logic.strip_assums_concl (Thm.term_of (Thm.cprem_of st i))) |
42364 | 63 |
in |
64 |
if j = 0 then @{thm meta_eq_to_obj_eq} |
|
21163 | 65 |
else |
66 |
let |
|
67 |
val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); |
|
45622 | 68 |
val mk_simp_implies = fold_rev (fn R => fn S => |
69 |
Const (@{const_name HOL.simp_implies}, propT --> propT --> propT) $ R $ S) Ps; |
|
42364 | 70 |
in |
71 |
Goal.prove_global (Thm.theory_of_thm st) [] |
|
45622 | 72 |
[mk_simp_implies @{prop "(x::'a) == y"}] |
73 |
(mk_simp_implies @{prop "(x::'a) = y"}) |
|
42364 | 74 |
(fn {prems, ...} => EVERY |
75 |
[rewrite_goals_tac @{thms simp_implies_def}, |
|
76 |
REPEAT (ares_tac (@{thm meta_eq_to_obj_eq} :: |
|
77 |
map (rewrite_rule @{thms simp_implies_def}) prems) 1)]) |
|
21163 | 78 |
end |
79 |
end; |
|
80 |
||
81 |
(*Congruence rules for = (instead of ==)*) |
|
36543
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
wenzelm
parents:
35364
diff
changeset
|
82 |
fun mk_meta_cong (_: simpset) rl = zero_var_indexes |
21163 | 83 |
(let val rl' = Seq.hd (TRYALL (fn i => fn st => |
84 |
rtac (lift_meta_eq_to_obj_eq i st) i st) rl) |
|
85 |
in mk_meta_eq rl' handle THM _ => |
|
86 |
if can Logic.dest_equals (concl_of rl') then rl' |
|
87 |
else error "Conclusion of congruence rules must be =-equality" |
|
88 |
end); |
|
89 |
||
90 |
fun mk_atomize pairs = |
|
91 |
let |
|
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
92 |
fun atoms thm = |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
93 |
let |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
94 |
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
|
95 |
fun res_fixed rls = |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
96 |
if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls |
36603
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
wenzelm
parents:
36543
diff
changeset
|
97 |
else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.global_thm_context thm) [thm]; |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
98 |
in |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
99 |
case concl_of thm |
35364 | 100 |
of Const (@{const_name Trueprop}, _) $ p => (case head_of p |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
101 |
of Const (a, _) => (case AList.lookup (op =) pairs a |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
102 |
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
|
103 |
| NONE => [thm]) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
104 |
| _ => [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 |
end; |
21163 | 107 |
in atoms end; |
108 |
||
36543
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
wenzelm
parents:
35364
diff
changeset
|
109 |
fun mksimps pairs (_: simpset) = |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
110 |
map_filter (try mk_eq) o mk_atomize pairs o gen_all; |
21163 | 111 |
|
43596 | 112 |
fun unsafe_solver_tac ss = |
22147 | 113 |
(fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' |
43597 | 114 |
FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ss), |
115 |
atac, etac @{thm FalseE}]; |
|
22147 | 116 |
|
21163 | 117 |
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; |
118 |
||
22838 | 119 |
|
21163 | 120 |
(*No premature instantiation of variables during simplification*) |
43596 | 121 |
fun safe_solver_tac ss = |
22147 | 122 |
(fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' |
43597 | 123 |
FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ss), |
124 |
eq_assume_tac, ematch_tac @{thms FalseE}]; |
|
22147 | 125 |
|
21163 | 126 |
val safe_solver = mk_solver "HOL safe" safe_solver_tac; |
127 |
||
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
128 |
structure Splitter = Splitter |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
129 |
( |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
130 |
val thy = @{theory} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
131 |
val mk_eq = mk_eq |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
132 |
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
|
133 |
val iffD = @{thm iffD2} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
134 |
val disjE = @{thm disjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
135 |
val conjE = @{thm conjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
136 |
val exE = @{thm exE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
137 |
val contrapos = @{thm contrapos_nn} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
138 |
val contrapos2 = @{thm contrapos_pp} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
139 |
val notnotD = @{thm notnotD} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
140 |
); |
21163 | 141 |
|
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
142 |
val split_tac = Splitter.split_tac; |
21674 | 143 |
val split_inside_tac = Splitter.split_inside_tac; |
144 |
||
145 |
||
21163 | 146 |
(* integration of simplifier with classical reasoner *) |
147 |
||
42478 | 148 |
structure Clasimp = Clasimp |
149 |
( |
|
150 |
structure Simplifier = Simplifier |
|
151 |
and Splitter = Splitter |
|
152 |
and Classical = Classical |
|
153 |
and Blast = Blast |
|
154 |
val iffD1 = @{thm iffD1} |
|
155 |
val iffD2 = @{thm iffD2} |
|
156 |
val notE = @{thm notE} |
|
157 |
); |
|
21674 | 158 |
open Clasimp; |
21163 | 159 |
|
160 |
val mksimps_pairs = |
|
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38715
diff
changeset
|
161 |
[(@{const_name HOL.implies}, [@{thm mp}]), |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
162 |
(@{const_name HOL.conj}, [@{thm conjunct1}, @{thm conjunct2}]), |
35364 | 163 |
(@{const_name All}, [@{thm spec}]), |
164 |
(@{const_name True}, []), |
|
165 |
(@{const_name False}, []), |
|
166 |
(@{const_name If}, [@{thm if_bool_eq_conj} RS @{thm iffD1}])]; |
|
21163 | 167 |
|
21674 | 168 |
val HOL_basic_ss = |
35232
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents:
33339
diff
changeset
|
169 |
Simplifier.global_context @{theory} empty_ss |
45625
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
170 |
setSSolver safe_solver |
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
171 |
setSolver unsafe_solver |
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
172 |
|> Simplifier.set_subgoaler asm_simp_tac |
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
173 |
|> Simplifier.set_mksimps (mksimps mksimps_pairs) |
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
174 |
|> Simplifier.set_mkeqTrue mk_eq_True |
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
175 |
|> Simplifier.set_mkcong mk_meta_cong; |
21163 | 176 |
|
21674 | 177 |
fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); |
21163 | 178 |
|
179 |
fun unfold_tac ths = |
|
21674 | 180 |
let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths |
21163 | 181 |
in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; |
182 |
||
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
183 |
end; |
21551 | 184 |
|
185 |
structure Splitter = Simpdata.Splitter; |
|
186 |
structure Clasimp = Simpdata.Clasimp; |