author | berghofe |
Fri, 28 Sep 2001 17:19:46 +0200 | |
changeset 11624 | 8a45c7abef04 |
parent 11534 | 0494d0347f18 |
child 11684 | abd396ca7ef9 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/simpdata.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Tobias Nipkow |
923 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
5304 | 6 |
Instantiation of the generic simplifier for HOL. |
923 | 7 |
*) |
8 |
||
1984 | 9 |
section "Simplifier"; |
10 |
||
7357 | 11 |
val [prem] = goal (the_context ()) "x==y ==> x=y"; |
7031 | 12 |
by (rewtac prem); |
13 |
by (rtac refl 1); |
|
14 |
qed "meta_eq_to_obj_eq"; |
|
4640 | 15 |
|
9023 | 16 |
Goal "(%s. f s) = f"; |
17 |
br refl 1; |
|
18 |
qed "eta_contract_eq"; |
|
19 |
||
923 | 20 |
local |
21 |
||
7357 | 22 |
fun prover s = prove_goal (the_context ()) s (fn _ => [(Blast_tac 1)]); |
923 | 23 |
|
2134 | 24 |
in |
25 |
||
5552 | 26 |
(*Make meta-equalities. The operator below is Trueprop*) |
27 |
||
6128 | 28 |
fun mk_meta_eq r = r RS eq_reflection; |
9832 | 29 |
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; |
6128 | 30 |
|
31 |
val Eq_TrueI = mk_meta_eq(prover "P --> (P = True)" RS mp); |
|
32 |
val Eq_FalseI = mk_meta_eq(prover "~P --> (P = False)" RS mp); |
|
5304 | 33 |
|
6128 | 34 |
fun mk_eq th = case concl_of th of |
35 |
Const("==",_)$_$_ => th |
|
36 |
| _$(Const("op =",_)$_$_) => mk_meta_eq th |
|
37 |
| _$(Const("Not",_)$_) => th RS Eq_FalseI |
|
38 |
| _ => th RS Eq_TrueI; |
|
39 |
(* last 2 lines requires all formulae to be of the from Trueprop(.) *) |
|
5304 | 40 |
|
11624
8a45c7abef04
mksimps and mk_eq_True no longer raise THM exception.
berghofe
parents:
11534
diff
changeset
|
41 |
fun mk_eq_True r = |
8a45c7abef04
mksimps and mk_eq_True no longer raise THM exception.
berghofe
parents:
11534
diff
changeset
|
42 |
Some (r RS meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => None; |
5552 | 43 |
|
9713 | 44 |
(*Congruence rules for = (instead of ==)*) |
6128 | 45 |
fun mk_meta_cong rl = |
46 |
standard(mk_meta_eq(replicate (nprems_of rl) meta_eq_to_obj_eq MRS rl)) |
|
47 |
handle THM _ => |
|
48 |
error("Premises and conclusion of congruence rules must be =-equalities"); |
|
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
49 |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
50 |
val not_not = prover "(~ ~ P) = P"; |
923 | 51 |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
52 |
val simp_thms = [not_not] @ map prover |
2082 | 53 |
[ "(x=x) = True", |
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
54 |
"(~True) = False", "(~False) = True", |
2082 | 55 |
"(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", |
4640 | 56 |
"(True=P) = P", "(P=True) = P", "(False=P) = (~P)", "(P=False) = (~P)", |
9713 | 57 |
"(True --> P) = P", "(False --> P) = True", |
2082 | 58 |
"(P --> True) = True", "(P --> P) = True", |
59 |
"(P --> False) = (~P)", "(P --> ~P) = (~P)", |
|
9713 | 60 |
"(P & True) = P", "(True & P) = P", |
2800 | 61 |
"(P & False) = False", "(False & P) = False", |
62 |
"(P & P) = P", "(P & (P & Q)) = (P & Q)", |
|
3913 | 63 |
"(P & ~P) = False", "(~P & P) = False", |
9713 | 64 |
"(P | True) = True", "(True | P) = True", |
2800 | 65 |
"(P | False) = P", "(False | P) = P", |
66 |
"(P | P) = P", "(P | (P | Q)) = (P | Q)", |
|
3913 | 67 |
"(P | ~P) = True", "(~P | P) = True", |
2082 | 68 |
"((~P) = (~Q)) = (P=Q)", |
9713 | 69 |
"(!x. P) = P", "(? x. P) = P", "? x. x=t", "? x. t=x", |
11232 | 70 |
(* needed for the one-point-rule quantifier simplification procs*) |
71 |
(*essential for termination!!*) |
|
72 |
"(? x. x=t & P(x)) = P(t)", |
|
73 |
"(? x. t=x & P(x)) = P(t)", |
|
74 |
"(! x. x=t --> P(x)) = P(t)", |
|
75 |
"(! x. t=x --> P(x)) = P(t)" ]; |
|
923 | 76 |
|
9875 | 77 |
val imp_cong = standard(impI RSN |
7357 | 78 |
(2, prove_goal (the_context ()) "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" |
9875 | 79 |
(fn _=> [(Blast_tac 1)]) RS mp RS mp)); |
1922 | 80 |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
81 |
(*Miniscoping: pushing in existential quantifiers*) |
7648 | 82 |
val ex_simps = map prover |
3842 | 83 |
["(EX x. P x & Q) = ((EX x. P x) & Q)", |
84 |
"(EX x. P & Q x) = (P & (EX x. Q x))", |
|
85 |
"(EX x. P x | Q) = ((EX x. P x) | Q)", |
|
86 |
"(EX x. P | Q x) = (P | (EX x. Q x))", |
|
87 |
"(EX x. P x --> Q) = ((ALL x. P x) --> Q)", |
|
88 |
"(EX x. P --> Q x) = (P --> (EX x. Q x))"]; |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
89 |
|
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
90 |
(*Miniscoping: pushing in universal quantifiers*) |
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
91 |
val all_simps = map prover |
3842 | 92 |
["(ALL x. P x & Q) = ((ALL x. P x) & Q)", |
93 |
"(ALL x. P & Q x) = (P & (ALL x. Q x))", |
|
94 |
"(ALL x. P x | Q) = ((ALL x. P x) | Q)", |
|
95 |
"(ALL x. P | Q x) = (P | (ALL x. Q x))", |
|
96 |
"(ALL x. P x --> Q) = ((EX x. P x) --> Q)", |
|
97 |
"(ALL x. P --> Q x) = (P --> (ALL x. Q x))"]; |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
98 |
|
923 | 99 |
|
100 |
end; |
|
101 |
||
7648 | 102 |
bind_thms ("ex_simps", ex_simps); |
103 |
bind_thms ("all_simps", all_simps); |
|
7711
4dae7a4fceaf
Rule not_not is now stored in theory (needed by Inductive).
berghofe
parents:
7648
diff
changeset
|
104 |
bind_thm ("not_not", not_not); |
9875 | 105 |
bind_thm ("imp_cong", imp_cong); |
7648 | 106 |
|
3654 | 107 |
(* Elimination of True from asumptions: *) |
108 |
||
11534
0494d0347f18
Proof of True_implies_equals is stored with "open" derivation to
berghofe
parents:
11451
diff
changeset
|
109 |
local fun rd s = read_cterm (sign_of (the_context())) (s, propT); |
0494d0347f18
Proof of True_implies_equals is stored with "open" derivation to
berghofe
parents:
11451
diff
changeset
|
110 |
in val True_implies_equals = standard' (equal_intr |
0494d0347f18
Proof of True_implies_equals is stored with "open" derivation to
berghofe
parents:
11451
diff
changeset
|
111 |
(implies_intr_hyps (implies_elim (assume (rd "True ==> PROP P")) TrueI)) |
0494d0347f18
Proof of True_implies_equals is stored with "open" derivation to
berghofe
parents:
11451
diff
changeset
|
112 |
(implies_intr_hyps (implies_intr (rd "True") (assume (rd "PROP P"))))); |
0494d0347f18
Proof of True_implies_equals is stored with "open" derivation to
berghofe
parents:
11451
diff
changeset
|
113 |
end; |
3654 | 114 |
|
7357 | 115 |
fun prove nm thm = qed_goal nm (the_context ()) thm (fn _ => [(Blast_tac 1)]); |
923 | 116 |
|
9511 | 117 |
prove "eq_commute" "(a=b) = (b=a)"; |
7623 | 118 |
prove "eq_left_commute" "(P=(Q=R)) = (Q=(P=R))"; |
119 |
prove "eq_assoc" "((P=Q)=R) = (P=(Q=R))"; |
|
120 |
val eq_ac = [eq_commute, eq_left_commute, eq_assoc]; |
|
121 |
||
9511 | 122 |
prove "neq_commute" "(a~=b) = (b~=a)"; |
123 |
||
923 | 124 |
prove "conj_commute" "(P&Q) = (Q&P)"; |
125 |
prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))"; |
|
126 |
val conj_comms = [conj_commute, conj_left_commute]; |
|
2134 | 127 |
prove "conj_assoc" "((P&Q)&R) = (P&(Q&R))"; |
923 | 128 |
|
1922 | 129 |
prove "disj_commute" "(P|Q) = (Q|P)"; |
130 |
prove "disj_left_commute" "(P|(Q|R)) = (Q|(P|R))"; |
|
131 |
val disj_comms = [disj_commute, disj_left_commute]; |
|
2134 | 132 |
prove "disj_assoc" "((P|Q)|R) = (P|(Q|R))"; |
1922 | 133 |
|
923 | 134 |
prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)"; |
135 |
prove "conj_disj_distribR" "((P|Q)&R) = (P&R | Q&R)"; |
|
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
136 |
|
1892 | 137 |
prove "disj_conj_distribL" "(P|(Q&R)) = ((P|Q) & (P|R))"; |
138 |
prove "disj_conj_distribR" "((P&Q)|R) = ((P|R) & (Q|R))"; |
|
139 |
||
2134 | 140 |
prove "imp_conjR" "(P --> (Q&R)) = ((P-->Q) & (P-->R))"; |
141 |
prove "imp_conjL" "((P&Q) -->R) = (P --> (Q --> R))"; |
|
142 |
prove "imp_disjL" "((P|Q) --> R) = ((P-->R)&(Q-->R))"; |
|
1892 | 143 |
|
3448 | 144 |
(*These two are specialized, but imp_disj_not1 is useful in Auth/Yahalom.ML*) |
8114 | 145 |
prove "imp_disj_not1" "(P --> Q | R) = (~Q --> P --> R)"; |
146 |
prove "imp_disj_not2" "(P --> Q | R) = (~R --> P --> Q)"; |
|
3448 | 147 |
|
3904 | 148 |
prove "imp_disj1" "((P-->Q)|R) = (P--> Q|R)"; |
149 |
prove "imp_disj2" "(Q|(P-->R)) = (P--> Q|R)"; |
|
150 |
||
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
151 |
prove "de_Morgan_disj" "(~(P | Q)) = (~P & ~Q)"; |
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
152 |
prove "de_Morgan_conj" "(~(P & Q)) = (~P | ~Q)"; |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
153 |
prove "not_imp" "(~(P --> Q)) = (P & ~Q)"; |
1922 | 154 |
prove "not_iff" "(P~=Q) = (P = (~Q))"; |
4743 | 155 |
prove "disj_not1" "(~P | Q) = (P --> Q)"; |
156 |
prove "disj_not2" "(P | ~Q) = (Q --> P)"; (* changes orientation :-( *) |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
157 |
prove "imp_conv_disj" "(P --> Q) = ((~P) | Q)"; |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
158 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
159 |
prove "iff_conv_conj_imp" "(P = Q) = ((P --> Q) & (Q --> P))"; |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
160 |
|
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
161 |
|
9713 | 162 |
(*Avoids duplication of subgoals after split_if, when the true and false |
163 |
cases boil down to the same thing.*) |
|
2134 | 164 |
prove "cases_simp" "((P --> Q) & (~P --> Q)) = Q"; |
165 |
||
3842 | 166 |
prove "not_all" "(~ (! x. P(x))) = (? x.~P(x))"; |
1922 | 167 |
prove "imp_all" "((! x. P x) --> Q) = (? x. P x --> Q)"; |
3842 | 168 |
prove "not_ex" "(~ (? x. P(x))) = (! x.~P(x))"; |
1922 | 169 |
prove "imp_ex" "((? x. P x) --> Q) = (! x. P x --> Q)"; |
1660 | 170 |
|
1655 | 171 |
prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))"; |
172 |
prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; |
|
173 |
||
2134 | 174 |
(* '&' congruence rule: not included by default! |
175 |
May slow rewrite proofs down by as much as 50% *) |
|
176 |
||
9713 | 177 |
let val th = prove_goal (the_context ()) |
2134 | 178 |
"(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" |
7031 | 179 |
(fn _=> [(Blast_tac 1)]) |
2134 | 180 |
in bind_thm("conj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
181 |
||
9713 | 182 |
let val th = prove_goal (the_context ()) |
2134 | 183 |
"(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))" |
7031 | 184 |
(fn _=> [(Blast_tac 1)]) |
2134 | 185 |
in bind_thm("rev_conj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
186 |
||
187 |
(* '|' congruence rule: not included by default! *) |
|
188 |
||
9713 | 189 |
let val th = prove_goal (the_context ()) |
2134 | 190 |
"(P=P')--> (~P'--> (Q=Q'))--> ((P|Q) = (P'|Q'))" |
7031 | 191 |
(fn _=> [(Blast_tac 1)]) |
2134 | 192 |
in bind_thm("disj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
193 |
||
194 |
prove "eq_sym_conv" "(x=y) = (y=x)"; |
|
195 |
||
5278 | 196 |
|
197 |
(** if-then-else rules **) |
|
198 |
||
7031 | 199 |
Goalw [if_def] "(if True then x else y) = x"; |
200 |
by (Blast_tac 1); |
|
201 |
qed "if_True"; |
|
2134 | 202 |
|
7031 | 203 |
Goalw [if_def] "(if False then x else y) = y"; |
204 |
by (Blast_tac 1); |
|
205 |
qed "if_False"; |
|
2134 | 206 |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7031
diff
changeset
|
207 |
Goalw [if_def] "P ==> (if P then x else y) = x"; |
7031 | 208 |
by (Blast_tac 1); |
209 |
qed "if_P"; |
|
5304 | 210 |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7031
diff
changeset
|
211 |
Goalw [if_def] "~P ==> (if P then x else y) = y"; |
7031 | 212 |
by (Blast_tac 1); |
213 |
qed "if_not_P"; |
|
2134 | 214 |
|
7031 | 215 |
Goal "P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))"; |
216 |
by (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1); |
|
217 |
by (stac if_P 2); |
|
218 |
by (stac if_not_P 1); |
|
219 |
by (ALLGOALS (Blast_tac)); |
|
220 |
qed "split_if"; |
|
221 |
||
222 |
Goal "P(if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))"; |
|
223 |
by (stac split_if 1); |
|
224 |
by (Blast_tac 1); |
|
225 |
qed "split_if_asm"; |
|
2134 | 226 |
|
9384
8e8941c491e6
* HOL: removed obsolete expand_if = split_if; theorems if_splits =
wenzelm
parents:
9164
diff
changeset
|
227 |
bind_thms ("if_splits", [split_if, split_if_asm]); |
8e8941c491e6
* HOL: removed obsolete expand_if = split_if; theorems if_splits =
wenzelm
parents:
9164
diff
changeset
|
228 |
|
11003 | 229 |
bind_thm ("if_def2", read_instantiate [("P","\\<lambda>x. x")] split_if); |
230 |
||
7031 | 231 |
Goal "(if c then x else x) = x"; |
232 |
by (stac split_if 1); |
|
233 |
by (Blast_tac 1); |
|
234 |
qed "if_cancel"; |
|
5304 | 235 |
|
7031 | 236 |
Goal "(if x = y then y else x) = x"; |
237 |
by (stac split_if 1); |
|
238 |
by (Blast_tac 1); |
|
239 |
qed "if_eq_cancel"; |
|
5304 | 240 |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
241 |
(*This form is useful for expanding IFs on the RIGHT of the ==> symbol*) |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7031
diff
changeset
|
242 |
Goal "(if P then Q else R) = ((P-->Q) & (~P-->R))"; |
7031 | 243 |
by (rtac split_if 1); |
244 |
qed "if_bool_eq_conj"; |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
245 |
|
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
246 |
(*And this form is useful for expanding IFs on the LEFT*) |
7031 | 247 |
Goal "(if P then Q else R) = ((P&Q) | (~P&R))"; |
248 |
by (stac split_if 1); |
|
249 |
by (Blast_tac 1); |
|
250 |
qed "if_bool_eq_disj"; |
|
2134 | 251 |
|
11232 | 252 |
local |
253 |
val uncurry = prove_goal (the_context()) "P --> Q --> R ==> P & Q --> R" |
|
254 |
(fn prems => [cut_facts_tac prems 1, Blast_tac 1]); |
|
255 |
||
256 |
val iff_allI = allI RS |
|
257 |
prove_goal (the_context()) "!x. P x = Q x ==> (!x. P x) = (!x. Q x)" |
|
258 |
(fn prems => [cut_facts_tac prems 1, Blast_tac 1]) |
|
259 |
in |
|
4351 | 260 |
|
261 |
(*** make simplification procedures for quantifier elimination ***) |
|
262 |
||
9851 | 263 |
structure Quantifier1 = Quantifier1Fun |
264 |
(struct |
|
4351 | 265 |
(*abstract syntax*) |
266 |
fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t) |
|
267 |
| dest_eq _ = None; |
|
268 |
fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t) |
|
269 |
| dest_conj _ = None; |
|
11232 | 270 |
fun dest_imp((c as Const("op -->",_)) $ s $ t) = Some(c,s,t) |
271 |
| dest_imp _ = None; |
|
4351 | 272 |
val conj = HOLogic.conj |
273 |
val imp = HOLogic.imp |
|
274 |
(*rules*) |
|
275 |
val iff_reflection = eq_reflection |
|
276 |
val iffI = iffI |
|
277 |
val conjI= conjI |
|
278 |
val conjE= conjE |
|
279 |
val impI = impI |
|
280 |
val mp = mp |
|
11232 | 281 |
val uncurry = uncurry |
4351 | 282 |
val exI = exI |
283 |
val exE = exE |
|
11232 | 284 |
val iff_allI = iff_allI |
4351 | 285 |
end); |
286 |
||
11232 | 287 |
end; |
288 |
||
4320
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
289 |
local |
11220 | 290 |
val ex_pattern = Thm.read_cterm (Theory.sign_of (the_context ())) |
291 |
("EX x. P(x) & Q(x)",HOLogic.boolT) |
|
292 |
val all_pattern = Thm.read_cterm (Theory.sign_of (the_context ())) |
|
11232 | 293 |
("ALL x. P(x) --> Q(x)",HOLogic.boolT) |
4320
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
294 |
in |
11220 | 295 |
val defEX_regroup = mk_simproc "defined EX" [ex_pattern] |
296 |
Quantifier1.rearrange_ex |
|
297 |
val defALL_regroup = mk_simproc "defined ALL" [all_pattern] |
|
298 |
Quantifier1.rearrange_all |
|
4320
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
299 |
end; |
3913 | 300 |
|
4351 | 301 |
|
302 |
(*** Case splitting ***) |
|
3913 | 303 |
|
5304 | 304 |
structure SplitterData = |
305 |
struct |
|
306 |
structure Simplifier = Simplifier |
|
5552 | 307 |
val mk_eq = mk_eq |
5304 | 308 |
val meta_eq_to_iff = meta_eq_to_obj_eq |
309 |
val iffD = iffD2 |
|
310 |
val disjE = disjE |
|
311 |
val conjE = conjE |
|
312 |
val exE = exE |
|
10231 | 313 |
val contrapos = contrapos_nn |
314 |
val contrapos2 = contrapos_pp |
|
5304 | 315 |
val notnotD = notnotD |
316 |
end; |
|
4681 | 317 |
|
5304 | 318 |
structure Splitter = SplitterFun(SplitterData); |
2263 | 319 |
|
5304 | 320 |
val split_tac = Splitter.split_tac; |
321 |
val split_inside_tac = Splitter.split_inside_tac; |
|
322 |
val split_asm_tac = Splitter.split_asm_tac; |
|
5307 | 323 |
val op addsplits = Splitter.addsplits; |
324 |
val op delsplits = Splitter.delsplits; |
|
5304 | 325 |
val Addsplits = Splitter.Addsplits; |
326 |
val Delsplits = Splitter.Delsplits; |
|
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
327 |
|
2134 | 328 |
(*In general it seems wrong to add distributive laws by default: they |
329 |
might cause exponential blow-up. But imp_disjL has been in for a while |
|
9713 | 330 |
and cannot be removed without affecting existing proofs. Moreover, |
2134 | 331 |
rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the |
332 |
grounds that it allows simplification of R in the two cases.*) |
|
333 |
||
334 |
val mksimps_pairs = |
|
335 |
[("op -->", [mp]), ("op &", [conjunct1,conjunct2]), |
|
336 |
("All", [spec]), ("True", []), ("False", []), |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
337 |
("If", [if_bool_eq_conj RS iffD1])]; |
1758 | 338 |
|
5552 | 339 |
(* ###FIXME: move to Provers/simplifier.ML |
5304 | 340 |
val mk_atomize: (string * thm list) list -> thm -> thm list |
341 |
*) |
|
5552 | 342 |
(* ###FIXME: move to Provers/simplifier.ML *) |
5304 | 343 |
fun mk_atomize pairs = |
344 |
let fun atoms th = |
|
345 |
(case concl_of th of |
|
346 |
Const("Trueprop",_) $ p => |
|
347 |
(case head_of p of |
|
348 |
Const(a,_) => |
|
349 |
(case assoc(pairs,a) of |
|
350 |
Some(rls) => flat (map atoms ([th] RL rls)) |
|
351 |
| None => [th]) |
|
352 |
| _ => [th]) |
|
353 |
| _ => [th]) |
|
354 |
in atoms end; |
|
355 |
||
11624
8a45c7abef04
mksimps and mk_eq_True no longer raise THM exception.
berghofe
parents:
11534
diff
changeset
|
356 |
fun mksimps pairs = |
8a45c7abef04
mksimps and mk_eq_True no longer raise THM exception.
berghofe
parents:
11534
diff
changeset
|
357 |
(mapfilter (try mk_eq) o mk_atomize pairs o forall_elim_vars_safe); |
5304 | 358 |
|
7570 | 359 |
fun unsafe_solver_tac prems = |
360 |
FIRST'[resolve_tac(reflexive_thm::TrueI::refl::prems), atac, etac FalseE]; |
|
361 |
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; |
|
362 |
||
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
363 |
(*No premature instantiation of variables during simplification*) |
7570 | 364 |
fun safe_solver_tac prems = |
365 |
FIRST'[match_tac(reflexive_thm::TrueI::refl::prems), |
|
366 |
eq_assume_tac, ematch_tac [FalseE]]; |
|
367 |
val safe_solver = mk_solver "HOL safe" safe_solver_tac; |
|
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
368 |
|
9713 | 369 |
val HOL_basic_ss = |
370 |
empty_ss setsubgoaler asm_simp_tac |
|
371 |
setSSolver safe_solver |
|
372 |
setSolver unsafe_solver |
|
373 |
setmksimps (mksimps mksimps_pairs) |
|
374 |
setmkeqTrue mk_eq_True |
|
375 |
setmkcong mk_meta_cong; |
|
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
376 |
|
9713 | 377 |
val HOL_ss = |
378 |
HOL_basic_ss addsimps |
|
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
379 |
([triv_forall_equality, (* prunes params *) |
3654 | 380 |
True_implies_equals, (* prune asms `True' *) |
9023 | 381 |
eta_contract_eq, (* prunes eta-expansions *) |
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
382 |
if_True, if_False, if_cancel, if_eq_cancel, |
5304 | 383 |
imp_disjL, conj_assoc, disj_assoc, |
3904 | 384 |
de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp, |
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11434
diff
changeset
|
385 |
disj_not1, not_all, not_ex, cases_simp, |
11434 | 386 |
thm "the_eq_trivial", the_sym_eq_trivial, thm "plus_ac0.zero", thm "plus_ac0_zero_right"] |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
387 |
@ ex_simps @ all_simps @ simp_thms) |
4032
4b1c69d8b767
For each datatype `t' there is now a theorem `split_t_case' of the form
nipkow
parents:
3919
diff
changeset
|
388 |
addsimprocs [defALL_regroup,defEX_regroup] |
4744
4469d498cd48
moved addsplits [expand_if] from HOL_basic_ss to HOL_ss;
wenzelm
parents:
4743
diff
changeset
|
389 |
addcongs [imp_cong] |
4830 | 390 |
addsplits [split_if]; |
2082 | 391 |
|
11034 | 392 |
fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); |
393 |
fun hol_rewrite_cterm rews = |
|
394 |
#2 o Thm.dest_comb o #prop o Thm.crep_thm o Simplifier.full_rewrite (HOL_basic_ss addsimps rews); |
|
395 |
||
396 |
||
6293 | 397 |
(*Simplifies x assuming c and y assuming ~c*) |
398 |
val prems = Goalw [if_def] |
|
399 |
"[| b=c; c ==> x=u; ~c ==> y=v |] ==> \ |
|
400 |
\ (if b then x else y) = (if c then u else v)"; |
|
401 |
by (asm_simp_tac (HOL_ss addsimps prems) 1); |
|
402 |
qed "if_cong"; |
|
403 |
||
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7031
diff
changeset
|
404 |
(*Prevents simplification of x and y: faster and allows the execution |
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7031
diff
changeset
|
405 |
of functional programs. NOW THE DEFAULT.*) |
7031 | 406 |
Goal "b=c ==> (if b then x else y) = (if c then x else y)"; |
407 |
by (etac arg_cong 1); |
|
408 |
qed "if_weak_cong"; |
|
6293 | 409 |
|
410 |
(*Prevents simplification of t: much faster*) |
|
7031 | 411 |
Goal "a = b ==> (let x=a in t(x)) = (let x=b in t(x))"; |
412 |
by (etac arg_cong 1); |
|
413 |
qed "let_weak_cong"; |
|
6293 | 414 |
|
7031 | 415 |
Goal "f(if c then x else y) = (if c then f x else f y)"; |
416 |
by (simp_tac (HOL_ss setloop (split_tac [split_if])) 1); |
|
417 |
qed "if_distrib"; |
|
1655 | 418 |
|
4327 | 419 |
(*For expand_case_tac*) |
7584 | 420 |
val prems = Goal "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)"; |
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
421 |
by (case_tac "P" 1); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
422 |
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems))); |
7584 | 423 |
qed "expand_case"; |
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
424 |
|
4327 | 425 |
(*Used in Auth proofs. Typically P contains Vars that become instantiated |
426 |
during unification.*) |
|
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
427 |
fun expand_case_tac P i = |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
428 |
res_inst_tac [("P",P)] expand_case i THEN |
9713 | 429 |
Simp_tac (i+1) THEN |
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
430 |
Simp_tac i; |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
431 |
|
7584 | 432 |
(*This lemma restricts the effect of the rewrite rule u=v to the left-hand |
433 |
side of an equality. Used in {Integ,Real}/simproc.ML*) |
|
434 |
Goal "x=y ==> (x=z) = (y=z)"; |
|
435 |
by (asm_simp_tac HOL_ss 1); |
|
436 |
qed "restrict_to_left"; |
|
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
437 |
|
7357 | 438 |
(* default simpset *) |
9713 | 439 |
val simpsetup = |
440 |
[fn thy => (simpset_ref_of thy := HOL_ss addcongs [if_weak_cong]; thy)]; |
|
3615
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
441 |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4651
diff
changeset
|
442 |
|
5219 | 443 |
(*** integration of simplifier with classical reasoner ***) |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
444 |
|
5219 | 445 |
structure Clasimp = ClasimpFun |
8473 | 446 |
(structure Simplifier = Simplifier and Splitter = Splitter |
9851 | 447 |
and Classical = Classical and Blast = Blast |
11344 | 448 |
val iffD1 = iffD1 val iffD2 = iffD2 val notE = notE |
9851 | 449 |
val cla_make_elim = cla_make_elim); |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4651
diff
changeset
|
450 |
open Clasimp; |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
451 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
452 |
val HOL_css = (HOL_cs, HOL_ss); |
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
453 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
454 |
|
8641
978db2870862
change_global/local_css move to Provers/clasimp.ML;
wenzelm
parents:
8473
diff
changeset
|
455 |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
456 |
(*** A general refutation procedure ***) |
9713 | 457 |
|
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
458 |
(* Parameters: |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
459 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
460 |
test: term -> bool |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
461 |
tests if a term is at all relevant to the refutation proof; |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
462 |
if not, then it can be discarded. Can improve performance, |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
463 |
esp. if disjunctions can be discarded (no case distinction needed!). |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
464 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
465 |
prep_tac: int -> tactic |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
466 |
A preparation tactic to be applied to the goal once all relevant premises |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
467 |
have been moved to the conclusion. |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
468 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
469 |
ref_tac: int -> tactic |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
470 |
the actual refutation tactic. Should be able to deal with goals |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
471 |
[| A1; ...; An |] ==> False |
9876 | 472 |
where the Ai are atomic, i.e. no top-level &, | or EX |
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
473 |
*) |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
474 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
475 |
fun refute_tac test prep_tac ref_tac = |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
476 |
let val nnf_simps = |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
477 |
[imp_conv_disj,iff_conv_conj_imp,de_Morgan_disj,de_Morgan_conj, |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
478 |
not_all,not_ex,not_not]; |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
479 |
val nnf_simpset = |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
480 |
empty_ss setmkeqTrue mk_eq_True |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
481 |
setmksimps (mksimps mksimps_pairs) |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
482 |
addsimps nnf_simps; |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
483 |
val prem_nnf_tac = full_simp_tac nnf_simpset; |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
484 |
|
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
485 |
val refute_prems_tac = |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
486 |
REPEAT(eresolve_tac [conjE, exE] 1 ORELSE |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
487 |
filter_prems_tac test 1 ORELSE |
6301 | 488 |
etac disjE 1) THEN |
11200
f43fa07536c0
arith_tac now copes with propositional reasoning as well.
nipkow
parents:
11162
diff
changeset
|
489 |
((etac notE 1 THEN eq_assume_tac 1) ORELSE |
f43fa07536c0
arith_tac now copes with propositional reasoning as well.
nipkow
parents:
11162
diff
changeset
|
490 |
ref_tac 1); |
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
491 |
in EVERY'[TRY o filter_prems_tac test, |
6128 | 492 |
DETERM o REPEAT o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac, |
5975
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
493 |
SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] |
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents:
5552
diff
changeset
|
494 |
end; |