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