author | wenzelm |
Thu, 30 Jul 1998 19:02:52 +0200 | |
changeset 5219 | 924359415f09 |
parent 5190 | 4ae031622592 |
child 5220 | 07f80f447b27 |
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 |
||
4794 | 6 |
Instantiation of the generic simplifier. |
923 | 7 |
*) |
8 |
||
1984 | 9 |
section "Simplifier"; |
10 |
||
11 |
(*** Addition of rules to simpsets and clasets simultaneously ***) |
|
12 |
||
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
13 |
infix 4 addIffs delIffs; |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
14 |
|
1984 | 15 |
(*Takes UNCONDITIONAL theorems of the form A<->B to |
2031 | 16 |
the Safe Intr rule B==>A and |
17 |
the Safe Destruct rule A==>B. |
|
1984 | 18 |
Also ~A goes to the Safe Elim rule A ==> ?R |
19 |
Failing other cases, A is added as a Safe Intr rule*) |
|
20 |
local |
|
21 |
val iff_const = HOLogic.eq_const HOLogic.boolT; |
|
22 |
||
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
23 |
fun addIff ((cla, simp), th) = |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
24 |
(case HOLogic.dest_Trueprop (#prop (rep_thm th)) of |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
25 |
(Const("Not", _) $ A) => |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
26 |
cla addSEs [zero_var_indexes (th RS notE)] |
2031 | 27 |
| (con $ _ $ _) => |
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
28 |
if con = iff_const |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
29 |
then cla addSIs [zero_var_indexes (th RS iffD2)] |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
30 |
addSDs [zero_var_indexes (th RS iffD1)] |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
31 |
else cla addSIs [th] |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
32 |
| _ => cla addSIs [th], |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
33 |
simp addsimps [th]) |
1984 | 34 |
handle _ => error ("AddIffs: theorem must be unconditional\n" ^ |
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
35 |
string_of_thm th); |
1984 | 36 |
|
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
37 |
fun delIff ((cla, simp), th) = |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
38 |
(case HOLogic.dest_Trueprop (#prop (rep_thm th)) of |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
39 |
(Const ("Not", _) $ A) => |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
40 |
cla delrules [zero_var_indexes (th RS notE)] |
2031 | 41 |
| (con $ _ $ _) => |
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
42 |
if con = iff_const |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
43 |
then cla delrules [zero_var_indexes (th RS iffD2), |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
44 |
make_elim (zero_var_indexes (th RS iffD1))] |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
45 |
else cla delrules [th] |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
46 |
| _ => cla delrules [th], |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
47 |
simp delsimps [th]) |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
48 |
handle _ => (warning("DelIffs: ignoring conditional theorem\n" ^ |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
49 |
string_of_thm th); (cla, simp)); |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
50 |
|
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
51 |
fun store_clasimp (cla, simp) = (claset_ref () := cla; simpset_ref () := simp) |
1984 | 52 |
in |
5190
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
53 |
val op addIffs = foldl addIff; |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
54 |
val op delIffs = foldl delIff; |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
55 |
fun AddIffs thms = store_clasimp ((claset (), simpset ()) addIffs thms); |
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
berghofe
parents:
4930
diff
changeset
|
56 |
fun DelIffs thms = store_clasimp ((claset (), simpset ()) delIffs thms); |
1984 | 57 |
end; |
58 |
||
4640 | 59 |
qed_goal "meta_eq_to_obj_eq" HOL.thy "x==y ==> x=y" |
60 |
(fn [prem] => [rewtac prem, rtac refl 1]); |
|
61 |
||
923 | 62 |
local |
63 |
||
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
64 |
fun prover s = prove_goal HOL.thy s (K [Blast_tac 1]); |
923 | 65 |
|
1922 | 66 |
val P_imp_P_iff_True = prover "P --> (P = True)" RS mp; |
67 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; |
|
923 | 68 |
|
1922 | 69 |
val not_P_imp_P_iff_F = prover "~P --> (P = False)" RS mp; |
70 |
val not_P_imp_P_eq_False = not_P_imp_P_iff_F RS eq_reflection; |
|
923 | 71 |
|
1922 | 72 |
fun atomize pairs = |
73 |
let fun atoms th = |
|
2031 | 74 |
(case concl_of th of |
75 |
Const("Trueprop",_) $ p => |
|
76 |
(case head_of p of |
|
77 |
Const(a,_) => |
|
78 |
(case assoc(pairs,a) of |
|
79 |
Some(rls) => flat (map atoms ([th] RL rls)) |
|
80 |
| None => [th]) |
|
81 |
| _ => [th]) |
|
82 |
| _ => [th]) |
|
1922 | 83 |
in atoms end; |
923 | 84 |
|
2134 | 85 |
fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; |
86 |
||
87 |
in |
|
88 |
||
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
89 |
fun mk_meta_eq r = r RS eq_reflection; |
4677 | 90 |
fun mk_meta_eq_True r = Some(r RS meta_eq_to_obj_eq RS P_imp_P_eq_True); |
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
91 |
|
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
92 |
fun mk_meta_eq_simp r = case concl_of r of |
2031 | 93 |
Const("==",_)$_$_ => r |
4677 | 94 |
| _$(Const("op =",_)$lhs$rhs) => mk_meta_eq r |
2718 | 95 |
| _$(Const("Not",_)$_) => r RS not_P_imp_P_eq_False |
1922 | 96 |
| _ => r RS P_imp_P_eq_True; |
97 |
(* last 2 lines requires all formulae to be of the from Trueprop(.) *) |
|
923 | 98 |
|
2082 | 99 |
val simp_thms = map prover |
100 |
[ "(x=x) = True", |
|
101 |
"(~True) = False", "(~False) = True", "(~ ~ P) = P", |
|
102 |
"(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", |
|
4640 | 103 |
"(True=P) = P", "(P=True) = P", "(False=P) = (~P)", "(P=False) = (~P)", |
2082 | 104 |
"(True --> P) = P", "(False --> P) = True", |
105 |
"(P --> True) = True", "(P --> P) = True", |
|
106 |
"(P --> False) = (~P)", "(P --> ~P) = (~P)", |
|
107 |
"(P & True) = P", "(True & P) = P", |
|
2800 | 108 |
"(P & False) = False", "(False & P) = False", |
109 |
"(P & P) = P", "(P & (P & Q)) = (P & Q)", |
|
3913 | 110 |
"(P & ~P) = False", "(~P & P) = False", |
2082 | 111 |
"(P | True) = True", "(True | P) = True", |
2800 | 112 |
"(P | False) = P", "(False | P) = P", |
113 |
"(P | P) = P", "(P | (P | Q)) = (P | Q)", |
|
3913 | 114 |
"(P | ~P) = True", "(~P | P) = True", |
2082 | 115 |
"((~P) = (~Q)) = (P=Q)", |
3842 | 116 |
"(!x. P) = P", "(? x. P) = P", "? x. x=t", "? x. t=x", |
4351 | 117 |
(*two needed for the one-point-rule quantifier simplification procs*) |
118 |
"(? x. x=t & P(x)) = P(t)", (*essential for termination!!*) |
|
119 |
"(! x. t=x --> P(x)) = P(t)" ]; (*covers a stray case*) |
|
923 | 120 |
|
988 | 121 |
(*Add congruence rules for = (instead of ==) *) |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
122 |
infix 4 addcongs delcongs; |
4351 | 123 |
|
4640 | 124 |
fun mk_meta_cong rl = |
125 |
standard(mk_meta_eq(replicate (nprems_of rl) meta_eq_to_obj_eq MRS rl)) |
|
126 |
handle THM _ => |
|
127 |
error("Premises and conclusion of congruence rules must be =-equalities"); |
|
128 |
||
129 |
fun ss addcongs congs = ss addeqcongs (map mk_meta_cong congs); |
|
130 |
||
131 |
fun ss delcongs congs = ss deleqcongs (map mk_meta_cong congs); |
|
923 | 132 |
|
4086 | 133 |
fun Addcongs congs = (simpset_ref() := simpset() addcongs congs); |
134 |
fun Delcongs congs = (simpset_ref() := simpset() delcongs congs); |
|
1264 | 135 |
|
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
136 |
fun mksimps pairs = map mk_meta_eq_simp o atomize pairs o gen_all; |
923 | 137 |
|
1922 | 138 |
val imp_cong = impI RSN |
139 |
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
140 |
(fn _=> [Blast_tac 1]) RS mp RS mp); |
1922 | 141 |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
142 |
(*Miniscoping: pushing in existential quantifiers*) |
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
143 |
val ex_simps = map prover |
3842 | 144 |
["(EX x. P x & Q) = ((EX x. P x) & Q)", |
145 |
"(EX x. P & Q x) = (P & (EX x. Q x))", |
|
146 |
"(EX x. P x | Q) = ((EX x. P x) | Q)", |
|
147 |
"(EX x. P | Q x) = (P | (EX x. Q x))", |
|
148 |
"(EX x. P x --> Q) = ((ALL x. P x) --> Q)", |
|
149 |
"(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
|
150 |
|
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
151 |
(*Miniscoping: pushing in universal quantifiers*) |
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
152 |
val all_simps = map prover |
3842 | 153 |
["(ALL x. P x & Q) = ((ALL x. P x) & Q)", |
154 |
"(ALL x. P & Q x) = (P & (ALL x. Q x))", |
|
155 |
"(ALL x. P x | Q) = ((ALL x. P x) | Q)", |
|
156 |
"(ALL x. P | Q x) = (P | (ALL x. Q x))", |
|
157 |
"(ALL x. P x --> Q) = ((EX x. P x) --> Q)", |
|
158 |
"(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
|
159 |
|
923 | 160 |
|
2022 | 161 |
(* elimination of existential quantifiers in assumptions *) |
923 | 162 |
|
163 |
val ex_all_equiv = |
|
164 |
let val lemma1 = prove_goal HOL.thy |
|
165 |
"(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)" |
|
166 |
(fn prems => [resolve_tac prems 1, etac exI 1]); |
|
167 |
val lemma2 = prove_goalw HOL.thy [Ex_def] |
|
168 |
"(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)" |
|
169 |
(fn prems => [REPEAT(resolve_tac prems 1)]) |
|
170 |
in equal_intr lemma1 lemma2 end; |
|
171 |
||
172 |
end; |
|
173 |
||
3654 | 174 |
(* Elimination of True from asumptions: *) |
175 |
||
176 |
val True_implies_equals = prove_goal HOL.thy |
|
177 |
"(True ==> PROP P) == PROP P" |
|
4525 | 178 |
(K [rtac equal_intr_rule 1, atac 2, |
3654 | 179 |
METAHYPS (fn prems => resolve_tac prems 1) 1, |
180 |
rtac TrueI 1]); |
|
181 |
||
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
182 |
fun prove nm thm = qed_goal nm HOL.thy thm (K [Blast_tac 1]); |
923 | 183 |
|
184 |
prove "conj_commute" "(P&Q) = (Q&P)"; |
|
185 |
prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))"; |
|
186 |
val conj_comms = [conj_commute, conj_left_commute]; |
|
2134 | 187 |
prove "conj_assoc" "((P&Q)&R) = (P&(Q&R))"; |
923 | 188 |
|
1922 | 189 |
prove "disj_commute" "(P|Q) = (Q|P)"; |
190 |
prove "disj_left_commute" "(P|(Q|R)) = (Q|(P|R))"; |
|
191 |
val disj_comms = [disj_commute, disj_left_commute]; |
|
2134 | 192 |
prove "disj_assoc" "((P|Q)|R) = (P|(Q|R))"; |
1922 | 193 |
|
923 | 194 |
prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)"; |
195 |
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
|
196 |
|
1892 | 197 |
prove "disj_conj_distribL" "(P|(Q&R)) = ((P|Q) & (P|R))"; |
198 |
prove "disj_conj_distribR" "((P&Q)|R) = ((P|R) & (Q|R))"; |
|
199 |
||
2134 | 200 |
prove "imp_conjR" "(P --> (Q&R)) = ((P-->Q) & (P-->R))"; |
201 |
prove "imp_conjL" "((P&Q) -->R) = (P --> (Q --> R))"; |
|
202 |
prove "imp_disjL" "((P|Q) --> R) = ((P-->R)&(Q-->R))"; |
|
1892 | 203 |
|
3448 | 204 |
(*These two are specialized, but imp_disj_not1 is useful in Auth/Yahalom.ML*) |
205 |
prove "imp_disj_not1" "((P --> Q | R)) = (~Q --> P --> R)"; |
|
206 |
prove "imp_disj_not2" "((P --> Q | R)) = (~R --> P --> Q)"; |
|
207 |
||
3904 | 208 |
prove "imp_disj1" "((P-->Q)|R) = (P--> Q|R)"; |
209 |
prove "imp_disj2" "(Q|(P-->R)) = (P--> Q|R)"; |
|
210 |
||
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
211 |
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
|
212 |
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
|
213 |
prove "not_imp" "(~(P --> Q)) = (P & ~Q)"; |
1922 | 214 |
prove "not_iff" "(P~=Q) = (P = (~Q))"; |
4743 | 215 |
prove "disj_not1" "(~P | Q) = (P --> Q)"; |
216 |
prove "disj_not2" "(P | ~Q) = (Q --> P)"; (* changes orientation :-( *) |
|
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
217 |
|
4830 | 218 |
(*Avoids duplication of subgoals after split_if, when the true and false |
2134 | 219 |
cases boil down to the same thing.*) |
220 |
prove "cases_simp" "((P --> Q) & (~P --> Q)) = Q"; |
|
221 |
||
3842 | 222 |
prove "not_all" "(~ (! x. P(x))) = (? x.~P(x))"; |
1922 | 223 |
prove "imp_all" "((! x. P x) --> Q) = (? x. P x --> Q)"; |
3842 | 224 |
prove "not_ex" "(~ (? x. P(x))) = (! x.~P(x))"; |
1922 | 225 |
prove "imp_ex" "((? x. P x) --> Q) = (! x. P x --> Q)"; |
1660 | 226 |
|
1655 | 227 |
prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))"; |
228 |
prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; |
|
229 |
||
2134 | 230 |
(* '&' congruence rule: not included by default! |
231 |
May slow rewrite proofs down by as much as 50% *) |
|
232 |
||
233 |
let val th = prove_goal HOL.thy |
|
234 |
"(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
235 |
(fn _=> [Blast_tac 1]) |
2134 | 236 |
in bind_thm("conj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
237 |
||
238 |
let val th = prove_goal HOL.thy |
|
239 |
"(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))" |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
240 |
(fn _=> [Blast_tac 1]) |
2134 | 241 |
in bind_thm("rev_conj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
242 |
||
243 |
(* '|' congruence rule: not included by default! *) |
|
244 |
||
245 |
let val th = prove_goal HOL.thy |
|
246 |
"(P=P')--> (~P'--> (Q=Q'))--> ((P|Q) = (P'|Q'))" |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
247 |
(fn _=> [Blast_tac 1]) |
2134 | 248 |
in bind_thm("disj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
249 |
||
250 |
prove "eq_sym_conv" "(x=y) = (y=x)"; |
|
251 |
||
252 |
qed_goalw "o_apply" HOL.thy [o_def] "(f o g) x = f (g x)" |
|
4525 | 253 |
(K [rtac refl 1]); |
2134 | 254 |
|
255 |
qed_goalw "if_True" HOL.thy [if_def] "(if True then x else y) = x" |
|
4525 | 256 |
(K [Blast_tac 1]); |
2134 | 257 |
|
258 |
qed_goalw "if_False" HOL.thy [if_def] "(if False then x else y) = y" |
|
4525 | 259 |
(K [Blast_tac 1]); |
2134 | 260 |
|
261 |
qed_goal "if_P" HOL.thy "P ==> (if P then x else y) = x" |
|
262 |
(fn [prem] => [ stac (prem RS eqTrueI) 1, rtac if_True 1 ]); |
|
263 |
(* |
|
264 |
qed_goal "if_not_P" HOL.thy "~P ==> (if P then x else y) = y" |
|
265 |
(fn [prem] => [ stac (prem RS not_P_imp_P_iff_F) 1, rtac if_False 1 ]); |
|
266 |
*) |
|
267 |
qed_goalw "if_not_P" HOL.thy [if_def] "!!P. ~P ==> (if P then x else y) = y" |
|
4525 | 268 |
(K [Blast_tac 1]); |
2134 | 269 |
|
4830 | 270 |
qed_goal "split_if" HOL.thy |
4205
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
oheimb
parents:
4189
diff
changeset
|
271 |
"P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" (K [ |
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
oheimb
parents:
4189
diff
changeset
|
272 |
res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1, |
2134 | 273 |
stac if_P 2, |
274 |
stac if_not_P 1, |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
275 |
ALLGOALS (Blast_tac)]); |
4830 | 276 |
(* for backwards compatibility: *) |
277 |
val expand_if = split_if; |
|
4205
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
oheimb
parents:
4189
diff
changeset
|
278 |
|
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
oheimb
parents:
4189
diff
changeset
|
279 |
qed_goal "split_if_asm" HOL.thy |
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
280 |
"P(if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" |
4830 | 281 |
(K [stac split_if 1, |
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
282 |
Blast_tac 1]); |
2134 | 283 |
|
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
284 |
(*This form is useful for expanding IFs on the RIGHT of the ==> symbol*) |
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
285 |
qed_goal "if_bool_eq_conj" HOL.thy |
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
286 |
"(if P then Q else R) = ((P-->Q) & (~P-->R))" |
4830 | 287 |
(K [rtac split_if 1]); |
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
288 |
|
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
289 |
(*And this form is useful for expanding IFs on the LEFT*) |
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
290 |
qed_goal "if_bool_eq_disj" HOL.thy |
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
291 |
"(if P then Q else R) = ((P&Q) | (~P&R))" |
4830 | 292 |
(K [stac split_if 1, |
4769
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
paulson
parents:
4744
diff
changeset
|
293 |
Blast_tac 1]); |
2134 | 294 |
|
4351 | 295 |
|
296 |
(*** make simplification procedures for quantifier elimination ***) |
|
297 |
||
298 |
structure Quantifier1 = Quantifier1Fun( |
|
299 |
struct |
|
300 |
(*abstract syntax*) |
|
301 |
fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t) |
|
302 |
| dest_eq _ = None; |
|
303 |
fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t) |
|
304 |
| dest_conj _ = None; |
|
305 |
val conj = HOLogic.conj |
|
306 |
val imp = HOLogic.imp |
|
307 |
(*rules*) |
|
308 |
val iff_reflection = eq_reflection |
|
309 |
val iffI = iffI |
|
310 |
val sym = sym |
|
311 |
val conjI= conjI |
|
312 |
val conjE= conjE |
|
313 |
val impI = impI |
|
314 |
val impE = impE |
|
315 |
val mp = mp |
|
316 |
val exI = exI |
|
317 |
val exE = exE |
|
318 |
val allI = allI |
|
319 |
val allE = allE |
|
320 |
end); |
|
321 |
||
4320
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
322 |
local |
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
323 |
val ex_pattern = |
4351 | 324 |
read_cterm (sign_of HOL.thy) ("EX x. P(x) & Q(x)",HOLogic.boolT) |
3913 | 325 |
|
4320
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
326 |
val all_pattern = |
4351 | 327 |
read_cterm (sign_of HOL.thy) ("ALL x. P(x) & P'(x) --> Q(x)",HOLogic.boolT) |
4320
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
328 |
|
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
329 |
in |
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
330 |
val defEX_regroup = |
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
331 |
mk_simproc "defined EX" [ex_pattern] Quantifier1.rearrange_ex; |
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
332 |
val defALL_regroup = |
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
333 |
mk_simproc "defined ALL" [all_pattern] Quantifier1.rearrange_all; |
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
nipkow
parents:
4205
diff
changeset
|
334 |
end; |
3913 | 335 |
|
4351 | 336 |
|
337 |
(*** Case splitting ***) |
|
3913 | 338 |
|
2263 | 339 |
local val mktac = mk_case_split_tac (meta_eq_to_obj_eq RS iffD2) |
340 |
in |
|
341 |
fun split_tac splits = mktac (map mk_meta_eq splits) |
|
342 |
end; |
|
343 |
||
344 |
local val mktac = mk_case_split_inside_tac (meta_eq_to_obj_eq RS iffD2) |
|
345 |
in |
|
346 |
fun split_inside_tac splits = mktac (map mk_meta_eq splits) |
|
347 |
end; |
|
348 |
||
4205
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
oheimb
parents:
4189
diff
changeset
|
349 |
val split_asm_tac = mk_case_split_asm_tac split_tac |
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
oheimb
parents:
4189
diff
changeset
|
350 |
(disjE,conjE,exE,contrapos,contrapos2,notnotD); |
4189 | 351 |
|
4681 | 352 |
infix 4 addsplits delsplits; |
353 |
||
4669 | 354 |
fun ss addsplits splits = |
4930
89271bc4e7ed
extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents:
4830
diff
changeset
|
355 |
let fun addsplit (ss,split) = |
89271bc4e7ed
extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents:
4830
diff
changeset
|
356 |
let val (name,asm) = split_thm_info split |
89271bc4e7ed
extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents:
4830
diff
changeset
|
357 |
in ss addloop ("split "^ name ^ (if asm then " asm" else ""), |
89271bc4e7ed
extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents:
4830
diff
changeset
|
358 |
(if asm then split_asm_tac else split_tac) [split]) end |
4669 | 359 |
in foldl addsplit (ss,splits) end; |
2263 | 360 |
|
4681 | 361 |
fun ss delsplits splits = |
362 |
let fun delsplit(ss,split) = |
|
4930
89271bc4e7ed
extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents:
4830
diff
changeset
|
363 |
let val (name,asm) = split_thm_info split |
89271bc4e7ed
extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents:
4830
diff
changeset
|
364 |
in ss delloop ("split "^ name ^ (if asm then " asm" else "")) end |
4681 | 365 |
in foldl delsplit (ss,splits) end; |
366 |
||
367 |
fun Addsplits splits = (simpset_ref() := simpset() addsplits splits); |
|
368 |
fun Delsplits splits = (simpset_ref() := simpset() delsplits splits); |
|
369 |
||
2251 | 370 |
qed_goal "if_cancel" HOL.thy "(if c then x else x) = x" |
4830 | 371 |
(K [split_tac [split_if] 1, Blast_tac 1]); |
2251 | 372 |
|
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
373 |
qed_goal "if_eq_cancel" HOL.thy "(if x = y then y else x) = x" |
4830 | 374 |
(K [split_tac [split_if] 1, Blast_tac 1]); |
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
375 |
|
2134 | 376 |
(** 'if' congruence rules: neither included by default! *) |
377 |
||
378 |
(*Simplifies x assuming c and y assuming ~c*) |
|
379 |
qed_goal "if_cong" HOL.thy |
|
380 |
"[| b=c; c ==> x=u; ~c ==> y=v |] ==>\ |
|
381 |
\ (if b then x else y) = (if c then u else v)" |
|
382 |
(fn rew::prems => |
|
4830 | 383 |
[stac rew 1, stac split_if 1, stac split_if 1, |
2935 | 384 |
blast_tac (HOL_cs addDs prems) 1]); |
2134 | 385 |
|
386 |
(*Prevents simplification of x and y: much faster*) |
|
387 |
qed_goal "if_weak_cong" HOL.thy |
|
388 |
"b=c ==> (if b then x else y) = (if c then x else y)" |
|
389 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
390 |
||
391 |
(*Prevents simplification of t: much faster*) |
|
392 |
qed_goal "let_weak_cong" HOL.thy |
|
393 |
"a = b ==> (let x=a in t(x)) = (let x=b in t(x))" |
|
394 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
395 |
||
396 |
(*In general it seems wrong to add distributive laws by default: they |
|
397 |
might cause exponential blow-up. But imp_disjL has been in for a while |
|
398 |
and cannot be removed without affecting existing proofs. Moreover, |
|
399 |
rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the |
|
400 |
grounds that it allows simplification of R in the two cases.*) |
|
401 |
||
402 |
val mksimps_pairs = |
|
403 |
[("op -->", [mp]), ("op &", [conjunct1,conjunct2]), |
|
404 |
("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
|
405 |
("If", [if_bool_eq_conj RS iffD1])]; |
1758 | 406 |
|
4640 | 407 |
fun unsafe_solver prems = FIRST'[resolve_tac (reflexive_thm::TrueI::refl::prems), |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
408 |
atac, etac FalseE]; |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
409 |
(*No premature instantiation of variables during simplification*) |
4640 | 410 |
fun safe_solver prems = FIRST'[match_tac (reflexive_thm::TrueI::prems), |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
411 |
eq_assume_tac, ematch_tac [FalseE]]; |
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
412 |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
413 |
val HOL_basic_ss = empty_ss setsubgoaler asm_simp_tac |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
414 |
setSSolver safe_solver |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
415 |
setSolver unsafe_solver |
4677 | 416 |
setmksimps (mksimps mksimps_pairs) |
4744
4469d498cd48
moved addsplits [expand_if] from HOL_basic_ss to HOL_ss;
wenzelm
parents:
4743
diff
changeset
|
417 |
setmkeqTrue mk_meta_eq_True; |
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
418 |
|
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
419 |
val HOL_ss = |
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
420 |
HOL_basic_ss addsimps |
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
421 |
([triv_forall_equality, (* prunes params *) |
3654 | 422 |
True_implies_equals, (* prune asms `True' *) |
4718
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
oheimb
parents:
4681
diff
changeset
|
423 |
if_True, if_False, if_cancel, if_eq_cancel, |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
424 |
o_apply, imp_disjL, conj_assoc, disj_assoc, |
3904 | 425 |
de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp, |
4743 | 426 |
disj_not1, not_all, not_ex, cases_simp] |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
427 |
@ 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
|
428 |
addsimprocs [defALL_regroup,defEX_regroup] |
4744
4469d498cd48
moved addsplits [expand_if] from HOL_basic_ss to HOL_ss;
wenzelm
parents:
4743
diff
changeset
|
429 |
addcongs [imp_cong] |
4830 | 430 |
addsplits [split_if]; |
2082 | 431 |
|
1655 | 432 |
qed_goal "if_distrib" HOL.thy |
433 |
"f(if c then x else y) = (if c then f x else f y)" |
|
4830 | 434 |
(K [simp_tac (HOL_ss setloop (split_tac [split_if])) 1]); |
1655 | 435 |
|
2097 | 436 |
qed_goalw "o_assoc" HOL.thy [o_def] "f o (g o h) = f o g o h" |
4525 | 437 |
(K [rtac ext 1, rtac refl 1]); |
1984 | 438 |
|
439 |
||
4327 | 440 |
(*For expand_case_tac*) |
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
441 |
val prems = goal HOL.thy "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)"; |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
442 |
by (case_tac "P" 1); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
443 |
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems))); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
444 |
val expand_case = result(); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
445 |
|
4327 | 446 |
(*Used in Auth proofs. Typically P contains Vars that become instantiated |
447 |
during unification.*) |
|
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
448 |
fun expand_case_tac P i = |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
449 |
res_inst_tac [("P",P)] expand_case i THEN |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
450 |
Simp_tac (i+1) THEN |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
451 |
Simp_tac i; |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
452 |
|
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
453 |
|
4119 | 454 |
(* install implicit simpset *) |
1984 | 455 |
|
4086 | 456 |
simpset_ref() := HOL_ss; |
1984 | 457 |
|
3615
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
458 |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4651
diff
changeset
|
459 |
|
5219 | 460 |
(*** integration of simplifier with classical reasoner ***) |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
461 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
462 |
(* rot_eq_tac rotates the first equality premise of subgoal i to the front, |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
463 |
fails if there is no equaliy or if an equality is already at the front *) |
3538 | 464 |
local |
465 |
fun is_eq (Const ("Trueprop", _) $ (Const("op =" ,_) $ _ $ _)) = true |
|
466 |
| is_eq _ = false; |
|
4188
1025a27b08f9
changed libraray function find to find_index_eq, currying it
oheimb
parents:
4119
diff
changeset
|
467 |
val find_eq = find_index is_eq; |
3538 | 468 |
in |
469 |
val rot_eq_tac = |
|
4188
1025a27b08f9
changed libraray function find to find_index_eq, currying it
oheimb
parents:
4119
diff
changeset
|
470 |
SUBGOAL (fn (Bi,i) => let val n = find_eq (Logic.strip_assums_hyp Bi) in |
1025a27b08f9
changed libraray function find to find_index_eq, currying it
oheimb
parents:
4119
diff
changeset
|
471 |
if n>0 then rotate_tac n i else no_tac end) |
3538 | 472 |
end; |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
473 |
|
5219 | 474 |
|
475 |
structure Clasimp = ClasimpFun |
|
476 |
(structure Simplifier = Simplifier and Classical = Classical and Blast = Blast |
|
477 |
val addcongs = op addcongs and delcongs = op delcongs |
|
478 |
and addSaltern = op addSaltern and addbefore = op addbefore); |
|
479 |
||
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4651
diff
changeset
|
480 |
open Clasimp; |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
481 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
482 |
val HOL_css = (HOL_cs, HOL_ss); |