author | oheimb |
Wed, 12 Aug 1998 17:40:18 +0200 | |
changeset 5307 | 6a699d5cdef4 |
parent 5306 | 3d1368a3a2a2 |
child 5316 | 7a8975451a89 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/Fun |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
Lemmas about functions. |
|
7 |
*) |
|
8 |
||
4656 | 9 |
|
5069 | 10 |
Goal "(f = g) = (!x. f(x)=g(x))"; |
923 | 11 |
by (rtac iffI 1); |
1264 | 12 |
by (Asm_simp_tac 1); |
13 |
by (rtac ext 1 THEN Asm_simp_tac 1); |
|
923 | 14 |
qed "expand_fun_eq"; |
15 |
||
4656 | 16 |
val prems = goal thy |
923 | 17 |
"[| f(x)=u; !!x. P(x) ==> g(f(x)) = x; P(x) |] ==> x=g(u)"; |
18 |
by (rtac (arg_cong RS box_equals) 1); |
|
19 |
by (REPEAT (resolve_tac (prems@[refl]) 1)); |
|
20 |
qed "apply_inverse"; |
|
21 |
||
22 |
||
4656 | 23 |
(** "Axiom" of Choice, proved using the description operator **) |
24 |
||
25 |
goal HOL.thy "!!Q. ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"; |
|
26 |
by (fast_tac (claset() addEs [selectI]) 1); |
|
27 |
qed "choice"; |
|
28 |
||
29 |
goal Set.thy "!!S. ALL x:S. EX y. Q x y ==> EX f. ALL x:S. Q x (f x)"; |
|
30 |
by (fast_tac (claset() addEs [selectI]) 1); |
|
31 |
qed "bchoice"; |
|
32 |
||
33 |
||
5306 | 34 |
section "o"; |
35 |
||
36 |
qed_goalw "o_apply" thy [o_def] "(f o g) x = f (g x)" |
|
37 |
(K [rtac refl 1]); |
|
38 |
Addsimps [o_apply]; |
|
39 |
||
40 |
qed_goalw "o_assoc" thy [o_def] "f o (g o h) = f o g o h" |
|
41 |
(K [rtac ext 1, rtac refl 1]); |
|
42 |
||
43 |
qed_goalw "Id_o" thy [Id_def] "Id o g = g" |
|
44 |
(K [rtac ext 1, Simp_tac 1]); |
|
45 |
Addsimps [Id_o]; |
|
46 |
||
47 |
qed_goalw "o_Id" thy [Id_def] "f o Id = f" |
|
48 |
(K [rtac ext 1, Simp_tac 1]); |
|
49 |
Addsimps [o_Id]; |
|
50 |
||
51 |
Goalw [o_def] "(f o g)``r = f``(g``r)"; |
|
52 |
by (Blast_tac 1); |
|
53 |
qed "image_compose"; |
|
54 |
||
55 |
||
56 |
section "inj"; |
|
57 |
||
923 | 58 |
(*** inj(f): f is a one-to-one function ***) |
59 |
||
4656 | 60 |
val prems = goalw thy [inj_def] |
923 | 61 |
"[| !! x y. f(x) = f(y) ==> x=y |] ==> inj(f)"; |
4089 | 62 |
by (blast_tac (claset() addIs prems) 1); |
923 | 63 |
qed "injI"; |
64 |
||
4656 | 65 |
val [major] = goal thy "(!!x. g(f(x)) = x) ==> inj(f)"; |
923 | 66 |
by (rtac injI 1); |
67 |
by (etac (arg_cong RS box_equals) 1); |
|
68 |
by (rtac major 1); |
|
69 |
by (rtac major 1); |
|
70 |
qed "inj_inverseI"; |
|
71 |
||
4656 | 72 |
val [major,minor] = goalw thy [inj_def] |
923 | 73 |
"[| inj(f); f(x) = f(y) |] ==> x=y"; |
74 |
by (rtac (major RS spec RS spec RS mp) 1); |
|
75 |
by (rtac minor 1); |
|
76 |
qed "injD"; |
|
77 |
||
78 |
(*Useful with the simplifier*) |
|
4656 | 79 |
val [major] = goal thy "inj(f) ==> (f(x) = f(y)) = (x=y)"; |
923 | 80 |
by (rtac iffI 1); |
81 |
by (etac (major RS injD) 1); |
|
82 |
by (etac arg_cong 1); |
|
83 |
qed "inj_eq"; |
|
84 |
||
4656 | 85 |
val [major] = goal thy "inj(f) ==> (@x. f(x)=f(y)) = y"; |
923 | 86 |
by (rtac (major RS injD) 1); |
87 |
by (rtac selectI 1); |
|
88 |
by (rtac refl 1); |
|
89 |
qed "inj_select"; |
|
90 |
||
91 |
(*A one-to-one function has an inverse (given using select).*) |
|
4656 | 92 |
val [major] = goalw thy [inv_def] "inj(f) ==> inv f (f x) = x"; |
923 | 93 |
by (EVERY1 [rtac (major RS inj_select)]); |
2912 | 94 |
qed "inv_f_f"; |
923 | 95 |
|
96 |
(* Useful??? *) |
|
4656 | 97 |
val [oneone,minor] = goal thy |
2912 | 98 |
"[| inj(f); !!y. y: range(f) ==> P(inv f y) |] ==> P(x)"; |
99 |
by (res_inst_tac [("t", "x")] (oneone RS (inv_f_f RS subst)) 1); |
|
923 | 100 |
by (rtac (rangeI RS minor) 1); |
101 |
qed "inj_transfer"; |
|
102 |
||
103 |
||
4830 | 104 |
(*** inj_on f A: f is one-to-one over A ***) |
923 | 105 |
|
4830 | 106 |
val prems = goalw thy [inj_on_def] |
107 |
"(!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> inj_on f A"; |
|
4089 | 108 |
by (blast_tac (claset() addIs prems) 1); |
4830 | 109 |
qed "inj_onI"; |
923 | 110 |
|
4656 | 111 |
val [major] = goal thy |
4830 | 112 |
"(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A"; |
113 |
by (rtac inj_onI 1); |
|
923 | 114 |
by (etac (apply_inverse RS trans) 1); |
115 |
by (REPEAT (eresolve_tac [asm_rl,major] 1)); |
|
4830 | 116 |
qed "inj_on_inverseI"; |
923 | 117 |
|
4830 | 118 |
val major::prems = goalw thy [inj_on_def] |
119 |
"[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y"; |
|
923 | 120 |
by (rtac (major RS bspec RS bspec RS mp) 1); |
121 |
by (REPEAT (resolve_tac prems 1)); |
|
4830 | 122 |
qed "inj_onD"; |
923 | 123 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
124 |
Goal "[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)"; |
4830 | 125 |
by (blast_tac (claset() addSDs [inj_onD]) 1); |
126 |
qed "inj_on_iff"; |
|
923 | 127 |
|
4656 | 128 |
val major::prems = goal thy |
4830 | 129 |
"[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)"; |
923 | 130 |
by (rtac contrapos 1); |
4830 | 131 |
by (etac (major RS inj_onD) 2); |
923 | 132 |
by (REPEAT (resolve_tac prems 1)); |
4830 | 133 |
qed "inj_on_contraD"; |
923 | 134 |
|
5069 | 135 |
Goalw [inj_on_def] |
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
136 |
"[| A<=B; inj_on f B |] ==> inj_on f A"; |
3341 | 137 |
by (Blast_tac 1); |
4830 | 138 |
qed "subset_inj_on"; |
3341 | 139 |
|
923 | 140 |
|
141 |
(*** Lemmas about inj ***) |
|
142 |
||
5069 | 143 |
Goalw [o_def] |
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
144 |
"[| inj(f); inj_on g (range f) |] ==> inj(g o f)"; |
4830 | 145 |
by (fast_tac (claset() addIs [injI] addEs [injD, inj_onD]) 1); |
923 | 146 |
qed "comp_inj"; |
147 |
||
4830 | 148 |
val [prem] = goal thy "inj(f) ==> inj_on f A"; |
149 |
by (blast_tac (claset() addIs [prem RS injD, inj_onI]) 1); |
|
923 | 150 |
qed "inj_imp"; |
151 |
||
4656 | 152 |
val [prem] = goalw thy [inv_def] "y : range(f) ==> f(inv f y) = y"; |
923 | 153 |
by (EVERY1 [rtac (prem RS rangeE), rtac selectI, etac sym]); |
2912 | 154 |
qed "f_inv_f"; |
923 | 155 |
|
4656 | 156 |
val prems = goal thy |
2912 | 157 |
"[| inv f x=inv f y; x: range(f); y: range(f) |] ==> x=y"; |
158 |
by (rtac (arg_cong RS box_equals) 1); |
|
159 |
by (REPEAT (resolve_tac (prems @ [f_inv_f]) 1)); |
|
160 |
qed "inv_injective"; |
|
161 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
162 |
Goal "[| inj(f); A<=range(f) |] ==> inj_on (inv f) A"; |
4830 | 163 |
by (fast_tac (claset() addIs [inj_onI] |
2912 | 164 |
addEs [inv_injective,injD]) 1); |
4830 | 165 |
qed "inj_on_inv"; |
923 | 166 |
|
5069 | 167 |
Goalw [inj_on_def] |
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
168 |
"[| inj_on f C; A<=C; B<=C |] ==> f``(A Int B) = f``A Int f``B"; |
4059 | 169 |
by (Blast_tac 1); |
4830 | 170 |
qed "inj_on_image_Int"; |
4059 | 171 |
|
5069 | 172 |
Goalw [inj_on_def] |
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
173 |
"[| inj_on f C; A<=C; B<=C |] ==> f``(A-B) = f``A - f``B"; |
4059 | 174 |
by (Blast_tac 1); |
4830 | 175 |
qed "inj_on_image_set_diff"; |
4059 | 176 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
177 |
Goalw [inj_def] "inj f ==> f``(A Int B) = f``A Int f``B"; |
4059 | 178 |
by (Blast_tac 1); |
179 |
qed "image_Int"; |
|
180 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
181 |
Goalw [inj_def] "inj f ==> f``(A-B) = f``A - f``B"; |
4059 | 182 |
by (Blast_tac 1); |
183 |
qed "image_set_diff"; |
|
184 |
||
923 | 185 |
|
4089 | 186 |
val set_cs = claset() delrules [equalityI]; |
5305 | 187 |
|
188 |
||
189 |
section "fun_upd"; |
|
190 |
||
191 |
Goalw [fun_upd_def] "(f(x:=y) = f) = (f x = y)"; |
|
192 |
by Safe_tac; |
|
193 |
by (etac subst 1); |
|
194 |
by (rtac ext 2); |
|
195 |
by Auto_tac; |
|
196 |
qed "fun_upd_idem_iff"; |
|
197 |
||
198 |
(* f x = y ==> f(x:=y) = f *) |
|
199 |
bind_thm("fun_upd_idem", fun_upd_idem_iff RS iffD2); |
|
200 |
||
201 |
(* f(x := f x) = f *) |
|
202 |
AddIffs [refl RS fun_upd_idem]; |
|
203 |
||
204 |
Goal "(f(x:=y))z = (if z=x then y else f z)"; |
|
205 |
by (simp_tac (simpset() addsimps [fun_upd_def]) 1); |
|
206 |
qed "fun_upd_apply"; |
|
207 |
Addsimps [fun_upd_apply]; |
|
208 |
||
209 |
qed_goal "fun_upd_same" thy "(f(x:=y)) x = y" |
|
210 |
(K [Simp_tac 1]); |
|
5306 | 211 |
qed_goal "fun_upd_other" thy "!!X. z~=x ==> (f(x:=y)) z = f z" |
5305 | 212 |
(K [Asm_simp_tac 1]); |
213 |
(*Addsimps [fun_upd_same, fun_upd_other];*) |
|
214 |
||
215 |
Goal "a ~= c ==> m(a:=b)(c:=d) = m(c:=d)(a:=b)"; |
|
216 |
by (rtac ext 1); |
|
217 |
by (Auto_tac); |
|
218 |
qed "fun_upd_twist"; |