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