| author | paulson | 
| Mon, 31 Jan 2000 16:18:09 +0100 | |
| changeset 8173 | a9966d5ab84d | 
| parent 8156 | 33d23d0a300e | 
| child 8226 | 07284f7ad262 | 
| 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  | 
|
| 7089 | 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  | 
||
| 5316 | 16  | 
val prems = Goal  | 
| 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  | 
||
| 5316 | 25  | 
Goal "!!Q. ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)";  | 
| 4656 | 26  | 
by (fast_tac (claset() addEs [selectI]) 1);  | 
27  | 
qed "choice";  | 
|
28  | 
||
| 5316 | 29  | 
Goal "!!S. ALL x:S. EX y. Q x y ==> EX f. ALL x:S. Q x (f x)";  | 
| 4656 | 30  | 
by (fast_tac (claset() addEs [selectI]) 1);  | 
31  | 
qed "bchoice";  | 
|
32  | 
||
33  | 
||
| 5608 | 34  | 
section "id";  | 
| 5441 | 35  | 
|
| 7089 | 36  | 
Goalw [id_def] "id x = x";  | 
37  | 
by (rtac refl 1);  | 
|
38  | 
qed "id_apply";  | 
|
| 5608 | 39  | 
Addsimps [id_apply];  | 
| 5441 | 40  | 
|
41  | 
||
| 5306 | 42  | 
section "o";  | 
43  | 
||
| 7089 | 44  | 
Goalw [o_def] "(f o g) x = f (g x)";  | 
45  | 
by (rtac refl 1);  | 
|
46  | 
qed "o_apply";  | 
|
| 5306 | 47  | 
Addsimps [o_apply];  | 
48  | 
||
| 7089 | 49  | 
Goalw [o_def] "f o (g o h) = f o g o h";  | 
50  | 
by (rtac ext 1);  | 
|
51  | 
by (rtac refl 1);  | 
|
52  | 
qed "o_assoc";  | 
|
| 5306 | 53  | 
|
| 7089 | 54  | 
Goalw [id_def] "id o g = g";  | 
55  | 
by (rtac ext 1);  | 
|
56  | 
by (Simp_tac 1);  | 
|
57  | 
qed "id_o";  | 
|
| 5608 | 58  | 
Addsimps [id_o];  | 
| 5306 | 59  | 
|
| 7089 | 60  | 
Goalw [id_def] "f o id = f";  | 
61  | 
by (rtac ext 1);  | 
|
62  | 
by (Simp_tac 1);  | 
|
63  | 
qed "o_id";  | 
|
| 5608 | 64  | 
Addsimps [o_id];  | 
| 5306 | 65  | 
|
66  | 
Goalw [o_def] "(f o g)``r = f``(g``r)";  | 
|
67  | 
by (Blast_tac 1);  | 
|
68  | 
qed "image_compose";  | 
|
69  | 
||
| 7916 | 70  | 
Goal "f``A = (UN x:A. {f x})";
 | 
| 7536 | 71  | 
by (Blast_tac 1);  | 
| 7916 | 72  | 
qed "image_eq_UN";  | 
| 7536 | 73  | 
|
| 5852 | 74  | 
Goalw [o_def] "UNION A (g o f) = UNION (f``A) g";  | 
75  | 
by (Blast_tac 1);  | 
|
| 
6829
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
76  | 
qed "UN_o";  | 
| 5852 | 77  | 
|
| 
7014
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
78  | 
(** lemma for proving injectivity of representation functions for **)  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
79  | 
(** datatypes involving function types **)  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
80  | 
|
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
81  | 
Goalw [o_def]  | 
| 7089 | 82  | 
"[| ! x y. g (f x) = g y --> f x = y; g o f = g o fa |] ==> f = fa";  | 
83  | 
by (rtac ext 1);  | 
|
84  | 
by (etac allE 1);  | 
|
85  | 
by (etac allE 1);  | 
|
86  | 
by (etac mp 1);  | 
|
87  | 
by (etac fun_cong 1);  | 
|
| 
7014
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
88  | 
qed "inj_fun_lemma";  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
89  | 
|
| 5306 | 90  | 
|
91  | 
section "inj";  | 
|
| 6171 | 92  | 
(**NB: inj now just translates to inj_on**)  | 
| 5306 | 93  | 
|
| 923 | 94  | 
(*** inj(f): f is a one-to-one function ***)  | 
95  | 
||
| 6171 | 96  | 
(*for Tools/datatype_rep_proofs*)  | 
97  | 
val [prem] = Goalw [inj_on_def]  | 
|
98  | 
"(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)";  | 
|
99  | 
by (blast_tac (claset() addIs [prem RS spec RS mp]) 1);  | 
|
100  | 
qed "datatype_injI";  | 
|
| 923 | 101  | 
|
| 6171 | 102  | 
Goalw [inj_on_def] "[| inj(f); f(x) = f(y) |] ==> x=y";  | 
| 5316 | 103  | 
by (Blast_tac 1);  | 
| 923 | 104  | 
qed "injD";  | 
105  | 
||
106  | 
(*Useful with the simplifier*)  | 
|
| 5316 | 107  | 
Goal "inj(f) ==> (f(x) = f(y)) = (x=y)";  | 
| 923 | 108  | 
by (rtac iffI 1);  | 
| 5316 | 109  | 
by (etac arg_cong 2);  | 
110  | 
by (etac injD 1);  | 
|
| 5318 | 111  | 
by (assume_tac 1);  | 
| 923 | 112  | 
qed "inj_eq";  | 
113  | 
||
| 5316 | 114  | 
Goal "inj(f) ==> (@x. f(x)=f(y)) = y";  | 
115  | 
by (etac injD 1);  | 
|
| 923 | 116  | 
by (rtac selectI 1);  | 
117  | 
by (rtac refl 1);  | 
|
118  | 
qed "inj_select";  | 
|
119  | 
||
120  | 
(*A one-to-one function has an inverse (given using select).*)  | 
|
| 5316 | 121  | 
Goalw [inv_def] "inj(f) ==> inv f (f x) = x";  | 
122  | 
by (etac inj_select 1);  | 
|
| 2912 | 123  | 
qed "inv_f_f";  | 
| 7338 | 124  | 
Addsimps [inv_f_f];  | 
| 923 | 125  | 
|
| 7338 | 126  | 
Goal "[| inj(f); f x = y |] ==> inv f y = x";  | 
127  | 
by (etac subst 1);  | 
|
128  | 
by (etac inv_f_f 1);  | 
|
129  | 
qed "inv_f_eq";  | 
|
| 6235 | 130  | 
|
| 923 | 131  | 
(* Useful??? *)  | 
| 5316 | 132  | 
val [oneone,minor] = Goal  | 
| 2912 | 133  | 
"[| inj(f); !!y. y: range(f) ==> P(inv f y) |] ==> P(x)";  | 
134  | 
by (res_inst_tac [("t", "x")] (oneone RS (inv_f_f RS subst)) 1);
 | 
|
| 923 | 135  | 
by (rtac (rangeI RS minor) 1);  | 
136  | 
qed "inj_transfer";  | 
|
137  | 
||
| 
7014
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
138  | 
Goalw [o_def] "[| inj f; f o g = f o h |] ==> g = h";  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
139  | 
by (rtac ext 1);  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
140  | 
by (etac injD 1);  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
141  | 
by (etac fun_cong 1);  | 
| 
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6829 
diff
changeset
 | 
142  | 
qed "inj_o";  | 
| 923 | 143  | 
|
| 4830 | 144  | 
(*** inj_on f A: f is one-to-one over A ***)  | 
| 923 | 145  | 
|
| 5316 | 146  | 
val prems = Goalw [inj_on_def]  | 
| 4830 | 147  | 
"(!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> inj_on f A";  | 
| 4089 | 148  | 
by (blast_tac (claset() addIs prems) 1);  | 
| 4830 | 149  | 
qed "inj_onI";  | 
| 6171 | 150  | 
val injI = inj_onI; (*for compatibility*)  | 
| 923 | 151  | 
|
| 5316 | 152  | 
val [major] = Goal  | 
| 4830 | 153  | 
"(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A";  | 
154  | 
by (rtac inj_onI 1);  | 
|
| 923 | 155  | 
by (etac (apply_inverse RS trans) 1);  | 
156  | 
by (REPEAT (eresolve_tac [asm_rl,major] 1));  | 
|
| 4830 | 157  | 
qed "inj_on_inverseI";  | 
| 6171 | 158  | 
val inj_inverseI = inj_on_inverseI; (*for compatibility*)  | 
| 923 | 159  | 
|
| 5316 | 160  | 
Goalw [inj_on_def] "[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y";  | 
161  | 
by (Blast_tac 1);  | 
|
| 4830 | 162  | 
qed "inj_onD";  | 
| 923 | 163  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
164  | 
Goal "[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)";  | 
| 4830 | 165  | 
by (blast_tac (claset() addSDs [inj_onD]) 1);  | 
166  | 
qed "inj_on_iff";  | 
|
| 923 | 167  | 
|
| 5316 | 168  | 
Goalw [inj_on_def] "[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)";  | 
169  | 
by (Blast_tac 1);  | 
|
| 4830 | 170  | 
qed "inj_on_contraD";  | 
| 923 | 171  | 
|
| 8156 | 172  | 
Goal "inj (%s. {s})";
 | 
173  | 
br injI 1;  | 
|
174  | 
be singleton_inject 1;  | 
|
175  | 
qed "inj_singleton";  | 
|
176  | 
||
| 5316 | 177  | 
Goalw [inj_on_def] "[| A<=B; inj_on f B |] ==> inj_on f A";  | 
| 3341 | 178  | 
by (Blast_tac 1);  | 
| 4830 | 179  | 
qed "subset_inj_on";  | 
| 3341 | 180  | 
|
| 923 | 181  | 
|
| 6235 | 182  | 
(** surj **)  | 
183  | 
||
| 6267 | 184  | 
val [prem] = Goalw [surj_def] "(!! x. g(f x) = x) ==> surj g";  | 
185  | 
by (blast_tac (claset() addIs [prem RS sym]) 1);  | 
|
| 6235 | 186  | 
qed "surjI";  | 
187  | 
||
188  | 
Goalw [surj_def] "surj f ==> range f = UNIV";  | 
|
189  | 
by Auto_tac;  | 
|
190  | 
qed "surj_range";  | 
|
191  | 
||
| 6267 | 192  | 
Goalw [surj_def] "surj f ==> EX x. y = f x";  | 
193  | 
by (Blast_tac 1);  | 
|
194  | 
qed "surjD";  | 
|
195  | 
||
| 6235 | 196  | 
|
| 7374 | 197  | 
(** Bijections **)  | 
198  | 
||
199  | 
Goalw [bij_def] "[| inj f; surj f |] ==> bij f";  | 
|
200  | 
by (Blast_tac 1);  | 
|
201  | 
qed "bijI";  | 
|
202  | 
||
203  | 
Goalw [bij_def] "bij f ==> inj f";  | 
|
204  | 
by (Blast_tac 1);  | 
|
205  | 
qed "bij_is_inj";  | 
|
206  | 
||
207  | 
Goalw [bij_def] "bij f ==> surj f";  | 
|
208  | 
by (Blast_tac 1);  | 
|
209  | 
qed "bij_is_surj";  | 
|
210  | 
||
211  | 
||
| 6171 | 212  | 
(*** Lemmas about injective functions and inv ***)  | 
| 923 | 213  | 
|
| 7051 | 214  | 
Goalw [o_def] "[| inj_on f A; inj_on g (f``A) |] ==> inj_on (g o f) A";  | 
| 6171 | 215  | 
by (fast_tac (claset() addIs [inj_onI] addEs [inj_onD]) 1);  | 
216  | 
qed "comp_inj_on";  | 
|
| 923 | 217  | 
|
| 5316 | 218  | 
Goalw [inv_def] "y : range(f) ==> f(inv f y) = y";  | 
219  | 
by (fast_tac (claset() addIs [selectI]) 1);  | 
|
| 2912 | 220  | 
qed "f_inv_f";  | 
| 923 | 221  | 
|
| 6235 | 222  | 
Goal "surj f ==> f(inv f y) = y";  | 
223  | 
by (asm_simp_tac (simpset() addsimps [f_inv_f, surj_range]) 1);  | 
|
224  | 
qed "surj_f_inv_f";  | 
|
225  | 
||
| 6171 | 226  | 
Goal "[| inv f x = inv f y; x: range(f); y: range(f) |] ==> x=y";  | 
| 2912 | 227  | 
by (rtac (arg_cong RS box_equals) 1);  | 
| 5316 | 228  | 
by (REPEAT (ares_tac [f_inv_f] 1));  | 
| 2912 | 229  | 
qed "inv_injective";  | 
230  | 
||
| 6235 | 231  | 
Goal "A <= range(f) ==> inj_on (inv f) A";  | 
| 4830 | 232  | 
by (fast_tac (claset() addIs [inj_onI]  | 
| 6235 | 233  | 
addEs [inv_injective, injD]) 1);  | 
| 4830 | 234  | 
qed "inj_on_inv";  | 
| 923 | 235  | 
|
| 6235 | 236  | 
Goal "surj f ==> inj (inv f)";  | 
237  | 
by (asm_simp_tac (simpset() addsimps [inj_on_inv, surj_range]) 1);  | 
|
238  | 
qed "surj_imp_inj_inv";  | 
|
239  | 
||
| 7514 | 240  | 
(** We seem to need both the id-forms and the (%x. x) forms; the latter can  | 
241  | 
arise by rewriting, while id may be used explicitly. **)  | 
|
242  | 
||
243  | 
Goal "(%x. x) `` Y = Y";  | 
|
244  | 
by (Blast_tac 1);  | 
|
245  | 
qed "image_ident";  | 
|
246  | 
||
247  | 
Goalw [id_def] "id `` Y = Y";  | 
|
248  | 
by (Blast_tac 1);  | 
|
249  | 
qed "image_id";  | 
|
250  | 
Addsimps [image_ident, image_id];  | 
|
251  | 
||
252  | 
Goal "(%x. x) -`` Y = Y";  | 
|
253  | 
by (Blast_tac 1);  | 
|
254  | 
qed "vimage_ident";  | 
|
255  | 
||
256  | 
Goalw [id_def] "id -`` A = A";  | 
|
257  | 
by Auto_tac;  | 
|
258  | 
qed "vimage_id";  | 
|
259  | 
Addsimps [vimage_ident, vimage_id];  | 
|
260  | 
||
| 7876 | 261  | 
Goal "f -`` (f `` A) = {y. EX x:A. f x = f y}";
 | 
262  | 
by (blast_tac (claset() addIs [sym]) 1);  | 
|
263  | 
qed "vimage_image_eq";  | 
|
264  | 
||
| 8173 | 265  | 
Goal "f `` (f -`` A) <= A";  | 
266  | 
by (Blast_tac 1);  | 
|
267  | 
qed "image_vimage_subset";  | 
|
268  | 
||
269  | 
Goal "f `` (f -`` A) = A Int range f";  | 
|
270  | 
by (Blast_tac 1);  | 
|
271  | 
qed "image_vimage_eq";  | 
|
272  | 
Addsimps [image_vimage_eq];  | 
|
273  | 
||
274  | 
Goal "surj f ==> f `` (f -`` A) = A";  | 
|
275  | 
by (asm_simp_tac (simpset() addsimps [surj_range]) 1);  | 
|
276  | 
qed "surj_image_vimage_eq";  | 
|
277  | 
||
278  | 
Goalw [inj_on_def] "inj f ==> f -`` (f `` A) = A";  | 
|
279  | 
by (Blast_tac 1);  | 
|
280  | 
qed "inj_vimage_image_eq";  | 
|
281  | 
||
282  | 
Goalw [surj_def] "surj f ==> f -`` B <= A ==> B <= f `` A";  | 
|
283  | 
by (blast_tac (claset() addIs [sym]) 1);  | 
|
284  | 
qed "vimage_subsetD";  | 
|
285  | 
||
286  | 
Goalw [inj_on_def] "inj f ==> B <= f `` A ==> f -`` B <= A";  | 
|
287  | 
by (Blast_tac 1);  | 
|
288  | 
qed "vimage_subsetI";  | 
|
289  | 
||
290  | 
Goalw [bij_def] "bij f ==> (f -`` B <= A) = (B <= f `` A)";  | 
|
291  | 
by (blast_tac (claset() delrules [subsetI]  | 
|
292  | 
addIs [vimage_subsetI, vimage_subsetD]) 1);  | 
|
293  | 
qed "vimage_subset_eq";  | 
|
294  | 
||
| 6290 | 295  | 
Goal "f``(A Int B) <= f``A Int f``B";  | 
296  | 
by (Blast_tac 1);  | 
|
297  | 
qed "image_Int_subset";  | 
|
298  | 
||
299  | 
Goal "f``A - f``B <= f``(A - B)";  | 
|
300  | 
by (Blast_tac 1);  | 
|
301  | 
qed "image_diff_subset";  | 
|
302  | 
||
| 5069 | 303  | 
Goalw [inj_on_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
304  | 
"[| inj_on f C; A<=C; B<=C |] ==> f``(A Int B) = f``A Int f``B";  | 
| 4059 | 305  | 
by (Blast_tac 1);  | 
| 4830 | 306  | 
qed "inj_on_image_Int";  | 
| 4059 | 307  | 
|
| 5069 | 308  | 
Goalw [inj_on_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
309  | 
"[| inj_on f C; A<=C; B<=C |] ==> f``(A-B) = f``A - f``B";  | 
| 4059 | 310  | 
by (Blast_tac 1);  | 
| 4830 | 311  | 
qed "inj_on_image_set_diff";  | 
| 4059 | 312  | 
|
| 6171 | 313  | 
Goalw [inj_on_def] "inj f ==> f``(A Int B) = f``A Int f``B";  | 
| 4059 | 314  | 
by (Blast_tac 1);  | 
315  | 
qed "image_Int";  | 
|
316  | 
||
| 6171 | 317  | 
Goalw [inj_on_def] "inj f ==> f``(A-B) = f``A - f``B";  | 
| 4059 | 318  | 
by (Blast_tac 1);  | 
319  | 
qed "image_set_diff";  | 
|
320  | 
||
| 6235 | 321  | 
Goalw [image_def] "inj(f) ==> inv(f)``(f``X) = X";  | 
322  | 
by Auto_tac;  | 
|
323  | 
qed "inv_image_comp";  | 
|
| 5847 | 324  | 
|
| 6301 | 325  | 
Goal "inj f ==> (f a : f``A) = (a : A)";  | 
326  | 
by (blast_tac (claset() addDs [injD]) 1);  | 
|
327  | 
qed "inj_image_mem_iff";  | 
|
328  | 
||
329  | 
Goal "inj f ==> (f``A = f``B) = (A = B)";  | 
|
330  | 
by (blast_tac (claset() addSEs [equalityE] addDs [injD]) 1);  | 
|
331  | 
qed "inj_image_eq_iff";  | 
|
332  | 
||
| 
6829
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
333  | 
Goal "(f `` (UNION A B)) = (UN x:A.(f `` (B x)))";  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
334  | 
by (Blast_tac 1);  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
335  | 
qed "image_UN";  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
336  | 
|
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
337  | 
(*injectivity's required. Left-to-right inclusion holds even if A is empty*)  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
338  | 
Goalw [inj_on_def]  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
339  | 
"[| inj_on f C; ALL x:A. B x <= C; j:A |] \  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
340  | 
\ ==> f `` (INTER A B) = (INT x:A. f `` B x)";  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
341  | 
by (Blast_tac 1);  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
342  | 
qed "image_INT";  | 
| 
 
50459a995aa3
renamed UNION_o to UN_o (to fit the convention) and added image_UN, image_INT
 
paulson 
parents: 
6301 
diff
changeset
 | 
343  | 
|
| 4089 | 344  | 
val set_cs = claset() delrules [equalityI];  | 
| 5305 | 345  | 
|
346  | 
||
347  | 
section "fun_upd";  | 
|
348  | 
||
349  | 
Goalw [fun_upd_def] "(f(x:=y) = f) = (f x = y)";  | 
|
350  | 
by Safe_tac;  | 
|
351  | 
by (etac subst 1);  | 
|
352  | 
by (rtac ext 2);  | 
|
353  | 
by Auto_tac;  | 
|
354  | 
qed "fun_upd_idem_iff";  | 
|
355  | 
||
356  | 
(* f x = y ==> f(x:=y) = f *)  | 
|
357  | 
bind_thm("fun_upd_idem", fun_upd_idem_iff RS iffD2);
 | 
|
358  | 
||
359  | 
(* f(x := f x) = f *)  | 
|
360  | 
AddIffs [refl RS fun_upd_idem];  | 
|
361  | 
||
362  | 
Goal "(f(x:=y))z = (if z=x then y else f z)";  | 
|
363  | 
by (simp_tac (simpset() addsimps [fun_upd_def]) 1);  | 
|
364  | 
qed "fun_upd_apply";  | 
|
365  | 
Addsimps [fun_upd_apply];  | 
|
366  | 
||
| 7445 | 367  | 
(*fun_upd_apply supersedes these two*)  | 
| 7089 | 368  | 
Goal "(f(x:=y)) x = y";  | 
369  | 
by (Simp_tac 1);  | 
|
370  | 
qed "fun_upd_same";  | 
|
371  | 
||
372  | 
Goal "z~=x ==> (f(x:=y)) z = f z";  | 
|
373  | 
by (Asm_simp_tac 1);  | 
|
374  | 
qed "fun_upd_other";  | 
|
375  | 
||
| 7445 | 376  | 
Goal "f(x:=y,x:=z) = f(x:=z)";  | 
377  | 
by (rtac ext 1);  | 
|
378  | 
by (Simp_tac 1);  | 
|
379  | 
qed "fun_upd_upd";  | 
|
380  | 
Addsimps [fun_upd_upd];  | 
|
| 5305 | 381  | 
|
382  | 
Goal "a ~= c ==> m(a:=b)(c:=d) = m(c:=d)(a:=b)";  | 
|
383  | 
by (rtac ext 1);  | 
|
| 7089 | 384  | 
by Auto_tac;  | 
| 5305 | 385  | 
qed "fun_upd_twist";  | 
| 5852 | 386  | 
|
387  | 
||
388  | 
(*** -> and Pi, by Florian Kammueller and LCP ***)  | 
|
389  | 
||
390  | 
val prems = Goalw [Pi_def]  | 
|
391  | 
"[| !!x. x: A ==> f x: B x; !!x. x ~: A ==> f(x) = (@ y. True)|] \  | 
|
392  | 
\ ==> f: Pi A B";  | 
|
393  | 
by (auto_tac (claset(), simpset() addsimps prems));  | 
|
394  | 
qed "Pi_I";  | 
|
395  | 
||
396  | 
val prems = Goal  | 
|
397  | 
"[| !!x. x: A ==> f x: B; !!x. x ~: A ==> f(x) = (@ y. True)|] ==> f: A funcset B";  | 
|
398  | 
by (blast_tac (claset() addIs Pi_I::prems) 1);  | 
|
399  | 
qed "funcsetI";  | 
|
400  | 
||
401  | 
Goalw [Pi_def] "[|f: Pi A B; x: A|] ==> f x: B x";  | 
|
402  | 
by Auto_tac;  | 
|
403  | 
qed "Pi_mem";  | 
|
404  | 
||
405  | 
Goalw [Pi_def] "[|f: A funcset B; x: A|] ==> f x: B";  | 
|
406  | 
by Auto_tac;  | 
|
407  | 
qed "funcset_mem";  | 
|
408  | 
||
409  | 
Goalw [Pi_def] "[|f: Pi A B; x~: A|] ==> f x = (@ y. True)";  | 
|
410  | 
by Auto_tac;  | 
|
411  | 
qed "apply_arb";  | 
|
412  | 
||
413  | 
Goalw [Pi_def] "[| f: Pi A B; g: Pi A B; ! x: A. f x = g x |] ==> f = g";  | 
|
414  | 
by (rtac ext 1);  | 
|
415  | 
by Auto_tac;  | 
|
416  | 
val Pi_extensionality = ballI RSN (3, result());  | 
|
417  | 
||
| 8138 | 418  | 
|
| 5852 | 419  | 
(*** compose ***)  | 
420  | 
||
421  | 
Goalw [Pi_def, compose_def, restrict_def]  | 
|
422  | 
"[| f: A funcset B; g: B funcset C |]==> compose A g f: A funcset C";  | 
|
423  | 
by Auto_tac;  | 
|
424  | 
qed "funcset_compose";  | 
|
425  | 
||
426  | 
Goal "[| f: A funcset B; g: B funcset C; h: C funcset D |]\  | 
|
427  | 
\ ==> compose A h (compose A g f) = compose A (compose B h g) f";  | 
|
428  | 
by (res_inst_tac [("A","A")] Pi_extensionality 1);
 | 
|
429  | 
by (blast_tac (claset() addIs [funcset_compose]) 1);  | 
|
430  | 
by (blast_tac (claset() addIs [funcset_compose]) 1);  | 
|
431  | 
by (rewrite_goals_tac [Pi_def, compose_def, restrict_def]);  | 
|
432  | 
by Auto_tac;  | 
|
433  | 
qed "compose_assoc";  | 
|
434  | 
||
435  | 
Goal "[| f: A funcset B; g: B funcset C; x: A |]==> compose A g f x = g(f(x))";  | 
|
436  | 
by (asm_full_simp_tac (simpset() addsimps [compose_def, restrict_def]) 1);  | 
|
437  | 
qed "compose_eq";  | 
|
438  | 
||
439  | 
Goal "[| f : A funcset B; f `` A = B; g: B funcset C; g `` B = C |]\  | 
|
440  | 
\ ==> compose A g f `` A = C";  | 
|
441  | 
by (auto_tac (claset(),  | 
|
442  | 
simpset() addsimps [image_def, compose_eq]));  | 
|
443  | 
qed "surj_compose";  | 
|
444  | 
||
445  | 
Goal "[| f : A funcset B; g: B funcset C; f `` A = B; inj_on f A; inj_on g B |]\  | 
|
446  | 
\ ==> inj_on (compose A g f) A";  | 
|
447  | 
by (auto_tac (claset(),  | 
|
| 8081 | 448  | 
simpset() addsimps [inj_on_def, compose_eq]));  | 
| 5852 | 449  | 
qed "inj_on_compose";  | 
450  | 
||
451  | 
||
452  | 
(*** restrict / lam ***)  | 
|
| 8138 | 453  | 
|
454  | 
Goal "f``A <= B ==> (lam x: A. f x) : A funcset B";  | 
|
| 5852 | 455  | 
by (auto_tac (claset(),  | 
456  | 
simpset() addsimps [restrict_def, Pi_def]));  | 
|
457  | 
qed "restrict_in_funcset";  | 
|
458  | 
||
459  | 
val prems = Goalw [restrict_def, Pi_def]  | 
|
460  | 
"(!!x. x: A ==> f x: B x) ==> (lam x: A. f x) : Pi A B";  | 
|
461  | 
by (asm_simp_tac (simpset() addsimps prems) 1);  | 
|
462  | 
qed "restrictI";  | 
|
463  | 
||
464  | 
Goal "x: A ==> (lam y: A. f y) x = f x";  | 
|
465  | 
by (asm_simp_tac (simpset() addsimps [restrict_def]) 1);  | 
|
466  | 
qed "restrict_apply1";  | 
|
467  | 
||
468  | 
Goal "[| x: A; f : A funcset B |] ==> (lam y: A. f y) x : B";  | 
|
469  | 
by (asm_full_simp_tac (simpset() addsimps [restrict_apply1,Pi_def]) 1);  | 
|
470  | 
qed "restrict_apply1_mem";  | 
|
471  | 
||
472  | 
Goal "x ~: A ==> (lam y: A. f y) x = (@ y. True)";  | 
|
473  | 
by (asm_simp_tac (simpset() addsimps [restrict_def]) 1);  | 
|
474  | 
qed "restrict_apply2";  | 
|
475  | 
||
476  | 
val prems = Goal  | 
|
477  | 
"(!!x. x: A ==> f x = g x) ==> (lam x: A. f x) = (lam x: A. g x)";  | 
|
478  | 
by (rtac ext 1);  | 
|
479  | 
by (auto_tac (claset(),  | 
|
480  | 
simpset() addsimps prems@[restrict_def, Pi_def]));  | 
|
481  | 
qed "restrict_ext";  | 
|
482  | 
||
| 8138 | 483  | 
Goalw [inj_on_def, restrict_def] "inj_on (restrict f A) A = inj_on f A";  | 
484  | 
by Auto_tac;  | 
|
485  | 
qed "inj_on_restrict_eq";  | 
|
486  | 
||
| 5852 | 487  | 
|
488  | 
(*** Inverse ***)  | 
|
489  | 
||
490  | 
Goal "[|f `` A = B; x: B |] ==> ? y: A. f y = x";  | 
|
491  | 
by (Blast_tac 1);  | 
|
492  | 
qed "surj_image";  | 
|
493  | 
||
494  | 
Goalw [Inv_def] "[| f `` A = B; f : A funcset B |] \  | 
|
495  | 
\ ==> (lam x: B. (Inv A f) x) : B funcset A";  | 
|
496  | 
by (fast_tac (claset() addIs [restrict_in_funcset, selectI2]) 1);  | 
|
497  | 
qed "Inv_funcset";  | 
|
498  | 
||
499  | 
||
500  | 
Goal "[| f: A funcset B; inj_on f A; f `` A = B; x: A |] \  | 
|
501  | 
\ ==> (lam y: B. (Inv A f) y) (f x) = x";  | 
|
502  | 
by (asm_simp_tac (simpset() addsimps [restrict_apply1, funcset_mem]) 1);  | 
|
| 8081 | 503  | 
by (asm_full_simp_tac (simpset() addsimps [Inv_def, inj_on_def]) 1);  | 
| 5852 | 504  | 
by (rtac selectI2 1);  | 
505  | 
by Auto_tac;  | 
|
506  | 
qed "Inv_f_f";  | 
|
507  | 
||
508  | 
Goal "[| f: A funcset B; f `` A = B; x: B |] \  | 
|
509  | 
\ ==> f ((lam y: B. (Inv A f y)) x) = x";  | 
|
510  | 
by (asm_simp_tac (simpset() addsimps [Inv_def, restrict_apply1]) 1);  | 
|
511  | 
by (fast_tac (claset() addIs [selectI2]) 1);  | 
|
512  | 
qed "f_Inv_f";  | 
|
513  | 
||
514  | 
Goal "[| f: A funcset B; inj_on f A; f `` A = B |]\  | 
|
515  | 
\ ==> compose A (lam y:B. (Inv A f) y) f = (lam x: A. x)";  | 
|
516  | 
by (rtac Pi_extensionality 1);  | 
|
517  | 
by (blast_tac (claset() addIs [funcset_compose, Inv_funcset]) 1);  | 
|
518  | 
by (blast_tac (claset() addIs [restrict_in_funcset]) 1);  | 
|
519  | 
by (asm_simp_tac  | 
|
520  | 
(simpset() addsimps [restrict_apply1, compose_def, Inv_f_f]) 1);  | 
|
521  | 
qed "compose_Inv_id";  | 
|
522  | 
||
523  | 
||
524  | 
(*** Pi and Applyall ***)  | 
|
525  | 
||
526  | 
Goalw [Pi_def] "[| B(x) = {};  x: A |] ==> (PI x: A. B x) = {}";
 | 
|
527  | 
by Auto_tac;  | 
|
528  | 
qed "Pi_eq_empty";  | 
|
529  | 
||
530  | 
Goal "[| (PI x: A. B x) ~= {};  x: A |] ==> B(x) ~= {}";
 | 
|
531  | 
by (blast_tac (HOL_cs addIs [Pi_eq_empty]) 1);  | 
|
532  | 
qed "Pi_total1";  | 
|
533  | 
||
534  | 
Goal "[| a : A; Pi A B ~= {} |] ==> Applyall (Pi A B) a = B a";
 | 
|
535  | 
by (auto_tac (claset(), simpset() addsimps [Applyall_def, Pi_def]));  | 
|
536  | 
by (rename_tac "g z" 1);  | 
|
537  | 
by (res_inst_tac [("x","%y. if  (y = a) then z else g y")] exI 1);
 | 
|
538  | 
by (auto_tac (claset(), simpset() addsimps [split_if_mem1, split_if_eq1]));  | 
|
539  | 
qed "Applyall_beta";  | 
|
540  | 
||
| 
5865
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
541  | 
Goal "Pi {} B = { (%x. @ y. True) }";
 | 
| 
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
542  | 
by (auto_tac (claset() addIs [ext], simpset() addsimps [Pi_def]));  | 
| 
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
543  | 
qed "Pi_empty";  | 
| 5852 | 544  | 
|
| 
5865
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
545  | 
val [major] = Goalw [Pi_def] "(!!x. x: A ==> B x <= C x) ==> Pi A B <= Pi A C";  | 
| 
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
546  | 
by (auto_tac (claset(),  | 
| 
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
547  | 
simpset() addsimps [impOfSubs major]));  | 
| 
 
2303f5a3036d
moved some facts about Pi from ex/PiSets to Fun.ML
 
paulson 
parents: 
5852 
diff
changeset
 | 
548  | 
qed "Pi_mono";  |