author | nipkow |
Thu, 11 Apr 1996 08:30:25 +0200 | |
changeset 1655 | 5be64540f275 |
parent 1548 | afe750876848 |
child 1660 | 8cb42cd97579 |
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 |
||
6 |
Instantiation of the generic simplifier |
|
7 |
*) |
|
8 |
||
9 |
open Simplifier; |
|
10 |
||
11 |
local |
|
12 |
||
13 |
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]); |
|
14 |
||
15 |
val P_imp_P_iff_True = prover "P --> (P = True)" RS mp; |
|
16 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; |
|
17 |
||
18 |
val not_P_imp_P_iff_F = prover "~P --> (P = False)" RS mp; |
|
19 |
val not_P_imp_P_eq_False = not_P_imp_P_iff_F RS eq_reflection; |
|
20 |
||
21 |
fun atomize pairs = |
|
22 |
let fun atoms th = |
|
23 |
(case concl_of th of |
|
24 |
Const("Trueprop",_) $ p => |
|
25 |
(case head_of p of |
|
26 |
Const(a,_) => |
|
27 |
(case assoc(pairs,a) of |
|
28 |
Some(rls) => flat (map atoms ([th] RL rls)) |
|
29 |
| None => [th]) |
|
30 |
| _ => [th]) |
|
31 |
| _ => [th]) |
|
32 |
in atoms end; |
|
33 |
||
34 |
fun mk_meta_eq r = case concl_of r of |
|
1465 | 35 |
Const("==",_)$_$_ => r |
36 |
| _$(Const("op =",_)$_$_) => r RS eq_reflection |
|
37 |
| _$(Const("not",_)$_) => r RS not_P_imp_P_eq_False |
|
923 | 38 |
| _ => r RS P_imp_P_eq_True; |
39 |
(* last 2 lines requires all formulae to be of the from Trueprop(.) *) |
|
40 |
||
41 |
fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; |
|
42 |
||
43 |
val imp_cong = impI RSN |
|
44 |
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" |
|
1465 | 45 |
(fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); |
923 | 46 |
|
47 |
val o_apply = prove_goalw HOL.thy [o_def] "(f o g)(x) = f(g(x))" |
|
48 |
(fn _ => [rtac refl 1]); |
|
49 |
||
50 |
val simp_thms = map prover |
|
51 |
[ "(x=x) = True", |
|
52 |
"(~True) = False", "(~False) = True", "(~ ~ P) = P", |
|
53 |
"(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", |
|
54 |
"(True=P) = P", "(P=True) = P", |
|
55 |
"(True --> P) = P", "(False --> P) = True", |
|
56 |
"(P --> True) = True", "(P --> P) = True", |
|
57 |
"(P --> False) = (~P)", "(P --> ~P) = (~P)", |
|
58 |
"(P & True) = P", "(True & P) = P", |
|
59 |
"(P & False) = False", "(False & P) = False", "(P & P) = P", |
|
60 |
"(P | True) = True", "(True | P) = True", |
|
61 |
"(P | False) = P", "(False | P) = P", "(P | P) = P", |
|
62 |
"(!x.P) = P", "(? x.P) = P", "? x. x=t", "(? x. x=t & P(x)) = P(t)", |
|
63 |
"(P|Q --> R) = ((P-->R)&(Q-->R))" ]; |
|
64 |
||
65 |
in |
|
66 |
||
67 |
val meta_eq_to_obj_eq = prove_goal HOL.thy "x==y ==> x=y" |
|
68 |
(fn [prem] => [rewtac prem, rtac refl 1]); |
|
69 |
||
70 |
val eq_sym_conv = prover "(x=y) = (y=x)"; |
|
71 |
||
72 |
val conj_assoc = prover "((P&Q)&R) = (P&(Q&R))"; |
|
73 |
||
965 | 74 |
val if_True = prove_goalw HOL.thy [if_def] "(if True then x else y) = x" |
923 | 75 |
(fn _=>[fast_tac (HOL_cs addIs [select_equality]) 1]); |
76 |
||
965 | 77 |
val if_False = prove_goalw HOL.thy [if_def] "(if False then x else y) = y" |
923 | 78 |
(fn _=>[fast_tac (HOL_cs addIs [select_equality]) 1]); |
79 |
||
965 | 80 |
val if_P = prove_goal HOL.thy "P ==> (if P then x else y) = x" |
923 | 81 |
(fn [prem] => [ stac (prem RS eqTrueI) 1, rtac if_True 1 ]); |
82 |
||
965 | 83 |
val if_not_P = prove_goal HOL.thy "~P ==> (if P then x else y) = y" |
923 | 84 |
(fn [prem] => [ stac (prem RS not_P_imp_P_iff_F) 1, rtac if_False 1 ]); |
85 |
||
86 |
val expand_if = prove_goal HOL.thy |
|
965 | 87 |
"P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" |
923 | 88 |
(fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1), |
1465 | 89 |
rtac (if_P RS ssubst) 2, |
90 |
rtac (if_not_P RS ssubst) 1, |
|
91 |
REPEAT(fast_tac HOL_cs 1) ]); |
|
923 | 92 |
|
965 | 93 |
val if_bool_eq = prove_goal HOL.thy |
94 |
"(if P then Q else R) = ((P-->Q) & (~P-->R))" |
|
95 |
(fn _ => [rtac expand_if 1]); |
|
923 | 96 |
|
988 | 97 |
(*Add congruence rules for = (instead of ==) *) |
98 |
infix 4 addcongs; |
|
923 | 99 |
fun ss addcongs congs = ss addeqcongs (congs RL [eq_reflection]); |
100 |
||
1264 | 101 |
fun Addcongs congs = (simpset := !simpset addcongs congs); |
102 |
||
988 | 103 |
(*Add a simpset to a classical set!*) |
104 |
infix 4 addss; |
|
105 |
fun cs addss ss = cs addbefore asm_full_simp_tac ss 1; |
|
106 |
||
923 | 107 |
val mksimps_pairs = |
108 |
[("op -->", [mp]), ("op &", [conjunct1,conjunct2]), |
|
109 |
("All", [spec]), ("True", []), ("False", []), |
|
965 | 110 |
("If", [if_bool_eq RS iffD1])]; |
923 | 111 |
|
112 |
fun mksimps pairs = map mk_meta_eq o atomize pairs o gen_all; |
|
113 |
||
114 |
val HOL_ss = empty_ss |
|
115 |
setmksimps (mksimps mksimps_pairs) |
|
116 |
setsolver (fn prems => resolve_tac (TrueI::refl::prems) ORELSE' atac |
|
117 |
ORELSE' etac FalseE) |
|
118 |
setsubgoaler asm_simp_tac |
|
119 |
addsimps ([if_True, if_False, o_apply, conj_assoc] @ simp_thms) |
|
120 |
addcongs [imp_cong]; |
|
121 |
||
941 | 122 |
local val mktac = mk_case_split_tac (meta_eq_to_obj_eq RS iffD2) |
123 |
in |
|
124 |
fun split_tac splits = mktac (map mk_meta_eq splits) |
|
125 |
end; |
|
126 |
||
923 | 127 |
|
128 |
(* eliminiation of existential quantifiers in assumptions *) |
|
129 |
||
130 |
val ex_all_equiv = |
|
131 |
let val lemma1 = prove_goal HOL.thy |
|
132 |
"(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)" |
|
133 |
(fn prems => [resolve_tac prems 1, etac exI 1]); |
|
134 |
val lemma2 = prove_goalw HOL.thy [Ex_def] |
|
135 |
"(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)" |
|
136 |
(fn prems => [REPEAT(resolve_tac prems 1)]) |
|
137 |
in equal_intr lemma1 lemma2 end; |
|
138 |
||
139 |
(* '&' congruence rule: not included by default! |
|
140 |
May slow rewrite proofs down by as much as 50% *) |
|
141 |
||
142 |
val conj_cong = impI RSN |
|
143 |
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" |
|
1465 | 144 |
(fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); |
923 | 145 |
|
1548 | 146 |
val rev_conj_cong = impI RSN |
147 |
(2, prove_goal HOL.thy "(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))" |
|
148 |
(fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); |
|
149 |
||
923 | 150 |
(** 'if' congruence rules: neither included by default! *) |
151 |
||
152 |
(*Simplifies x assuming c and y assuming ~c*) |
|
153 |
val if_cong = prove_goal HOL.thy |
|
965 | 154 |
"[| b=c; c ==> x=u; ~c ==> y=v |] ==>\ |
155 |
\ (if b then x else y) = (if c then u else v)" |
|
923 | 156 |
(fn rew::prems => |
157 |
[stac rew 1, stac expand_if 1, stac expand_if 1, |
|
158 |
fast_tac (HOL_cs addDs prems) 1]); |
|
159 |
||
160 |
(*Prevents simplification of x and y: much faster*) |
|
161 |
val if_weak_cong = prove_goal HOL.thy |
|
965 | 162 |
"b=c ==> (if b then x else y) = (if c then x else y)" |
923 | 163 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
164 |
||
165 |
(*Prevents simplification of t: much faster*) |
|
166 |
val let_weak_cong = prove_goal HOL.thy |
|
167 |
"a = b ==> (let x=a in t(x)) = (let x=b in t(x))" |
|
168 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
169 |
||
170 |
end; |
|
171 |
||
172 |
fun prove nm thm = qed_goal nm HOL.thy thm (fn _ => [fast_tac HOL_cs 1]); |
|
173 |
||
174 |
prove "conj_commute" "(P&Q) = (Q&P)"; |
|
175 |
prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))"; |
|
176 |
val conj_comms = [conj_commute, conj_left_commute]; |
|
177 |
||
178 |
prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)"; |
|
179 |
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
|
180 |
|
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
181 |
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
|
182 |
prove "de_Morgan_conj" "(~(P & Q)) = (~P | ~Q)"; |
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
183 |
|
1655 | 184 |
prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))"; |
185 |
prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; |
|
186 |
||
187 |
qed_goal "if_cancel" HOL.thy "(if c then x else x) = x" |
|
188 |
(fn _ => [simp_tac (HOL_ss setloop (split_tac [expand_if])) 1]); |
|
189 |
||
190 |
qed_goal "if_distrib" HOL.thy |
|
191 |
"f(if c then x else y) = (if c then f x else f y)" |
|
192 |
(fn _ => [simp_tac (HOL_ss setloop (split_tac [expand_if])) 1]); |
|
193 |
||
194 |
qed_goalw "o_assoc" HOL.thy [o_def] "(f o g) o h = (f o g o h)" |
|
195 |
(fn _=>[rtac ext 1, rtac refl 1]); |