author | wenzelm |
Tue, 17 Sep 2024 11:14:25 +0200 | |
changeset 80894 | 3e0ca6af6738 |
parent 71916 | 435cdc772110 |
child 82641 | d22294b20573 |
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*) |
69593 | 13 |
fun dest_eq (Const(\<^const_name>\<open>HOL.eq\<close>,_) $ s $ t) = SOME (s, t) |
21163 | 14 |
| dest_eq _ = NONE; |
69593 | 15 |
fun dest_conj (Const(\<^const_name>\<open>HOL.conj\<close>,_) $ s $ t) = SOME (s, t) |
21163 | 16 |
| dest_conj _ = NONE; |
69593 | 17 |
fun dest_imp (Const(\<^const_name>\<open>HOL.implies\<close>,_) $ 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} |
|
71916
435cdc772110
specific atomization inert to later rule set modifications
haftmann
parents:
69593
diff
changeset
|
36 |
val atomize = |
435cdc772110
specific atomization inert to later rule set modifications
haftmann
parents:
69593
diff
changeset
|
37 |
let val rules = @{thms atomize_all atomize_imp atomize_eq atomize_conj} |
435cdc772110
specific atomization inert to later rule set modifications
haftmann
parents:
69593
diff
changeset
|
38 |
in fn ctxt => Raw_Simplifier.rewrite ctxt true rules end |
42458 | 39 |
); |
21163 | 40 |
|
21551 | 41 |
structure Simpdata = |
21163 | 42 |
struct |
43 |
||
22147 | 44 |
fun mk_meta_eq r = r RS @{thm eq_reflection}; |
21163 | 45 |
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; |
46 |
||
59582 | 47 |
fun mk_eq th = |
48 |
(case Thm.concl_of th of |
|
21163 | 49 |
(*expects Trueprop if not == *) |
69593 | 50 |
Const (\<^const_name>\<open>Pure.eq\<close>,_) $ _ $ _ => th |
51 |
| _ $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _) => mk_meta_eq th |
|
52 |
| _ $ (Const (\<^const_name>\<open>Not\<close>, _) $ _) => th RS @{thm Eq_FalseI} |
|
59582 | 53 |
| _ => th RS @{thm Eq_TrueI}) |
21163 | 54 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
55 |
fun mk_eq_True (_: Proof.context) r = |
67710
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents:
67329
diff
changeset
|
56 |
SOME (HOLogic.mk_obj_eq r RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE; |
21163 | 57 |
|
58 |
(* Produce theorems of the form |
|
59 |
(P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y) |
|
60 |
*) |
|
22838 | 61 |
|
60651 | 62 |
fun lift_meta_eq_to_obj_eq ctxt i st = |
21163 | 63 |
let |
69593 | 64 |
fun count_imp (Const (\<^const_name>\<open>HOL.simp_implies\<close>, _) $ _ $ P) = 1 + count_imp P |
21163 | 65 |
| count_imp _ = 0; |
45622 | 66 |
val j = count_imp (Logic.strip_assums_concl (Thm.term_of (Thm.cprem_of st i))) |
42364 | 67 |
in |
68 |
if j = 0 then @{thm meta_eq_to_obj_eq} |
|
21163 | 69 |
else |
70 |
let |
|
71 |
val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); |
|
45622 | 72 |
val mk_simp_implies = fold_rev (fn R => fn S => |
69593 | 73 |
Const (\<^const_name>\<open>HOL.simp_implies\<close>, propT --> propT --> propT) $ R $ S) Ps; |
42364 | 74 |
in |
60651 | 75 |
Goal.prove_global (Proof_Context.theory_of ctxt) [] |
69593 | 76 |
[mk_simp_implies \<^prop>\<open>(x::'a) == y\<close>] |
77 |
(mk_simp_implies \<^prop>\<open>(x::'a) = y\<close>) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
78 |
(fn {context = ctxt, prems} => EVERY |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
79 |
[rewrite_goals_tac ctxt @{thms simp_implies_def}, |
59499 | 80 |
REPEAT (assume_tac ctxt 1 ORELSE |
81 |
resolve_tac ctxt |
|
82 |
(@{thm meta_eq_to_obj_eq} :: |
|
83 |
map (rewrite_rule ctxt @{thms simp_implies_def}) prems) 1)]) |
|
21163 | 84 |
end |
85 |
end; |
|
86 |
||
87 |
(*Congruence rules for = (instead of ==)*) |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
88 |
fun mk_meta_cong ctxt rl = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
89 |
let |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
90 |
val rl' = Seq.hd (TRYALL (fn i => fn st => |
60651 | 91 |
resolve_tac ctxt [lift_meta_eq_to_obj_eq ctxt i st] i st) rl) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
92 |
in |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
93 |
mk_meta_eq rl' handle THM _ => |
59582 | 94 |
if can Logic.dest_equals (Thm.concl_of rl') then rl' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
95 |
else error "Conclusion of congruence rules must be =-equality" |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
96 |
end |> zero_var_indexes; |
21163 | 97 |
|
59169 | 98 |
fun mk_atomize ctxt pairs = |
21163 | 99 |
let |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
100 |
fun atoms thm = |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
101 |
let |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
102 |
fun res th = map (fn rl => th RS rl); (*exception THM*) |
59169 | 103 |
val thm_ctxt = Variable.declare_thm thm ctxt; |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
104 |
fun res_fixed rls = |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
105 |
if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls |
59169 | 106 |
else Variable.trade (K (fn [thm'] => res thm' rls)) thm_ctxt [thm]; |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
107 |
in |
59582 | 108 |
case Thm.concl_of thm |
69593 | 109 |
of Const (\<^const_name>\<open>Trueprop\<close>, _) $ p => (case head_of p |
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
110 |
of Const (a, _) => (case AList.lookup (op =) pairs a |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
111 |
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
|
112 |
| NONE => [thm]) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
113 |
| _ => [thm]) |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
114 |
| _ => [thm] |
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
115 |
end; |
21163 | 116 |
in atoms end; |
117 |
||
60822 | 118 |
fun mksimps pairs ctxt = map_filter (try mk_eq) o mk_atomize ctxt pairs o Variable.gen_all ctxt; |
21163 | 119 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
120 |
fun unsafe_solver_tac ctxt = |
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54998
diff
changeset
|
121 |
let |
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54998
diff
changeset
|
122 |
val sol_thms = |
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54998
diff
changeset
|
123 |
reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt; |
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54998
diff
changeset
|
124 |
fun sol_tac i = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
125 |
FIRST |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
126 |
[resolve_tac ctxt sol_thms i, |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
127 |
assume_tac ctxt i, |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59169
diff
changeset
|
128 |
eresolve_tac ctxt @{thms FalseE} i] ORELSE |
60171 | 129 |
(match_tac ctxt [@{thm conjI}] |
130 |
THEN_ALL_NEW sol_tac) i |
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54998
diff
changeset
|
131 |
in |
58957 | 132 |
(fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' sol_tac |
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54998
diff
changeset
|
133 |
end; |
22147 | 134 |
|
21163 | 135 |
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; |
136 |
||
22838 | 137 |
|
21163 | 138 |
(*No premature instantiation of variables during simplification*) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
139 |
fun safe_solver_tac ctxt = |
58957 | 140 |
(fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' |
141 |
FIRST' [match_tac ctxt (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt), |
|
142 |
eq_assume_tac, ematch_tac ctxt @{thms FalseE}]; |
|
22147 | 143 |
|
21163 | 144 |
val safe_solver = mk_solver "HOL safe" safe_solver_tac; |
145 |
||
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
146 |
structure Splitter = Splitter |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
147 |
( |
69593 | 148 |
val context = \<^context> |
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
149 |
val mk_eq = mk_eq |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
150 |
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
|
151 |
val iffD = @{thm iffD2} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
152 |
val disjE = @{thm disjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
153 |
val conjE = @{thm conjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
154 |
val exE = @{thm exE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
155 |
val contrapos = @{thm contrapos_nn} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
156 |
val contrapos2 = @{thm contrapos_pp} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
157 |
val notnotD = @{thm notnotD} |
63636 | 158 |
val safe_tac = Classical.safe_tac |
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
159 |
); |
21163 | 160 |
|
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
161 |
val split_tac = Splitter.split_tac; |
21674 | 162 |
val split_inside_tac = Splitter.split_inside_tac; |
163 |
||
164 |
||
21163 | 165 |
(* integration of simplifier with classical reasoner *) |
166 |
||
42478 | 167 |
structure Clasimp = Clasimp |
168 |
( |
|
169 |
structure Simplifier = Simplifier |
|
170 |
and Splitter = Splitter |
|
171 |
and Classical = Classical |
|
172 |
and Blast = Blast |
|
173 |
val iffD1 = @{thm iffD1} |
|
174 |
val iffD2 = @{thm iffD2} |
|
175 |
val notE = @{thm notE} |
|
176 |
); |
|
21674 | 177 |
open Clasimp; |
21163 | 178 |
|
179 |
val mksimps_pairs = |
|
69593 | 180 |
[(\<^const_name>\<open>HOL.implies\<close>, [@{thm mp}]), |
181 |
(\<^const_name>\<open>HOL.conj\<close>, [@{thm conjunct1}, @{thm conjunct2}]), |
|
182 |
(\<^const_name>\<open>All\<close>, [@{thm spec}]), |
|
183 |
(\<^const_name>\<open>True\<close>, []), |
|
184 |
(\<^const_name>\<open>False\<close>, []), |
|
185 |
(\<^const_name>\<open>If\<close>, [@{thm if_bool_eq_conj} RS @{thm iffD1}])]; |
|
21163 | 186 |
|
21674 | 187 |
val HOL_basic_ss = |
69593 | 188 |
empty_simpset \<^context> |
45625
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45622
diff
changeset
|
189 |
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
|
190 |
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
|
191 |
|> 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
|
192 |
|> 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
|
193 |
|> Simplifier.set_mkeqTrue mk_eq_True |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
194 |
|> Simplifier.set_mkcong mk_meta_cong |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
195 |
|> simpset_of; |
21163 | 196 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
197 |
fun hol_simplify ctxt rews = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
198 |
Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps rews); |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
199 |
|
54998 | 200 |
fun unfold_tac ctxt ths = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45625
diff
changeset
|
201 |
ALLGOALS (full_simp_tac (clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths)); |
21163 | 202 |
|
21313
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
wenzelm
parents:
21163
diff
changeset
|
203 |
end; |
21551 | 204 |
|
205 |
structure Splitter = Simpdata.Splitter; |
|
206 |
structure Clasimp = Simpdata.Clasimp; |