| author | nipkow | 
| Wed, 29 May 1996 13:47:43 +0200 | |
| changeset 1772 | ee2be39126d2 | 
| parent 218 | be834b9a0c72 | 
| permissions | -rw-r--r-- | 
| 0 | 1  | 
(* Title: ZF/perm.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1991 University of Cambridge  | 
|
5  | 
||
6  | 
For perm.thy. The theory underlying permutation groups  | 
|
7  | 
-- Composition of relations, the identity relation  | 
|
8  | 
-- Injections, surjections, bijections  | 
|
9  | 
-- Lemmas for the Schroeder-Bernstein Theorem  | 
|
10  | 
*)  | 
|
11  | 
||
12  | 
open Perm;  | 
|
13  | 
||
14  | 
(** Surjective function space **)  | 
|
15  | 
||
16  | 
goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B";  | 
|
17  | 
by (etac CollectD1 1);  | 
|
18  | 
val surj_is_fun = result();  | 
|
19  | 
||
20  | 
goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))";  | 
|
21  | 
by (fast_tac (ZF_cs addIs [apply_equality]  | 
|
22  | 
addEs [range_of_fun,domain_type]) 1);  | 
|
23  | 
val fun_is_surj = result();  | 
|
24  | 
||
25  | 
goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B";  | 
|
26  | 
by (best_tac (ZF_cs addIs [equalityI,apply_Pair] addEs [range_type]) 1);  | 
|
27  | 
val surj_range = result();  | 
|
28  | 
||
29  | 
||
30  | 
(** Injective function space **)  | 
|
31  | 
||
32  | 
goalw Perm.thy [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B";  | 
|
33  | 
by (etac CollectD1 1);  | 
|
34  | 
val inj_is_fun = result();  | 
|
35  | 
||
36  | 
goalw Perm.thy [inj_def]  | 
|
37  | 
"!!f A B. [| <a,b>:f; <c,b>:f; f: inj(A,B) |] ==> a=c";  | 
|
38  | 
by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1));  | 
|
39  | 
by (fast_tac ZF_cs 1);  | 
|
40  | 
val inj_equality = result();  | 
|
41  | 
||
42  | 
(** Bijections -- simple lemmas but no intro/elim rules -- use unfolding **)  | 
|
43  | 
||
44  | 
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)";  | 
|
45  | 
by (etac IntD1 1);  | 
|
46  | 
val bij_is_inj = result();  | 
|
47  | 
||
48  | 
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)";  | 
|
49  | 
by (etac IntD2 1);  | 
|
50  | 
val bij_is_surj = result();  | 
|
51  | 
||
52  | 
(* f: bij(A,B) ==> f: A->B *)  | 
|
53  | 
val bij_is_fun = standard (bij_is_inj RS inj_is_fun);  | 
|
54  | 
||
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
55  | 
|
| 0 | 56  | 
(** Identity function **)  | 
57  | 
||
58  | 
val [prem] = goalw Perm.thy [id_def] "a:A ==> <a,a> : id(A)";  | 
|
59  | 
by (rtac (prem RS lamI) 1);  | 
|
60  | 
val idI = result();  | 
|
61  | 
||
62  | 
val major::prems = goalw Perm.thy [id_def]  | 
|
63  | 
"[| p: id(A); !!x.[| x:A; p=<x,x> |] ==> P \  | 
|
64  | 
\ |] ==> P";  | 
|
65  | 
by (rtac (major RS lamE) 1);  | 
|
66  | 
by (REPEAT (ares_tac prems 1));  | 
|
67  | 
val idE = result();  | 
|
68  | 
||
69  | 
goalw Perm.thy [id_def] "id(A) : A->A";  | 
|
70  | 
by (rtac lam_type 1);  | 
|
71  | 
by (assume_tac 1);  | 
|
72  | 
val id_type = result();  | 
|
73  | 
||
74  | 
val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)";  | 
|
75  | 
by (rtac (prem RS lam_mono) 1);  | 
|
76  | 
val id_mono = result();  | 
|
77  | 
||
78  | 
goalw Perm.thy [inj_def,id_def] "id(A): inj(A,A)";  | 
|
79  | 
by (REPEAT (ares_tac [CollectI,lam_type] 1));  | 
|
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
80  | 
by (simp_tac ZF_ss 1);  | 
| 0 | 81  | 
val id_inj = result();  | 
82  | 
||
83  | 
goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)";  | 
|
84  | 
by (fast_tac (ZF_cs addIs [lam_type,beta]) 1);  | 
|
85  | 
val id_surj = result();  | 
|
86  | 
||
87  | 
goalw Perm.thy [bij_def] "id(A): bij(A,A)";  | 
|
88  | 
by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1);  | 
|
89  | 
val id_bij = result();  | 
|
90  | 
||
91  | 
||
92  | 
(** Converse of a relation **)  | 
|
93  | 
||
94  | 
val [prem] = goal Perm.thy "f: inj(A,B) ==> converse(f) : range(f)->A";  | 
|
95  | 
by (rtac (prem RS inj_is_fun RS PiE) 1);  | 
|
96  | 
by (rtac (converse_type RS PiI) 1);  | 
|
97  | 
by (fast_tac ZF_cs 1);  | 
|
98  | 
by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1);  | 
|
99  | 
by flexflex_tac;  | 
|
100  | 
val inj_converse_fun = result();  | 
|
101  | 
||
102  | 
val prems = goalw Perm.thy [surj_def]  | 
|
103  | 
"f: inj(A,B) ==> converse(f): surj(range(f), A)";  | 
|
104  | 
by (fast_tac (ZF_cs addIs (prems@[inj_converse_fun,apply_Pair,apply_equality,  | 
|
105  | 
converseI,inj_is_fun])) 1);  | 
|
106  | 
val inj_converse_surj = result();  | 
|
107  | 
||
108  | 
(*The premises are equivalent to saying that f is injective...*)  | 
|
109  | 
val prems = goal Perm.thy  | 
|
110  | 
"[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a";  | 
|
111  | 
by (fast_tac (ZF_cs addIs (prems@[apply_Pair,apply_equality,converseI])) 1);  | 
|
112  | 
val left_inverse_lemma = result();  | 
|
113  | 
||
114  | 
val prems = goal Perm.thy  | 
|
115  | 
"[| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a";  | 
|
116  | 
by (fast_tac (ZF_cs addIs (prems@  | 
|
117  | 
[left_inverse_lemma,inj_converse_fun,inj_is_fun])) 1);  | 
|
118  | 
val left_inverse = result();  | 
|
119  | 
||
120  | 
val prems = goal Perm.thy  | 
|
121  | 
"[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b";  | 
|
122  | 
by (rtac (apply_Pair RS (converseD RS apply_equality)) 1);  | 
|
123  | 
by (REPEAT (resolve_tac prems 1));  | 
|
124  | 
val right_inverse_lemma = result();  | 
|
125  | 
||
126  | 
val prems = goal Perm.thy  | 
|
127  | 
"[| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b";  | 
|
128  | 
by (rtac right_inverse_lemma 1);  | 
|
129  | 
by (REPEAT (resolve_tac (prems@ [inj_converse_fun,inj_is_fun]) 1));  | 
|
130  | 
val right_inverse = result();  | 
|
131  | 
||
132  | 
val prems = goal Perm.thy  | 
|
133  | 
"f: inj(A,B) ==> converse(f): inj(range(f), A)";  | 
|
134  | 
bw inj_def; (*rewrite subgoal but not prems!!*)  | 
|
135  | 
by (cut_facts_tac prems 1);  | 
|
136  | 
by (safe_tac ZF_cs);  | 
|
137  | 
(*apply f to both sides and simplify using right_inverse  | 
|
138  | 
-- could also use etac[subst_context RS box_equals] in this proof *)  | 
|
139  | 
by (rtac simp_equals 2);  | 
|
140  | 
by (REPEAT (eresolve_tac [inj_converse_fun, right_inverse RS sym, ssubst] 1  | 
|
141  | 
ORELSE ares_tac [refl,rangeI] 1));  | 
|
142  | 
val inj_converse_inj = result();  | 
|
143  | 
||
144  | 
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)";  | 
|
145  | 
by (etac IntE 1);  | 
|
146  | 
by (eresolve_tac [(surj_range RS subst)] 1);  | 
|
147  | 
by (rtac IntI 1);  | 
|
148  | 
by (etac inj_converse_inj 1);  | 
|
149  | 
by (etac inj_converse_surj 1);  | 
|
150  | 
val bij_converse_bij = result();  | 
|
151  | 
||
152  | 
||
153  | 
(** Composition of two relations **)  | 
|
154  | 
||
155  | 
(*The inductive definition package could derive these theorems for R o S*)  | 
|
156  | 
||
157  | 
goalw Perm.thy [comp_def] "!!r s. [| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";  | 
|
158  | 
by (fast_tac ZF_cs 1);  | 
|
159  | 
val compI = result();  | 
|
160  | 
||
161  | 
val prems = goalw Perm.thy [comp_def]  | 
|
162  | 
"[| xz : r O s; \  | 
|
163  | 
\ !!x y z. [| xz=<x,z>; <x,y>:s; <y,z>:r |] ==> P \  | 
|
164  | 
\ |] ==> P";  | 
|
165  | 
by (cut_facts_tac prems 1);  | 
|
166  | 
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1));  | 
|
167  | 
val compE = result();  | 
|
168  | 
||
169  | 
val compEpair =  | 
|
170  | 
rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac)  | 
|
171  | 
THEN prune_params_tac)  | 
|
172  | 
	(read_instantiate [("xz","<a,c>")] compE);
 | 
|
173  | 
||
174  | 
val comp_cs = ZF_cs addIs [compI,idI] addSEs [compE,idE];  | 
|
175  | 
||
176  | 
(** Domain and Range -- see Suppes, section 3.1 **)  | 
|
177  | 
||
178  | 
(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*)  | 
|
179  | 
goal Perm.thy "range(r O s) <= range(r)";  | 
|
180  | 
by (fast_tac comp_cs 1);  | 
|
181  | 
val range_comp = result();  | 
|
182  | 
||
183  | 
goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)";  | 
|
184  | 
by (rtac (range_comp RS equalityI) 1);  | 
|
185  | 
by (fast_tac comp_cs 1);  | 
|
186  | 
val range_comp_eq = result();  | 
|
187  | 
||
188  | 
goal Perm.thy "domain(r O s) <= domain(s)";  | 
|
189  | 
by (fast_tac comp_cs 1);  | 
|
190  | 
val domain_comp = result();  | 
|
191  | 
||
192  | 
goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)";  | 
|
193  | 
by (rtac (domain_comp RS equalityI) 1);  | 
|
194  | 
by (fast_tac comp_cs 1);  | 
|
195  | 
val domain_comp_eq = result();  | 
|
196  | 
||
| 218 | 197  | 
goal Perm.thy "(r O s)``A = r``(s``A)";  | 
198  | 
by (fast_tac (comp_cs addIs [equalityI]) 1);  | 
|
199  | 
val image_comp = result();  | 
|
200  | 
||
201  | 
||
| 0 | 202  | 
(** Other results **)  | 
203  | 
||
204  | 
goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";  | 
|
205  | 
by (fast_tac comp_cs 1);  | 
|
206  | 
val comp_mono = result();  | 
|
207  | 
||
208  | 
(*composition preserves relations*)  | 
|
209  | 
goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C";  | 
|
210  | 
by (fast_tac comp_cs 1);  | 
|
211  | 
val comp_rel = result();  | 
|
212  | 
||
213  | 
(*associative law for composition*)  | 
|
214  | 
goal Perm.thy "(r O s) O t = r O (s O t)";  | 
|
215  | 
by (fast_tac (comp_cs addIs [equalityI]) 1);  | 
|
216  | 
val comp_assoc = result();  | 
|
217  | 
||
218  | 
(*left identity of composition; provable inclusions are  | 
|
219  | 
id(A) O r <= r  | 
|
220  | 
and [| r<=A*B; B<=C |] ==> r <= id(C) O r *)  | 
|
221  | 
goal Perm.thy "!!r A B. r<=A*B ==> id(B) O r = r";  | 
|
222  | 
by (fast_tac (comp_cs addIs [equalityI]) 1);  | 
|
223  | 
val left_comp_id = result();  | 
|
224  | 
||
225  | 
(*right identity of composition; provable inclusions are  | 
|
226  | 
r O id(A) <= r  | 
|
227  | 
and [| r<=A*B; A<=C |] ==> r <= r O id(C) *)  | 
|
228  | 
goal Perm.thy "!!r A B. r<=A*B ==> r O id(A) = r";  | 
|
229  | 
by (fast_tac (comp_cs addIs [equalityI]) 1);  | 
|
230  | 
val right_comp_id = result();  | 
|
231  | 
||
232  | 
||
233  | 
(** Composition preserves functions, injections, and surjections **)  | 
|
234  | 
||
235  | 
goal Perm.thy "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C";  | 
|
236  | 
by (REPEAT (ares_tac [PiI,comp_rel,ex1I,compI] 1  | 
|
237  | 
ORELSE eresolve_tac [fun_is_rel,apply_Pair,apply_type] 1));  | 
|
238  | 
by (fast_tac (comp_cs addDs [apply_equality]) 1);  | 
|
239  | 
val comp_func = result();  | 
|
240  | 
||
241  | 
goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)";  | 
|
242  | 
by (REPEAT (ares_tac [comp_func,apply_equality,compI,  | 
|
243  | 
apply_Pair,apply_type] 1));  | 
|
244  | 
val comp_func_apply = result();  | 
|
245  | 
||
246  | 
goalw Perm.thy [inj_def]  | 
|
247  | 
"!!f g. [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)";  | 
|
248  | 
by (REPEAT (eresolve_tac [bspec RS bspec RS mp, box_equals] 1  | 
|
249  | 
ORELSE step_tac (ZF_cs addSIs [comp_func,apply_type,comp_func_apply]) 1));  | 
|
250  | 
val comp_inj = result();  | 
|
251  | 
||
252  | 
goalw Perm.thy [surj_def]  | 
|
253  | 
"!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)";  | 
|
254  | 
by (best_tac (ZF_cs addSIs [comp_func,comp_func_apply]) 1);  | 
|
255  | 
val comp_surj = result();  | 
|
256  | 
||
257  | 
goalw Perm.thy [bij_def]  | 
|
258  | 
"!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)";  | 
|
259  | 
by (fast_tac (ZF_cs addIs [comp_inj,comp_surj]) 1);  | 
|
260  | 
val comp_bij = result();  | 
|
261  | 
||
262  | 
||
263  | 
(** Dual properties of inj and surj -- useful for proofs from  | 
|
264  | 
D Pastre. Automatic theorem proving in set theory.  | 
|
265  | 
Artificial Intelligence, 10:1--27, 1978. **)  | 
|
266  | 
||
267  | 
goalw Perm.thy [inj_def]  | 
|
268  | 
"!!f g. [| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)";  | 
|
269  | 
by (safe_tac comp_cs);  | 
|
270  | 
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));  | 
|
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
271  | 
by (asm_simp_tac (FOL_ss addsimps [comp_func_apply]) 1);  | 
| 0 | 272  | 
val comp_mem_injD1 = result();  | 
273  | 
||
274  | 
goalw Perm.thy [inj_def,surj_def]  | 
|
275  | 
"!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)";  | 
|
276  | 
by (safe_tac comp_cs);  | 
|
277  | 
by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1);
 | 
|
278  | 
by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3);
 | 
|
279  | 
by (REPEAT (assume_tac 1));  | 
|
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
280  | 
by (safe_tac comp_cs);  | 
| 
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
281  | 
by (res_inst_tac [("t", "op `(g)")] subst_context 1);
 | 
| 0 | 282  | 
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));  | 
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
283  | 
by (asm_simp_tac (FOL_ss addsimps [comp_func_apply]) 1);  | 
| 0 | 284  | 
val comp_mem_injD2 = result();  | 
285  | 
||
286  | 
goalw Perm.thy [surj_def]  | 
|
287  | 
"!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)";  | 
|
288  | 
by (fast_tac (comp_cs addSIs [comp_func_apply RS sym, apply_type]) 1);  | 
|
289  | 
val comp_mem_surjD1 = result();  | 
|
290  | 
||
291  | 
goal Perm.thy  | 
|
292  | 
"!!f g. [| (f O g)`a = c; g: A->B; f: B->C; a:A |] ==> f`(g`a) = c";  | 
|
293  | 
by (REPEAT (ares_tac [comp_func_apply RS sym RS trans] 1));  | 
|
294  | 
val comp_func_applyD = result();  | 
|
295  | 
||
296  | 
goalw Perm.thy [inj_def,surj_def]  | 
|
297  | 
"!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)";  | 
|
298  | 
by (safe_tac comp_cs);  | 
|
299  | 
by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1);
 | 
|
300  | 
by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_func_applyD 1));  | 
|
301  | 
by (best_tac (comp_cs addSIs [apply_type]) 1);  | 
|
302  | 
val comp_mem_surjD2 = result();  | 
|
303  | 
||
304  | 
||
305  | 
(** inverses of composition **)  | 
|
306  | 
||
307  | 
(*left inverse of composition; one inclusion is  | 
|
308  | 
f: A->B ==> id(A) <= converse(f) O f *)  | 
|
309  | 
val [prem] = goal Perm.thy  | 
|
310  | 
"f: inj(A,B) ==> converse(f) O f = id(A)";  | 
|
311  | 
val injfD = prem RSN (3,inj_equality);  | 
|
312  | 
by (cut_facts_tac [prem RS inj_is_fun] 1);  | 
|
313  | 
by (fast_tac (comp_cs addIs [equalityI,apply_Pair]  | 
|
314  | 
addEs [domain_type, make_elim injfD]) 1);  | 
|
315  | 
val left_comp_inverse = result();  | 
|
316  | 
||
317  | 
(*right inverse of composition; one inclusion is  | 
|
318  | 
f: A->B ==> f O converse(f) <= id(B) *)  | 
|
319  | 
val [prem] = goalw Perm.thy [surj_def]  | 
|
320  | 
"f: surj(A,B) ==> f O converse(f) = id(B)";  | 
|
321  | 
val appfD = (prem RS CollectD1) RSN (3,apply_equality2);  | 
|
322  | 
by (cut_facts_tac [prem] 1);  | 
|
323  | 
by (rtac equalityI 1);  | 
|
324  | 
by (best_tac (comp_cs addEs [domain_type, range_type, make_elim appfD]) 1);  | 
|
325  | 
by (best_tac (comp_cs addIs [apply_Pair]) 1);  | 
|
326  | 
val right_comp_inverse = result();  | 
|
327  | 
||
328  | 
(*Injective case applies converse(f) to both sides then simplifies  | 
|
329  | 
using left_inverse_lemma*)  | 
|
330  | 
goalw Perm.thy [bij_def,inj_def,surj_def]  | 
|
331  | 
"!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)";  | 
|
332  | 
val cf_cong = read_instantiate_sg (sign_of Perm.thy)  | 
|
333  | 
                [("t","%x.?f`x")]   subst_context;
 | 
|
334  | 
by (fast_tac (ZF_cs addIs [left_inverse_lemma, right_inverse_lemma, apply_type]  | 
|
335  | 
addEs [cf_cong RS box_equals]) 1);  | 
|
336  | 
val invertible_imp_bijective = result();  | 
|
337  | 
||
338  | 
(** Unions of functions -- cf similar theorems on func.ML **)  | 
|
339  | 
||
340  | 
goal Perm.thy "converse(r Un s) = converse(r) Un converse(s)";  | 
|
341  | 
by (rtac equalityI 1);  | 
|
342  | 
by (DEPTH_SOLVE_1  | 
|
343  | 
(resolve_tac [Un_least,converse_mono, Un_upper1,Un_upper2] 2));  | 
|
344  | 
by (fast_tac ZF_cs 1);  | 
|
345  | 
val converse_of_Un = result();  | 
|
346  | 
||
347  | 
goalw Perm.thy [surj_def]  | 
|
348  | 
"!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \  | 
|
349  | 
\ (f Un g) : surj(A Un C, B Un D)";  | 
|
350  | 
by (DEPTH_SOLVE_1 (eresolve_tac [fun_disjoint_apply1, fun_disjoint_apply2] 1  | 
|
351  | 
ORELSE ball_tac 1  | 
|
352  | 
ORELSE (rtac trans 1 THEN atac 2)  | 
|
353  | 
ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1));  | 
|
354  | 
val surj_disjoint_Un = result();  | 
|
355  | 
||
356  | 
(*A simple, high-level proof; the version for injections follows from it,  | 
|
357  | 
using f:inj(A,B)<->f:bij(A,range(f)) *)  | 
|
358  | 
goal Perm.thy  | 
|
359  | 
"!!f g. [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==> \  | 
|
360  | 
\ (f Un g) : bij(A Un C, B Un D)";  | 
|
361  | 
by (rtac invertible_imp_bijective 1);  | 
|
362  | 
by (rtac (converse_of_Un RS ssubst) 1);  | 
|
363  | 
by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1));  | 
|
364  | 
val bij_disjoint_Un = result();  | 
|
365  | 
||
366  | 
||
367  | 
(** Restrictions as surjections and bijections *)  | 
|
368  | 
||
369  | 
val prems = goalw Perm.thy [surj_def]  | 
|
370  | 
"f: Pi(A,B) ==> f: surj(A, f``A)";  | 
|
371  | 
val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]);  | 
|
372  | 
by (fast_tac (ZF_cs addIs rls) 1);  | 
|
373  | 
val surj_image = result();  | 
|
374  | 
||
375  | 
goal Perm.thy  | 
|
376  | 
"!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A";  | 
|
377  | 
by (rtac equalityI 1);  | 
|
378  | 
by (SELECT_GOAL (rewtac restrict_def) 2);  | 
|
379  | 
by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2  | 
|
380  | 
ORELSE ares_tac [subsetI,lamI,imageI] 2));  | 
|
381  | 
by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1));  | 
|
382  | 
val restrict_image = result();  | 
|
383  | 
||
384  | 
goalw Perm.thy [inj_def]  | 
|
385  | 
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)";  | 
|
386  | 
by (safe_tac (ZF_cs addSEs [restrict_type2]));  | 
|
387  | 
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD,  | 
|
388  | 
box_equals, restrict] 1));  | 
|
389  | 
val restrict_inj = result();  | 
|
390  | 
||
391  | 
val prems = goal Perm.thy  | 
|
392  | 
"[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)";  | 
|
393  | 
by (rtac (restrict_image RS subst) 1);  | 
|
394  | 
by (rtac (restrict_type2 RS surj_image) 3);  | 
|
395  | 
by (REPEAT (resolve_tac prems 1));  | 
|
396  | 
val restrict_surj = result();  | 
|
397  | 
||
398  | 
goalw Perm.thy [inj_def,bij_def]  | 
|
399  | 
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)";  | 
|
400  | 
by (safe_tac ZF_cs);  | 
|
401  | 
by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD,  | 
|
402  | 
box_equals, restrict] 1  | 
|
403  | 
ORELSE ares_tac [surj_is_fun,restrict_surj] 1));  | 
|
404  | 
val restrict_bij = result();  | 
|
405  | 
||
406  | 
||
407  | 
(*** Lemmas for Ramsey's Theorem ***)  | 
|
408  | 
||
409  | 
goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)";  | 
|
410  | 
by (fast_tac (ZF_cs addSEs [fun_weaken_type]) 1);  | 
|
411  | 
val inj_weaken_type = result();  | 
|
412  | 
||
413  | 
val [major] = goal Perm.thy  | 
|
414  | 
    "[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})";
 | 
|
415  | 
by (rtac (major RS restrict_bij RS bij_is_inj RS inj_weaken_type) 1);  | 
|
416  | 
by (fast_tac ZF_cs 1);  | 
|
417  | 
by (cut_facts_tac [major] 1);  | 
|
418  | 
by (rewtac inj_def);  | 
|
419  | 
by (safe_tac ZF_cs);  | 
|
420  | 
by (etac range_type 1);  | 
|
421  | 
by (assume_tac 1);  | 
|
422  | 
by (dtac apply_equality 1);  | 
|
423  | 
by (assume_tac 1);  | 
|
424  | 
by (res_inst_tac [("a","m")] mem_anti_refl 1);
 | 
|
425  | 
by (fast_tac ZF_cs 1);  | 
|
426  | 
val inj_succ_restrict = result();  | 
|
427  | 
||
428  | 
goalw Perm.thy [inj_def]  | 
|
| 37 | 429  | 
"!!f. [| f: inj(A,B); a~:A; b~:B |] ==> \  | 
| 0 | 430  | 
\ cons(<a,b>,f) : inj(cons(a,A), cons(b,B))";  | 
431  | 
(*cannot use safe_tac: must preserve the implication*)  | 
|
432  | 
by (etac CollectE 1);  | 
|
433  | 
by (rtac CollectI 1);  | 
|
434  | 
by (etac fun_extend 1);  | 
|
435  | 
by (REPEAT (ares_tac [ballI] 1));  | 
|
436  | 
by (REPEAT_FIRST (eresolve_tac [consE,ssubst]));  | 
|
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
437  | 
(*Assumption ALL w:A. ALL x:A. f`w = f`x --> w=x would make asm_simp_tac loop  | 
| 
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
438  | 
using ZF_ss! But FOL_ss ignores the assumption...*)  | 
| 
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
439  | 
by (ALLGOALS (asm_simp_tac  | 
| 
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
440  | 
(FOL_ss addsimps [fun_extend_apply2,fun_extend_apply1])));  | 
| 0 | 441  | 
by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type])));  | 
442  | 
val inj_extend = result();  |