| author | nipkow |
| Thu, 21 Jul 1994 14:27:00 +0200 | |
| changeset 482 | 3a4e092ba69c |
| parent 437 | 435875e4b21d |
| child 484 | 70b789956bd3 |
| 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 |
||
| 435 | 78 |
goalw Perm.thy [inj_def,id_def] "!!A B. A<=B ==> id(A): inj(A,B)"; |
| 0 | 79 |
by (REPEAT (ares_tac [CollectI,lam_type] 1)); |
| 435 | 80 |
by (etac subsetD 1 THEN assume_tac 1); |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
81 |
by (simp_tac ZF_ss 1); |
| 435 | 82 |
val id_subset_inj = result(); |
83 |
||
84 |
val id_inj = subset_refl RS id_subset_inj; |
|
| 0 | 85 |
|
86 |
goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)"; |
|
87 |
by (fast_tac (ZF_cs addIs [lam_type,beta]) 1); |
|
88 |
val id_surj = result(); |
|
89 |
||
90 |
goalw Perm.thy [bij_def] "id(A): bij(A,A)"; |
|
91 |
by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1); |
|
92 |
val id_bij = result(); |
|
93 |
||
94 |
||
95 |
(** Converse of a relation **) |
|
96 |
||
97 |
val [prem] = goal Perm.thy "f: inj(A,B) ==> converse(f) : range(f)->A"; |
|
98 |
by (rtac (prem RS inj_is_fun RS PiE) 1); |
|
99 |
by (rtac (converse_type RS PiI) 1); |
|
100 |
by (fast_tac ZF_cs 1); |
|
101 |
by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1); |
|
102 |
by flexflex_tac; |
|
103 |
val inj_converse_fun = result(); |
|
104 |
||
105 |
val prems = goalw Perm.thy [surj_def] |
|
106 |
"f: inj(A,B) ==> converse(f): surj(range(f), A)"; |
|
107 |
by (fast_tac (ZF_cs addIs (prems@[inj_converse_fun,apply_Pair,apply_equality, |
|
108 |
converseI,inj_is_fun])) 1); |
|
109 |
val inj_converse_surj = result(); |
|
110 |
||
111 |
(*The premises are equivalent to saying that f is injective...*) |
|
112 |
val prems = goal Perm.thy |
|
113 |
"[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a"; |
|
114 |
by (fast_tac (ZF_cs addIs (prems@[apply_Pair,apply_equality,converseI])) 1); |
|
115 |
val left_inverse_lemma = result(); |
|
116 |
||
| 435 | 117 |
goal Perm.thy |
118 |
"!!f. [| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a"; |
|
119 |
by (fast_tac (ZF_cs addIs [left_inverse_lemma,inj_converse_fun,inj_is_fun]) 1); |
|
| 0 | 120 |
val left_inverse = result(); |
121 |
||
| 435 | 122 |
val left_inverse_bij = bij_is_inj RS left_inverse; |
123 |
||
| 0 | 124 |
val prems = goal Perm.thy |
125 |
"[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b"; |
|
126 |
by (rtac (apply_Pair RS (converseD RS apply_equality)) 1); |
|
127 |
by (REPEAT (resolve_tac prems 1)); |
|
128 |
val right_inverse_lemma = result(); |
|
129 |
||
| 435 | 130 |
goal Perm.thy |
131 |
"!!f. [| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b"; |
|
| 0 | 132 |
by (rtac right_inverse_lemma 1); |
| 435 | 133 |
by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1)); |
| 0 | 134 |
val right_inverse = result(); |
135 |
||
| 435 | 136 |
goalw Perm.thy [bij_def] |
137 |
"!!f. [| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b"; |
|
138 |
by (EVERY1 [etac IntE, etac right_inverse, |
|
139 |
etac (surj_range RS ssubst), |
|
140 |
assume_tac]); |
|
141 |
val right_inverse_bij = result(); |
|
142 |
||
| 0 | 143 |
val prems = goal Perm.thy |
144 |
"f: inj(A,B) ==> converse(f): inj(range(f), A)"; |
|
| 437 | 145 |
by (rewtac inj_def); (*rewrite subgoal but not prems!!*) |
| 0 | 146 |
by (cut_facts_tac prems 1); |
147 |
by (safe_tac ZF_cs); |
|
148 |
(*apply f to both sides and simplify using right_inverse |
|
149 |
-- could also use etac[subst_context RS box_equals] in this proof *) |
|
150 |
by (rtac simp_equals 2); |
|
151 |
by (REPEAT (eresolve_tac [inj_converse_fun, right_inverse RS sym, ssubst] 1 |
|
152 |
ORELSE ares_tac [refl,rangeI] 1)); |
|
153 |
val inj_converse_inj = result(); |
|
154 |
||
155 |
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)"; |
|
156 |
by (etac IntE 1); |
|
157 |
by (eresolve_tac [(surj_range RS subst)] 1); |
|
158 |
by (rtac IntI 1); |
|
159 |
by (etac inj_converse_inj 1); |
|
160 |
by (etac inj_converse_surj 1); |
|
161 |
val bij_converse_bij = result(); |
|
162 |
||
163 |
||
164 |
(** Composition of two relations **) |
|
165 |
||
166 |
(*The inductive definition package could derive these theorems for R o S*) |
|
167 |
||
168 |
goalw Perm.thy [comp_def] "!!r s. [| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s"; |
|
169 |
by (fast_tac ZF_cs 1); |
|
170 |
val compI = result(); |
|
171 |
||
172 |
val prems = goalw Perm.thy [comp_def] |
|
173 |
"[| xz : r O s; \ |
|
174 |
\ !!x y z. [| xz=<x,z>; <x,y>:s; <y,z>:r |] ==> P \ |
|
175 |
\ |] ==> P"; |
|
176 |
by (cut_facts_tac prems 1); |
|
177 |
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); |
|
178 |
val compE = result(); |
|
179 |
||
180 |
val compEpair = |
|
181 |
rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) |
|
182 |
THEN prune_params_tac) |
|
183 |
(read_instantiate [("xz","<a,c>")] compE);
|
|
184 |
||
185 |
val comp_cs = ZF_cs addIs [compI,idI] addSEs [compE,idE]; |
|
186 |
||
187 |
(** Domain and Range -- see Suppes, section 3.1 **) |
|
188 |
||
189 |
(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*) |
|
190 |
goal Perm.thy "range(r O s) <= range(r)"; |
|
191 |
by (fast_tac comp_cs 1); |
|
192 |
val range_comp = result(); |
|
193 |
||
194 |
goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)"; |
|
195 |
by (rtac (range_comp RS equalityI) 1); |
|
196 |
by (fast_tac comp_cs 1); |
|
197 |
val range_comp_eq = result(); |
|
198 |
||
199 |
goal Perm.thy "domain(r O s) <= domain(s)"; |
|
200 |
by (fast_tac comp_cs 1); |
|
201 |
val domain_comp = result(); |
|
202 |
||
203 |
goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)"; |
|
204 |
by (rtac (domain_comp RS equalityI) 1); |
|
205 |
by (fast_tac comp_cs 1); |
|
206 |
val domain_comp_eq = result(); |
|
207 |
||
| 218 | 208 |
goal Perm.thy "(r O s)``A = r``(s``A)"; |
209 |
by (fast_tac (comp_cs addIs [equalityI]) 1); |
|
210 |
val image_comp = result(); |
|
211 |
||
212 |
||
| 0 | 213 |
(** Other results **) |
214 |
||
215 |
goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; |
|
216 |
by (fast_tac comp_cs 1); |
|
217 |
val comp_mono = result(); |
|
218 |
||
219 |
(*composition preserves relations*) |
|
220 |
goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"; |
|
221 |
by (fast_tac comp_cs 1); |
|
222 |
val comp_rel = result(); |
|
223 |
||
224 |
(*associative law for composition*) |
|
225 |
goal Perm.thy "(r O s) O t = r O (s O t)"; |
|
226 |
by (fast_tac (comp_cs addIs [equalityI]) 1); |
|
227 |
val comp_assoc = result(); |
|
228 |
||
229 |
(*left identity of composition; provable inclusions are |
|
230 |
id(A) O r <= r |
|
231 |
and [| r<=A*B; B<=C |] ==> r <= id(C) O r *) |
|
232 |
goal Perm.thy "!!r A B. r<=A*B ==> id(B) O r = r"; |
|
233 |
by (fast_tac (comp_cs addIs [equalityI]) 1); |
|
234 |
val left_comp_id = result(); |
|
235 |
||
236 |
(*right identity of composition; provable inclusions are |
|
237 |
r O id(A) <= r |
|
238 |
and [| r<=A*B; A<=C |] ==> r <= r O id(C) *) |
|
239 |
goal Perm.thy "!!r A B. r<=A*B ==> r O id(A) = r"; |
|
240 |
by (fast_tac (comp_cs addIs [equalityI]) 1); |
|
241 |
val right_comp_id = result(); |
|
242 |
||
243 |
||
244 |
(** Composition preserves functions, injections, and surjections **) |
|
245 |
||
246 |
goal Perm.thy "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C"; |
|
247 |
by (REPEAT (ares_tac [PiI,comp_rel,ex1I,compI] 1 |
|
248 |
ORELSE eresolve_tac [fun_is_rel,apply_Pair,apply_type] 1)); |
|
249 |
by (fast_tac (comp_cs addDs [apply_equality]) 1); |
|
| 435 | 250 |
val comp_fun = result(); |
| 0 | 251 |
|
252 |
goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; |
|
| 435 | 253 |
by (REPEAT (ares_tac [comp_fun,apply_equality,compI, |
| 0 | 254 |
apply_Pair,apply_type] 1)); |
| 435 | 255 |
val comp_fun_apply = result(); |
| 0 | 256 |
|
257 |
goalw Perm.thy [inj_def] |
|
258 |
"!!f g. [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)"; |
|
259 |
by (REPEAT (eresolve_tac [bspec RS bspec RS mp, box_equals] 1 |
|
| 435 | 260 |
ORELSE step_tac (ZF_cs addSIs [comp_fun,apply_type,comp_fun_apply]) 1)); |
| 0 | 261 |
val comp_inj = result(); |
262 |
||
263 |
goalw Perm.thy [surj_def] |
|
264 |
"!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)"; |
|
| 435 | 265 |
by (best_tac (ZF_cs addSIs [comp_fun,comp_fun_apply]) 1); |
| 0 | 266 |
val comp_surj = result(); |
267 |
||
268 |
goalw Perm.thy [bij_def] |
|
269 |
"!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)"; |
|
270 |
by (fast_tac (ZF_cs addIs [comp_inj,comp_surj]) 1); |
|
271 |
val comp_bij = result(); |
|
272 |
||
273 |
||
274 |
(** Dual properties of inj and surj -- useful for proofs from |
|
275 |
D Pastre. Automatic theorem proving in set theory. |
|
276 |
Artificial Intelligence, 10:1--27, 1978. **) |
|
277 |
||
278 |
goalw Perm.thy [inj_def] |
|
279 |
"!!f g. [| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)"; |
|
280 |
by (safe_tac comp_cs); |
|
281 |
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); |
|
| 435 | 282 |
by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1); |
| 0 | 283 |
val comp_mem_injD1 = result(); |
284 |
||
285 |
goalw Perm.thy [inj_def,surj_def] |
|
286 |
"!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)"; |
|
287 |
by (safe_tac comp_cs); |
|
288 |
by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1);
|
|
289 |
by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3);
|
|
290 |
by (REPEAT (assume_tac 1)); |
|
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
291 |
by (safe_tac comp_cs); |
|
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
292 |
by (res_inst_tac [("t", "op `(g)")] subst_context 1);
|
| 0 | 293 |
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); |
| 435 | 294 |
by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1); |
| 0 | 295 |
val comp_mem_injD2 = result(); |
296 |
||
297 |
goalw Perm.thy [surj_def] |
|
298 |
"!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)"; |
|
| 435 | 299 |
by (fast_tac (comp_cs addSIs [comp_fun_apply RS sym, apply_type]) 1); |
| 0 | 300 |
val comp_mem_surjD1 = result(); |
301 |
||
302 |
goal Perm.thy |
|
303 |
"!!f g. [| (f O g)`a = c; g: A->B; f: B->C; a:A |] ==> f`(g`a) = c"; |
|
| 435 | 304 |
by (REPEAT (ares_tac [comp_fun_apply RS sym RS trans] 1)); |
305 |
val comp_fun_applyD = result(); |
|
| 0 | 306 |
|
307 |
goalw Perm.thy [inj_def,surj_def] |
|
308 |
"!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)"; |
|
309 |
by (safe_tac comp_cs); |
|
310 |
by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1);
|
|
| 435 | 311 |
by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_fun_applyD 1)); |
| 0 | 312 |
by (best_tac (comp_cs addSIs [apply_type]) 1); |
313 |
val comp_mem_surjD2 = result(); |
|
314 |
||
315 |
||
316 |
(** inverses of composition **) |
|
317 |
||
318 |
(*left inverse of composition; one inclusion is |
|
319 |
f: A->B ==> id(A) <= converse(f) O f *) |
|
320 |
val [prem] = goal Perm.thy |
|
321 |
"f: inj(A,B) ==> converse(f) O f = id(A)"; |
|
322 |
val injfD = prem RSN (3,inj_equality); |
|
323 |
by (cut_facts_tac [prem RS inj_is_fun] 1); |
|
324 |
by (fast_tac (comp_cs addIs [equalityI,apply_Pair] |
|
325 |
addEs [domain_type, make_elim injfD]) 1); |
|
326 |
val left_comp_inverse = result(); |
|
327 |
||
328 |
(*right inverse of composition; one inclusion is |
|
329 |
f: A->B ==> f O converse(f) <= id(B) *) |
|
330 |
val [prem] = goalw Perm.thy [surj_def] |
|
331 |
"f: surj(A,B) ==> f O converse(f) = id(B)"; |
|
332 |
val appfD = (prem RS CollectD1) RSN (3,apply_equality2); |
|
333 |
by (cut_facts_tac [prem] 1); |
|
334 |
by (rtac equalityI 1); |
|
335 |
by (best_tac (comp_cs addEs [domain_type, range_type, make_elim appfD]) 1); |
|
336 |
by (best_tac (comp_cs addIs [apply_Pair]) 1); |
|
337 |
val right_comp_inverse = result(); |
|
338 |
||
| 435 | 339 |
(** Proving that a function is a bijection **) |
340 |
||
341 |
val prems = |
|
342 |
goalw Perm.thy [bij_def, inj_def, surj_def] |
|
343 |
"[| !!x. x:A ==> c(x): B; \ |
|
344 |
\ !!y. y:B ==> d(y): A; \ |
|
345 |
\ !!x. x:A ==> d(c(x)) = x; \ |
|
346 |
\ !!y. y:B ==> c(d(y)) = y \ |
|
347 |
\ |] ==> (lam x:A.c(x)) : bij(A,B)"; |
|
348 |
by (simp_tac (ZF_ss addsimps ([lam_type]@prems)) 1); |
|
349 |
by (safe_tac ZF_cs); |
|
350 |
(*Apply d to both sides then simplify (type constraint is essential) *) |
|
351 |
by (dres_inst_tac [("t", "d::i=>i")] subst_context 1);
|
|
352 |
by (asm_full_simp_tac (ZF_ss addsimps prems) 1); |
|
353 |
by (fast_tac (ZF_cs addIs prems) 1); |
|
354 |
val lam_bijective = result(); |
|
355 |
||
356 |
goalw Perm.thy [id_def] |
|
357 |
"!!f A B. [| f: A->B; g: B->A |] ==> \ |
|
358 |
\ f O g = id(B) <-> (ALL y:B. f`(g`y)=y)"; |
|
359 |
by (safe_tac ZF_cs); |
|
360 |
by (dres_inst_tac [("t", "%h.h`y ")] subst_context 1);
|
|
361 |
by (asm_full_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1); |
|
| 437 | 362 |
by (rtac fun_extension 1); |
| 435 | 363 |
by (REPEAT (ares_tac [comp_fun, lam_type] 1)); |
364 |
by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1); |
|
365 |
val comp_eq_id_iff = result(); |
|
366 |
||
367 |
goalw Perm.thy [bij_def, inj_def, surj_def] |
|
368 |
"!!f A B. [| f: A->B; g: B->A; f O g = id(B); g O f = id(A) \ |
|
369 |
\ |] ==> f : bij(A,B)"; |
|
370 |
by (asm_full_simp_tac (ZF_ss addsimps [comp_eq_id_iff]) 1); |
|
371 |
by (safe_tac ZF_cs); |
|
372 |
(*Apply g to both sides then simplify*) |
|
373 |
by (dres_inst_tac [("t", "op `(g)"), ("a", "f`x")] subst_context 1);
|
|
374 |
by (asm_full_simp_tac ZF_ss 1); |
|
375 |
by (fast_tac (ZF_cs addIs [apply_type]) 1); |
|
376 |
val fg_imp_bijective = result(); |
|
377 |
||
378 |
goal Perm.thy "!!f A. [| f: A->A; f O f = id(A) |] ==> f : bij(A,A)"; |
|
379 |
by (REPEAT (ares_tac [fg_imp_bijective] 1)); |
|
380 |
val nilpotent_imp_bijective = result(); |
|
381 |
||
| 0 | 382 |
(*Injective case applies converse(f) to both sides then simplifies |
383 |
using left_inverse_lemma*) |
|
384 |
goalw Perm.thy [bij_def,inj_def,surj_def] |
|
385 |
"!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; |
|
386 |
val cf_cong = read_instantiate_sg (sign_of Perm.thy) |
|
387 |
[("t","%x.?f`x")] subst_context;
|
|
388 |
by (fast_tac (ZF_cs addIs [left_inverse_lemma, right_inverse_lemma, apply_type] |
|
389 |
addEs [cf_cong RS box_equals]) 1); |
|
390 |
val invertible_imp_bijective = result(); |
|
391 |
||
392 |
(** Unions of functions -- cf similar theorems on func.ML **) |
|
393 |
||
394 |
goal Perm.thy "converse(r Un s) = converse(r) Un converse(s)"; |
|
395 |
by (rtac equalityI 1); |
|
396 |
by (DEPTH_SOLVE_1 |
|
397 |
(resolve_tac [Un_least,converse_mono, Un_upper1,Un_upper2] 2)); |
|
398 |
by (fast_tac ZF_cs 1); |
|
399 |
val converse_of_Un = result(); |
|
400 |
||
401 |
goalw Perm.thy [surj_def] |
|
402 |
"!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \ |
|
403 |
\ (f Un g) : surj(A Un C, B Un D)"; |
|
404 |
by (DEPTH_SOLVE_1 (eresolve_tac [fun_disjoint_apply1, fun_disjoint_apply2] 1 |
|
405 |
ORELSE ball_tac 1 |
|
406 |
ORELSE (rtac trans 1 THEN atac 2) |
|
407 |
ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1)); |
|
408 |
val surj_disjoint_Un = result(); |
|
409 |
||
410 |
(*A simple, high-level proof; the version for injections follows from it, |
|
411 |
using f:inj(A,B)<->f:bij(A,range(f)) *) |
|
412 |
goal Perm.thy |
|
413 |
"!!f g. [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==> \ |
|
414 |
\ (f Un g) : bij(A Un C, B Un D)"; |
|
415 |
by (rtac invertible_imp_bijective 1); |
|
416 |
by (rtac (converse_of_Un RS ssubst) 1); |
|
417 |
by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1)); |
|
418 |
val bij_disjoint_Un = result(); |
|
419 |
||
420 |
||
421 |
(** Restrictions as surjections and bijections *) |
|
422 |
||
423 |
val prems = goalw Perm.thy [surj_def] |
|
424 |
"f: Pi(A,B) ==> f: surj(A, f``A)"; |
|
425 |
val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]); |
|
426 |
by (fast_tac (ZF_cs addIs rls) 1); |
|
427 |
val surj_image = result(); |
|
428 |
||
429 |
goal Perm.thy |
|
430 |
"!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A"; |
|
431 |
by (rtac equalityI 1); |
|
432 |
by (SELECT_GOAL (rewtac restrict_def) 2); |
|
433 |
by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2 |
|
434 |
ORELSE ares_tac [subsetI,lamI,imageI] 2)); |
|
435 |
by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1)); |
|
436 |
val restrict_image = result(); |
|
437 |
||
438 |
goalw Perm.thy [inj_def] |
|
439 |
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)"; |
|
440 |
by (safe_tac (ZF_cs addSEs [restrict_type2])); |
|
441 |
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD, |
|
442 |
box_equals, restrict] 1)); |
|
443 |
val restrict_inj = result(); |
|
444 |
||
445 |
val prems = goal Perm.thy |
|
446 |
"[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)"; |
|
447 |
by (rtac (restrict_image RS subst) 1); |
|
448 |
by (rtac (restrict_type2 RS surj_image) 3); |
|
449 |
by (REPEAT (resolve_tac prems 1)); |
|
450 |
val restrict_surj = result(); |
|
451 |
||
452 |
goalw Perm.thy [inj_def,bij_def] |
|
453 |
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)"; |
|
454 |
by (safe_tac ZF_cs); |
|
455 |
by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD, |
|
456 |
box_equals, restrict] 1 |
|
457 |
ORELSE ares_tac [surj_is_fun,restrict_surj] 1)); |
|
458 |
val restrict_bij = result(); |
|
459 |
||
460 |
||
461 |
(*** Lemmas for Ramsey's Theorem ***) |
|
462 |
||
463 |
goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)"; |
|
464 |
by (fast_tac (ZF_cs addSEs [fun_weaken_type]) 1); |
|
465 |
val inj_weaken_type = result(); |
|
466 |
||
467 |
val [major] = goal Perm.thy |
|
468 |
"[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})";
|
|
469 |
by (rtac (major RS restrict_bij RS bij_is_inj RS inj_weaken_type) 1); |
|
470 |
by (fast_tac ZF_cs 1); |
|
471 |
by (cut_facts_tac [major] 1); |
|
472 |
by (rewtac inj_def); |
|
473 |
by (safe_tac ZF_cs); |
|
474 |
by (etac range_type 1); |
|
475 |
by (assume_tac 1); |
|
476 |
by (dtac apply_equality 1); |
|
477 |
by (assume_tac 1); |
|
| 437 | 478 |
by (res_inst_tac [("a","m")] mem_irrefl 1);
|
| 0 | 479 |
by (fast_tac ZF_cs 1); |
480 |
val inj_succ_restrict = result(); |
|
481 |
||
482 |
goalw Perm.thy [inj_def] |
|
| 37 | 483 |
"!!f. [| f: inj(A,B); a~:A; b~:B |] ==> \ |
| 0 | 484 |
\ cons(<a,b>,f) : inj(cons(a,A), cons(b,B))"; |
485 |
(*cannot use safe_tac: must preserve the implication*) |
|
486 |
by (etac CollectE 1); |
|
487 |
by (rtac CollectI 1); |
|
488 |
by (etac fun_extend 1); |
|
489 |
by (REPEAT (ares_tac [ballI] 1)); |
|
490 |
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
|
491 |
(*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
|
492 |
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
|
493 |
by (ALLGOALS (asm_simp_tac |
|
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
494 |
(FOL_ss addsimps [fun_extend_apply2,fun_extend_apply1]))); |
| 0 | 495 |
by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type]))); |
496 |
val inj_extend = result(); |