author | wenzelm |
Fri, 11 Jan 2002 14:53:30 +0100 | |
changeset 12720 | f8a134b9a57f |
parent 9907 | 473a6604da94 |
child 12891 | 92af5c3a10fb |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Perm.ML |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
735 | 6 |
The theory underlying permutation groups |
0 | 7 |
-- Composition of relations, the identity relation |
8 |
-- Injections, surjections, bijections |
|
9 |
-- Lemmas for the Schroeder-Bernstein Theorem |
|
10 |
*) |
|
11 |
||
12 |
(** Surjective function space **) |
|
13 |
||
5137 | 14 |
Goalw [surj_def] "f: surj(A,B) ==> f: A->B"; |
0 | 15 |
by (etac CollectD1 1); |
760 | 16 |
qed "surj_is_fun"; |
0 | 17 |
|
5137 | 18 |
Goalw [surj_def] "f : Pi(A,B) ==> f: surj(A,range(f))"; |
4091 | 19 |
by (blast_tac (claset() addIs [apply_equality, range_of_fun, domain_type]) 1); |
760 | 20 |
qed "fun_is_surj"; |
0 | 21 |
|
5137 | 22 |
Goalw [surj_def] "f: surj(A,B) ==> range(f)=B"; |
4091 | 23 |
by (best_tac (claset() addIs [apply_Pair] addEs [range_type]) 1); |
760 | 24 |
qed "surj_range"; |
0 | 25 |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
26 |
(** A function with a right inverse is a surjection **) |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
27 |
|
9211 | 28 |
val prems = Goalw [surj_def] |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
29 |
"[| f: A->B; !!y. y:B ==> d(y): A; !!y. y:B ==> f`d(y) = y \ |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
30 |
\ |] ==> f: surj(A,B)"; |
4091 | 31 |
by (blast_tac (claset() addIs prems) 1); |
760 | 32 |
qed "f_imp_surjective"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
33 |
|
5321 | 34 |
val prems = Goal |
1461 | 35 |
"[| !!x. x:A ==> c(x): B; \ |
36 |
\ !!y. y:B ==> d(y): A; \ |
|
37 |
\ !!y. y:B ==> c(d(y)) = y \ |
|
3840 | 38 |
\ |] ==> (lam x:A. c(x)) : surj(A,B)"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
39 |
by (res_inst_tac [("d", "d")] f_imp_surjective 1); |
5529 | 40 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [lam_type]@prems) )); |
760 | 41 |
qed "lam_surjective"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
42 |
|
735 | 43 |
(*Cantor's theorem revisited*) |
5067 | 44 |
Goalw [surj_def] "f ~: surj(A,Pow(A))"; |
4152 | 45 |
by Safe_tac; |
735 | 46 |
by (cut_facts_tac [cantor] 1); |
47 |
by (fast_tac subset_cs 1); |
|
760 | 48 |
qed "cantor_surj"; |
735 | 49 |
|
0 | 50 |
|
51 |
(** Injective function space **) |
|
52 |
||
5137 | 53 |
Goalw [inj_def] "f: inj(A,B) ==> f: A->B"; |
0 | 54 |
by (etac CollectD1 1); |
760 | 55 |
qed "inj_is_fun"; |
0 | 56 |
|
1787 | 57 |
(*Good for dealing with sets of pairs, but a bit ugly in use [used in AC]*) |
5067 | 58 |
Goalw [inj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
59 |
"[| <a,b>:f; <c,b>:f; f: inj(A,B) |] ==> a=c"; |
9173 | 60 |
by (blast_tac (claset() addDs [Pair_mem_PiD]) 1); |
760 | 61 |
qed "inj_equality"; |
0 | 62 |
|
5137 | 63 |
Goalw [inj_def] "[| f:inj(A,B); a:A; b:A; f`a=f`b |] ==> a=b"; |
3016 | 64 |
by (Blast_tac 1); |
9907 | 65 |
qed "inj_apply_equality"; |
826 | 66 |
|
484 | 67 |
(** A function with a left inverse is an injection **) |
68 |
||
5137 | 69 |
Goal "[| f: A->B; ALL x:A. d(f`x)=x |] ==> f: inj(A,B)"; |
4091 | 70 |
by (asm_simp_tac (simpset() addsimps [inj_def]) 1); |
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5529
diff
changeset
|
71 |
by (blast_tac (claset() addIs [subst_context RS box_equals]) 1); |
1787 | 72 |
bind_thm ("f_imp_injective", ballI RSN (2,result())); |
484 | 73 |
|
5321 | 74 |
val prems = Goal |
1461 | 75 |
"[| !!x. x:A ==> c(x): B; \ |
76 |
\ !!x. x:A ==> d(c(x)) = x \ |
|
3840 | 77 |
\ |] ==> (lam x:A. c(x)) : inj(A,B)"; |
484 | 78 |
by (res_inst_tac [("d", "d")] f_imp_injective 1); |
5529 | 79 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [lam_type]@prems))); |
760 | 80 |
qed "lam_injective"; |
484 | 81 |
|
82 |
(** Bijections **) |
|
0 | 83 |
|
5137 | 84 |
Goalw [bij_def] "f: bij(A,B) ==> f: inj(A,B)"; |
0 | 85 |
by (etac IntD1 1); |
760 | 86 |
qed "bij_is_inj"; |
0 | 87 |
|
5137 | 88 |
Goalw [bij_def] "f: bij(A,B) ==> f: surj(A,B)"; |
0 | 89 |
by (etac IntD2 1); |
760 | 90 |
qed "bij_is_surj"; |
0 | 91 |
|
92 |
(* f: bij(A,B) ==> f: A->B *) |
|
6153 | 93 |
bind_thm ("bij_is_fun", bij_is_inj RS inj_is_fun); |
0 | 94 |
|
9211 | 95 |
val prems = Goalw [bij_def] |
1461 | 96 |
"[| !!x. x:A ==> c(x): B; \ |
97 |
\ !!y. y:B ==> d(y): A; \ |
|
98 |
\ !!x. x:A ==> d(c(x)) = x; \ |
|
99 |
\ !!y. y:B ==> c(d(y)) = y \ |
|
3840 | 100 |
\ |] ==> (lam x:A. c(x)) : bij(A,B)"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
101 |
by (REPEAT (ares_tac (prems @ [IntI, lam_injective, lam_surjective]) 1)); |
760 | 102 |
qed "lam_bijective"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
103 |
|
5284
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
104 |
Goal "(ALL y : x. EX! y'. f(y') = f(y)) \ |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
105 |
\ ==> (lam z:{f(y). y:x}. THE y. f(y) = z) : bij({f(y). y:x}, x)"; |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
106 |
by (res_inst_tac [("d","f")] lam_bijective 1); |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
107 |
by (auto_tac (claset(), |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
108 |
simpset() addsimps [the_equality2])); |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
109 |
qed "RepFun_bijective"; |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5268
diff
changeset
|
110 |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
111 |
|
0 | 112 |
(** Identity function **) |
113 |
||
9211 | 114 |
Goalw [id_def] "a:A ==> <a,a> : id(A)"; |
115 |
by (etac lamI 1); |
|
760 | 116 |
qed "idI"; |
0 | 117 |
|
9211 | 118 |
val major::prems = Goalw [id_def] |
0 | 119 |
"[| p: id(A); !!x.[| x:A; p=<x,x> |] ==> P \ |
120 |
\ |] ==> P"; |
|
121 |
by (rtac (major RS lamE) 1); |
|
122 |
by (REPEAT (ares_tac prems 1)); |
|
760 | 123 |
qed "idE"; |
0 | 124 |
|
5067 | 125 |
Goalw [id_def] "id(A) : A->A"; |
0 | 126 |
by (rtac lam_type 1); |
127 |
by (assume_tac 1); |
|
760 | 128 |
qed "id_type"; |
0 | 129 |
|
5137 | 130 |
Goalw [id_def] "x:A ==> id(A)`x = x"; |
2469 | 131 |
by (Asm_simp_tac 1); |
132 |
qed "id_conv"; |
|
133 |
||
134 |
Addsimps [id_conv]; |
|
826 | 135 |
|
9211 | 136 |
Goalw [id_def] "A<=B ==> id(A) <= id(B)"; |
137 |
by (etac lam_mono 1); |
|
760 | 138 |
qed "id_mono"; |
0 | 139 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5137
diff
changeset
|
140 |
Goalw [inj_def,id_def] "A<=B ==> id(A): inj(A,B)"; |
0 | 141 |
by (REPEAT (ares_tac [CollectI,lam_type] 1)); |
435 | 142 |
by (etac subsetD 1 THEN assume_tac 1); |
2469 | 143 |
by (Simp_tac 1); |
760 | 144 |
qed "id_subset_inj"; |
435 | 145 |
|
9907 | 146 |
bind_thm ("id_inj", subset_refl RS id_subset_inj); |
0 | 147 |
|
5067 | 148 |
Goalw [id_def,surj_def] "id(A): surj(A,A)"; |
4091 | 149 |
by (blast_tac (claset() addIs [lam_type, beta]) 1); |
760 | 150 |
qed "id_surj"; |
0 | 151 |
|
5067 | 152 |
Goalw [bij_def] "id(A): bij(A,A)"; |
4091 | 153 |
by (blast_tac (claset() addIs [id_inj, id_surj]) 1); |
760 | 154 |
qed "id_bij"; |
0 | 155 |
|
5067 | 156 |
Goalw [id_def] "A <= B <-> id(A) : A->B"; |
4091 | 157 |
by (fast_tac (claset() addSIs [lam_type] addDs [apply_type] |
158 |
addss (simpset())) 1); |
|
760 | 159 |
qed "subset_iff_id"; |
517 | 160 |
|
0 | 161 |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
162 |
(*** Converse of a function ***) |
0 | 163 |
|
5137 | 164 |
Goalw [inj_def] "f: inj(A,B) ==> converse(f) : range(f)->A"; |
4091 | 165 |
by (asm_simp_tac (simpset() addsimps [Pi_iff, function_def]) 1); |
2033 | 166 |
by (etac CollectE 1); |
4091 | 167 |
by (asm_simp_tac (simpset() addsimps [apply_iff]) 1); |
168 |
by (blast_tac (claset() addDs [fun_is_rel]) 1); |
|
760 | 169 |
qed "inj_converse_fun"; |
0 | 170 |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
171 |
(** Equations for converse(f) **) |
0 | 172 |
|
173 |
(*The premises are equivalent to saying that f is injective...*) |
|
5268 | 174 |
Goal "[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a"; |
4091 | 175 |
by (blast_tac (claset() addIs [apply_Pair, apply_equality, converseI]) 1); |
760 | 176 |
qed "left_inverse_lemma"; |
0 | 177 |
|
5268 | 178 |
Goal "[| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a"; |
4091 | 179 |
by (blast_tac (claset() addIs [left_inverse_lemma, inj_converse_fun, |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
180 |
inj_is_fun]) 1); |
760 | 181 |
qed "left_inverse"; |
0 | 182 |
|
9907 | 183 |
bind_thm ("left_inverse_bij", bij_is_inj RS left_inverse); |
435 | 184 |
|
5321 | 185 |
Goal "[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b"; |
0 | 186 |
by (rtac (apply_Pair RS (converseD RS apply_equality)) 1); |
5321 | 187 |
by (REPEAT (assume_tac 1)); |
760 | 188 |
qed "right_inverse_lemma"; |
0 | 189 |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
190 |
(*Should the premises be f:surj(A,B), b:B for symmetry with left_inverse? |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
191 |
No: they would not imply that converse(f) was a function! *) |
5137 | 192 |
Goal "[| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b"; |
0 | 193 |
by (rtac right_inverse_lemma 1); |
435 | 194 |
by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1)); |
760 | 195 |
qed "right_inverse"; |
0 | 196 |
|
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
197 |
Addsimps [left_inverse, right_inverse]; |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
198 |
|
2469 | 199 |
|
5137 | 200 |
Goal "[| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b"; |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
201 |
by (force_tac (claset(), simpset() addsimps [bij_def, surj_range]) 1); |
760 | 202 |
qed "right_inverse_bij"; |
435 | 203 |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
204 |
(** Converses of injections, surjections, bijections **) |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
205 |
|
5137 | 206 |
Goal "f: inj(A,B) ==> converse(f): inj(range(f), A)"; |
1461 | 207 |
by (rtac f_imp_injective 1); |
208 |
by (etac inj_converse_fun 1); |
|
209 |
by (rtac right_inverse 1); |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
210 |
by (REPEAT (assume_tac 1)); |
760 | 211 |
qed "inj_converse_inj"; |
0 | 212 |
|
5137 | 213 |
Goal "f: inj(A,B) ==> converse(f): surj(range(f), A)"; |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
214 |
by (blast_tac (claset() addIs [f_imp_surjective, inj_converse_fun, |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
215 |
left_inverse, |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
216 |
inj_is_fun, range_of_fun RS apply_type]) 1); |
760 | 217 |
qed "inj_converse_surj"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
218 |
|
5137 | 219 |
Goalw [bij_def] "f: bij(A,B) ==> converse(f): bij(B,A)"; |
4091 | 220 |
by (fast_tac (claset() addEs [surj_range RS subst, inj_converse_inj, |
3016 | 221 |
inj_converse_surj]) 1); |
760 | 222 |
qed "bij_converse_bij"; |
2469 | 223 |
(*Adding this as an SI seems to cause looping*) |
0 | 224 |
|
6153 | 225 |
AddTCs [bij_converse_bij]; |
226 |
||
0 | 227 |
|
228 |
(** Composition of two relations **) |
|
229 |
||
791 | 230 |
(*The inductive definition package could derive these theorems for (r O s)*) |
0 | 231 |
|
5137 | 232 |
Goalw [comp_def] "[| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s"; |
3016 | 233 |
by (Blast_tac 1); |
760 | 234 |
qed "compI"; |
0 | 235 |
|
9211 | 236 |
val prems = Goalw [comp_def] |
0 | 237 |
"[| xz : r O s; \ |
238 |
\ !!x y z. [| xz=<x,z>; <x,y>:s; <y,z>:r |] ==> P \ |
|
239 |
\ |] ==> P"; |
|
240 |
by (cut_facts_tac prems 1); |
|
241 |
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); |
|
760 | 242 |
qed "compE"; |
0 | 243 |
|
2688 | 244 |
bind_thm ("compEpair", |
0 | 245 |
rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) |
1461 | 246 |
THEN prune_params_tac) |
8551 | 247 |
(inst "xz" "<a,c>" compE)); |
0 | 248 |
|
2469 | 249 |
AddSIs [idI]; |
250 |
AddIs [compI]; |
|
251 |
AddSEs [compE,idE]; |
|
0 | 252 |
|
5202 | 253 |
Goal "converse(R O S) = converse(S) O converse(R)"; |
254 |
by (Blast_tac 1); |
|
255 |
qed "converse_comp"; |
|
256 |
||
257 |
||
0 | 258 |
(** Domain and Range -- see Suppes, section 3.1 **) |
259 |
||
260 |
(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*) |
|
5067 | 261 |
Goal "range(r O s) <= range(r)"; |
3016 | 262 |
by (Blast_tac 1); |
760 | 263 |
qed "range_comp"; |
0 | 264 |
|
5137 | 265 |
Goal "domain(r) <= range(s) ==> range(r O s) = range(r)"; |
0 | 266 |
by (rtac (range_comp RS equalityI) 1); |
3016 | 267 |
by (Blast_tac 1); |
760 | 268 |
qed "range_comp_eq"; |
0 | 269 |
|
5067 | 270 |
Goal "domain(r O s) <= domain(s)"; |
3016 | 271 |
by (Blast_tac 1); |
760 | 272 |
qed "domain_comp"; |
0 | 273 |
|
5137 | 274 |
Goal "range(s) <= domain(r) ==> domain(r O s) = domain(s)"; |
0 | 275 |
by (rtac (domain_comp RS equalityI) 1); |
3016 | 276 |
by (Blast_tac 1); |
760 | 277 |
qed "domain_comp_eq"; |
0 | 278 |
|
5067 | 279 |
Goal "(r O s)``A = r``(s``A)"; |
3016 | 280 |
by (Blast_tac 1); |
760 | 281 |
qed "image_comp"; |
218 | 282 |
|
283 |
||
0 | 284 |
(** Other results **) |
285 |
||
5137 | 286 |
Goal "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; |
3016 | 287 |
by (Blast_tac 1); |
760 | 288 |
qed "comp_mono"; |
0 | 289 |
|
290 |
(*composition preserves relations*) |
|
5137 | 291 |
Goal "[| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"; |
3016 | 292 |
by (Blast_tac 1); |
760 | 293 |
qed "comp_rel"; |
0 | 294 |
|
295 |
(*associative law for composition*) |
|
5067 | 296 |
Goal "(r O s) O t = r O (s O t)"; |
3016 | 297 |
by (Blast_tac 1); |
760 | 298 |
qed "comp_assoc"; |
0 | 299 |
|
300 |
(*left identity of composition; provable inclusions are |
|
301 |
id(A) O r <= r |
|
302 |
and [| r<=A*B; B<=C |] ==> r <= id(C) O r *) |
|
5137 | 303 |
Goal "r<=A*B ==> id(B) O r = r"; |
3016 | 304 |
by (Blast_tac 1); |
760 | 305 |
qed "left_comp_id"; |
0 | 306 |
|
307 |
(*right identity of composition; provable inclusions are |
|
308 |
r O id(A) <= r |
|
309 |
and [| r<=A*B; A<=C |] ==> r <= r O id(C) *) |
|
5137 | 310 |
Goal "r<=A*B ==> r O id(A) = r"; |
3016 | 311 |
by (Blast_tac 1); |
760 | 312 |
qed "right_comp_id"; |
0 | 313 |
|
314 |
||
315 |
(** Composition preserves functions, injections, and surjections **) |
|
316 |
||
5067 | 317 |
Goalw [function_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
318 |
"[| function(g); function(f) |] ==> function(f O g)"; |
3016 | 319 |
by (Blast_tac 1); |
760 | 320 |
qed "comp_function"; |
693
b89939545725
ZF/Perm/inj_converse_fun: tidied; removed uses of PiI/E
lcp
parents:
517
diff
changeset
|
321 |
|
5137 | 322 |
Goal "[| g: A->B; f: B->C |] ==> (f O g) : A->C"; |
1787 | 323 |
by (asm_full_simp_tac |
4091 | 324 |
(simpset() addsimps [Pi_def, comp_function, Pow_iff, comp_rel] |
1787 | 325 |
setloop etac conjE) 1); |
2033 | 326 |
by (stac (range_rel_subset RS domain_comp_eq) 1 THEN assume_tac 2); |
3016 | 327 |
by (Blast_tac 1); |
760 | 328 |
qed "comp_fun"; |
0 | 329 |
|
5137 | 330 |
Goal "[| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; |
435 | 331 |
by (REPEAT (ares_tac [comp_fun,apply_equality,compI, |
1461 | 332 |
apply_Pair,apply_type] 1)); |
760 | 333 |
qed "comp_fun_apply"; |
0 | 334 |
|
2469 | 335 |
Addsimps [comp_fun_apply]; |
336 |
||
862 | 337 |
(*Simplifies compositions of lambda-abstractions*) |
5321 | 338 |
val [prem] = Goal |
1461 | 339 |
"[| !!x. x:A ==> b(x): B \ |
3840 | 340 |
\ |] ==> (lam y:B. c(y)) O (lam x:A. b(x)) = (lam x:A. c(b(x)))"; |
1461 | 341 |
by (rtac fun_extension 1); |
342 |
by (rtac comp_fun 1); |
|
343 |
by (rtac lam_funtype 2); |
|
6153 | 344 |
by (typecheck_tac (tcset() addTCs [prem])); |
4091 | 345 |
by (asm_simp_tac (simpset() |
7570 | 346 |
setSolver (mk_solver "" |
347 |
(type_solver_tac (tcset() addTCs [prem, lam_funtype])))) 1); |
|
862 | 348 |
qed "comp_lam"; |
349 |
||
5137 | 350 |
Goal "[| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)"; |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
351 |
by (res_inst_tac [("d", "%y. converse(g) ` (converse(f) ` y)")] |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
352 |
f_imp_injective 1); |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
353 |
by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1)); |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
354 |
by (asm_simp_tac (simpset() |
7570 | 355 |
setSolver (mk_solver "" |
356 |
(type_solver_tac (tcset() addTCs [inj_is_fun])))) 1); |
|
760 | 357 |
qed "comp_inj"; |
0 | 358 |
|
5067 | 359 |
Goalw [surj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
360 |
"[| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)"; |
4091 | 361 |
by (blast_tac (claset() addSIs [comp_fun,comp_fun_apply]) 1); |
760 | 362 |
qed "comp_surj"; |
0 | 363 |
|
5067 | 364 |
Goalw [bij_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
365 |
"[| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)"; |
4091 | 366 |
by (blast_tac (claset() addIs [comp_inj,comp_surj]) 1); |
760 | 367 |
qed "comp_bij"; |
0 | 368 |
|
369 |
||
370 |
(** Dual properties of inj and surj -- useful for proofs from |
|
371 |
D Pastre. Automatic theorem proving in set theory. |
|
372 |
Artificial Intelligence, 10:1--27, 1978. **) |
|
373 |
||
5067 | 374 |
Goalw [inj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
375 |
"[| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)"; |
4152 | 376 |
by Safe_tac; |
0 | 377 |
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); |
4091 | 378 |
by (asm_simp_tac (simpset() ) 1); |
760 | 379 |
qed "comp_mem_injD1"; |
0 | 380 |
|
5067 | 381 |
Goalw [inj_def,surj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
382 |
"[| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)"; |
4152 | 383 |
by Safe_tac; |
0 | 384 |
by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1); |
385 |
by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3); |
|
386 |
by (REPEAT (assume_tac 1)); |
|
4152 | 387 |
by Safe_tac; |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
388 |
by (res_inst_tac [("t", "op `(g)")] subst_context 1); |
0 | 389 |
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); |
4091 | 390 |
by (asm_simp_tac (simpset() ) 1); |
760 | 391 |
qed "comp_mem_injD2"; |
0 | 392 |
|
5067 | 393 |
Goalw [surj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
394 |
"[| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)"; |
4091 | 395 |
by (blast_tac (claset() addSIs [comp_fun_apply RS sym, apply_funtype]) 1); |
760 | 396 |
qed "comp_mem_surjD1"; |
0 | 397 |
|
5268 | 398 |
Goal "[| (f O g)`a = c; g: A->B; f: B->C; a:A |] ==> f`(g`a) = c"; |
435 | 399 |
by (REPEAT (ares_tac [comp_fun_apply RS sym RS trans] 1)); |
760 | 400 |
qed "comp_fun_applyD"; |
0 | 401 |
|
5067 | 402 |
Goalw [inj_def,surj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
403 |
"[| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)"; |
4152 | 404 |
by Safe_tac; |
0 | 405 |
by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1); |
435 | 406 |
by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_fun_applyD 1)); |
4091 | 407 |
by (blast_tac (claset() addSIs [apply_funtype]) 1); |
760 | 408 |
qed "comp_mem_surjD2"; |
0 | 409 |
|
410 |
||
411 |
(** inverses of composition **) |
|
412 |
||
413 |
(*left inverse of composition; one inclusion is |
|
414 |
f: A->B ==> id(A) <= converse(f) O f *) |
|
5137 | 415 |
Goalw [inj_def] "f: inj(A,B) ==> converse(f) O f = id(A)"; |
4091 | 416 |
by (fast_tac (claset() addIs [apply_Pair] |
1787 | 417 |
addEs [domain_type] |
4091 | 418 |
addss (simpset() addsimps [apply_iff])) 1); |
760 | 419 |
qed "left_comp_inverse"; |
0 | 420 |
|
421 |
(*right inverse of composition; one inclusion is |
|
1461 | 422 |
f: A->B ==> f O converse(f) <= id(B) |
735 | 423 |
*) |
9907 | 424 |
val [prem] = goalw (the_context ()) [surj_def] |
0 | 425 |
"f: surj(A,B) ==> f O converse(f) = id(B)"; |
426 |
val appfD = (prem RS CollectD1) RSN (3,apply_equality2); |
|
427 |
by (cut_facts_tac [prem] 1); |
|
428 |
by (rtac equalityI 1); |
|
4091 | 429 |
by (best_tac (claset() addEs [domain_type, range_type, make_elim appfD]) 1); |
430 |
by (blast_tac (claset() addIs [apply_Pair]) 1); |
|
760 | 431 |
qed "right_comp_inverse"; |
0 | 432 |
|
435 | 433 |
(** Proving that a function is a bijection **) |
434 |
||
5067 | 435 |
Goalw [id_def] |
7379 | 436 |
"[| f: A->B; g: B->A |] ==> f O g = id(B) <-> (ALL y:B. f`(g`y)=y)"; |
4152 | 437 |
by Safe_tac; |
3840 | 438 |
by (dres_inst_tac [("t", "%h. h`y ")] subst_context 1); |
2469 | 439 |
by (Asm_full_simp_tac 1); |
437 | 440 |
by (rtac fun_extension 1); |
435 | 441 |
by (REPEAT (ares_tac [comp_fun, lam_type] 1)); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4152
diff
changeset
|
442 |
by Auto_tac; |
760 | 443 |
qed "comp_eq_id_iff"; |
435 | 444 |
|
5067 | 445 |
Goalw [bij_def] |
7379 | 446 |
"[| f: A->B; g: B->A; f O g = id(B); g O f = id(A) |] ==> f : bij(A,B)"; |
4091 | 447 |
by (asm_full_simp_tac (simpset() addsimps [comp_eq_id_iff]) 1); |
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
448 |
by (REPEAT (ares_tac [conjI, f_imp_injective, f_imp_surjective] 1 |
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
449 |
ORELSE eresolve_tac [bspec, apply_type] 1)); |
760 | 450 |
qed "fg_imp_bijective"; |
435 | 451 |
|
5137 | 452 |
Goal "[| f: A->A; f O f = id(A) |] ==> f : bij(A,A)"; |
435 | 453 |
by (REPEAT (ares_tac [fg_imp_bijective] 1)); |
760 | 454 |
qed "nilpotent_imp_bijective"; |
435 | 455 |
|
5137 | 456 |
Goal "[| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; |
4091 | 457 |
by (asm_simp_tac (simpset() addsimps [fg_imp_bijective, comp_eq_id_iff, |
1461 | 458 |
left_inverse_lemma, right_inverse_lemma]) 1); |
760 | 459 |
qed "invertible_imp_bijective"; |
0 | 460 |
|
461 |
(** Unions of functions -- cf similar theorems on func.ML **) |
|
462 |
||
1709 | 463 |
(*Theorem by KG, proof by LCP*) |
5466 | 464 |
Goal "[| f: inj(A,B); g: inj(C,D); B Int D = 0 |] \ |
6068 | 465 |
\ ==> (lam a: A Un C. if a:A then f`a else g`a) : inj(A Un C, B Un D)"; |
466 |
by (res_inst_tac [("d","%z. if z:B then converse(f)`z else converse(g)`z")] |
|
1709 | 467 |
lam_injective 1); |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
468 |
by (auto_tac (claset(), |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
469 |
simpset() addsimps [inj_is_fun RS apply_type])); |
5466 | 470 |
by (blast_tac (claset() addIs [inj_is_fun RS apply_type]) 1); |
1709 | 471 |
qed "inj_disjoint_Un"; |
1610 | 472 |
|
5067 | 473 |
Goalw [surj_def] |
7379 | 474 |
"[| f: surj(A,B); g: surj(C,D); A Int C = 0 |] \ |
475 |
\ ==> (f Un g) : surj(A Un C, B Un D)"; |
|
4091 | 476 |
by (blast_tac (claset() addIs [fun_disjoint_apply1, fun_disjoint_apply2, |
3016 | 477 |
fun_disjoint_Un, trans]) 1); |
760 | 478 |
qed "surj_disjoint_Un"; |
0 | 479 |
|
480 |
(*A simple, high-level proof; the version for injections follows from it, |
|
502
77e36960fd9e
ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents:
484
diff
changeset
|
481 |
using f:inj(A,B) <-> f:bij(A,range(f)) *) |
7379 | 482 |
Goal "[| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] \ |
483 |
\ ==> (f Un g) : bij(A Un C, B Un D)"; |
|
0 | 484 |
by (rtac invertible_imp_bijective 1); |
2033 | 485 |
by (stac converse_Un 1); |
0 | 486 |
by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1)); |
760 | 487 |
qed "bij_disjoint_Un"; |
0 | 488 |
|
489 |
||
490 |
(** Restrictions as surjections and bijections *) |
|
491 |
||
9907 | 492 |
val prems = goalw (the_context ()) [surj_def] |
0 | 493 |
"f: Pi(A,B) ==> f: surj(A, f``A)"; |
494 |
val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]); |
|
4091 | 495 |
by (fast_tac (claset() addIs rls) 1); |
760 | 496 |
qed "surj_image"; |
0 | 497 |
|
5137 | 498 |
Goal "[| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A"; |
0 | 499 |
by (rtac equalityI 1); |
500 |
by (SELECT_GOAL (rewtac restrict_def) 2); |
|
501 |
by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2 |
|
502 |
ORELSE ares_tac [subsetI,lamI,imageI] 2)); |
|
503 |
by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1)); |
|
760 | 504 |
qed "restrict_image"; |
0 | 505 |
|
5067 | 506 |
Goalw [inj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
507 |
"[| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)"; |
4091 | 508 |
by (safe_tac (claset() addSEs [restrict_type2])); |
0 | 509 |
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD, |
510 |
box_equals, restrict] 1)); |
|
760 | 511 |
qed "restrict_inj"; |
0 | 512 |
|
5321 | 513 |
Goal "[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)"; |
0 | 514 |
by (rtac (restrict_image RS subst) 1); |
515 |
by (rtac (restrict_type2 RS surj_image) 3); |
|
5321 | 516 |
by (REPEAT (assume_tac 1)); |
760 | 517 |
qed "restrict_surj"; |
0 | 518 |
|
5067 | 519 |
Goalw [inj_def,bij_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
520 |
"[| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)"; |
4091 | 521 |
by (blast_tac (claset() addSIs [restrict, restrict_surj] |
3016 | 522 |
addIs [box_equals, surj_is_fun]) 1); |
760 | 523 |
qed "restrict_bij"; |
0 | 524 |
|
525 |
||
526 |
(*** Lemmas for Ramsey's Theorem ***) |
|
527 |
||
5137 | 528 |
Goalw [inj_def] "[| f: inj(A,B); B<=D |] ==> f: inj(A,D)"; |
4091 | 529 |
by (blast_tac (claset() addIs [fun_weaken_type]) 1); |
760 | 530 |
qed "inj_weaken_type"; |
0 | 531 |
|
9180 | 532 |
Goal "[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})"; |
533 |
by (rtac (restrict_bij RS bij_is_inj RS inj_weaken_type) 1); |
|
534 |
by (assume_tac 1); |
|
3016 | 535 |
by (Blast_tac 1); |
0 | 536 |
by (rewtac inj_def); |
4091 | 537 |
by (fast_tac (claset() addEs [range_type, mem_irrefl] |
9180 | 538 |
addDs [apply_equality]) 1); |
760 | 539 |
qed "inj_succ_restrict"; |
0 | 540 |
|
5067 | 541 |
Goalw [inj_def] |
7379 | 542 |
"[| f: inj(A,B); a~:A; b~:B |] \ |
543 |
\ ==> cons(<a,b>,f) : inj(cons(a,A), cons(b,B))"; |
|
5525 | 544 |
by (force_tac (claset() addIs [apply_type], |
545 |
simpset() addsimps [fun_extend, fun_extend_apply2, |
|
546 |
fun_extend_apply1]) 1); |
|
760 | 547 |
qed "inj_extend"; |
1787 | 548 |