author | wenzelm |
Mon, 21 Jan 2002 15:29:06 +0100 | |
changeset 12828 | 57fb9d1ee34a |
parent 9211 | 6236c5285bd8 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/domrange |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
6 |
Converse, domain, range of a relation or function |
|
7 |
*) |
|
8 |
||
9 |
(*** converse ***) |
|
10 |
||
9211 | 11 |
Goalw [converse_def] "<a,b>: converse(r) <-> <b,a>:r"; |
12 |
by (Blast_tac 1) ; |
|
13 |
qed "converse_iff"; |
|
687 | 14 |
|
9211 | 15 |
Goalw [converse_def] "<a,b>:r ==> <b,a>:converse(r)"; |
16 |
by (Blast_tac 1) ; |
|
17 |
qed "converseI"; |
|
0 | 18 |
|
9211 | 19 |
Goalw [converse_def] "<a,b> : converse(r) ==> <b,a> : r"; |
20 |
by (Blast_tac 1) ; |
|
21 |
qed "converseD"; |
|
0 | 22 |
|
9211 | 23 |
val [major,minor]= Goalw [converse_def] |
0 | 24 |
"[| yx : converse(r); \ |
25 |
\ !!x y. [| yx=<y,x>; <x,y>:r |] ==> P \ |
|
9211 | 26 |
\ |] ==> P"; |
27 |
by (rtac (major RS ReplaceE) 1); |
|
28 |
by (REPEAT (eresolve_tac [exE, conjE, minor] 1)); |
|
29 |
by (hyp_subst_tac 1); |
|
30 |
by (assume_tac 1) ; |
|
31 |
qed "converseE"; |
|
0 | 32 |
|
2469 | 33 |
Addsimps [converse_iff]; |
34 |
AddSIs [converseI]; |
|
35 |
AddSEs [converseD,converseE]; |
|
0 | 36 |
|
9180 | 37 |
Goal "r<=Sigma(A,B) ==> converse(converse(r)) = r"; |
38 |
by (Blast_tac 1) ; |
|
39 |
qed "converse_converse"; |
|
40 |
||
41 |
Goal "r<=A*B ==> converse(r)<=B*A"; |
|
42 |
by (Blast_tac 1) ; |
|
43 |
qed "converse_type"; |
|
0 | 44 |
|
9180 | 45 |
Goal "converse(A*B) = B*A"; |
46 |
by (Blast_tac 1) ; |
|
47 |
qed "converse_prod"; |
|
0 | 48 |
|
9180 | 49 |
Goal "converse(0) = 0"; |
50 |
by (Blast_tac 1) ; |
|
51 |
qed "converse_empty"; |
|
2469 | 52 |
|
53 |
Addsimps [converse_prod, converse_empty]; |
|
0 | 54 |
|
9211 | 55 |
Goal "A <= Sigma(X,Y) ==> converse(A) <= converse(B) <-> A <= B"; |
56 |
by (Blast_tac 1) ; |
|
57 |
qed "converse_subset_iff"; |
|
5202 | 58 |
|
59 |
||
0 | 60 |
(*** domain ***) |
61 |
||
9211 | 62 |
Goalw [domain_def] "a: domain(r) <-> (EX y. <a,y>: r)"; |
63 |
by (Blast_tac 1) ; |
|
64 |
qed "domain_iff"; |
|
0 | 65 |
|
9180 | 66 |
Goal "<a,b>: r ==> a: domain(r)"; |
67 |
by (etac (exI RS (domain_iff RS iffD2)) 1) ; |
|
68 |
qed "domainI"; |
|
0 | 69 |
|
9180 | 70 |
val prems= Goal |
71 |
"[| a : domain(r); !!y. <a,y>: r ==> P |] ==> P"; |
|
72 |
by (rtac (domain_iff RS iffD1 RS exE) 1); |
|
73 |
by (REPEAT (ares_tac prems 1)) ; |
|
74 |
qed "domainE"; |
|
0 | 75 |
|
2469 | 76 |
AddIs [domainI]; |
77 |
AddSEs [domainE]; |
|
0 | 78 |
|
9180 | 79 |
Goal "domain(Sigma(A,B)) <= A"; |
80 |
by (Blast_tac 1) ; |
|
81 |
qed "domain_subset"; |
|
0 | 82 |
|
83 |
(*** range ***) |
|
84 |
||
9211 | 85 |
Goalw [range_def] "<a,b>: r ==> b : range(r)"; |
86 |
by (etac (converseI RS domainI) 1) ; |
|
87 |
qed "rangeI"; |
|
0 | 88 |
|
9211 | 89 |
val major::prems= Goalw [range_def] |
90 |
"[| b : range(r); !!x. <x,b>: r ==> P |] ==> P"; |
|
91 |
by (rtac (major RS domainE) 1); |
|
92 |
by (resolve_tac prems 1); |
|
93 |
by (etac converseD 1) ; |
|
94 |
qed "rangeE"; |
|
0 | 95 |
|
2469 | 96 |
AddIs [rangeI]; |
97 |
AddSEs [rangeE]; |
|
98 |
||
9211 | 99 |
Goalw [range_def] "range(A*B) <= B"; |
100 |
by (stac converse_prod 1); |
|
101 |
by (rtac domain_subset 1) ; |
|
102 |
qed "range_subset"; |
|
0 | 103 |
|
104 |
||
105 |
(*** field ***) |
|
106 |
||
9211 | 107 |
Goalw [field_def] "<a,b>: r ==> a : field(r)"; |
108 |
by (Blast_tac 1) ; |
|
109 |
qed "fieldI1"; |
|
0 | 110 |
|
9211 | 111 |
Goalw [field_def] "<a,b>: r ==> b : field(r)"; |
112 |
by (Blast_tac 1) ; |
|
113 |
qed "fieldI2"; |
|
0 | 114 |
|
9211 | 115 |
val [prem]= Goalw [field_def] |
116 |
"(~ <c,a>:r ==> <a,b>: r) ==> a : field(r)"; |
|
117 |
by (blast_tac (claset() addIs [prem]) 1) ; |
|
118 |
qed "fieldCI"; |
|
119 |
||
120 |
val major::prems= Goalw [field_def] |
|
0 | 121 |
"[| a : field(r); \ |
122 |
\ !!x. <a,x>: r ==> P; \ |
|
9211 | 123 |
\ !!x. <x,a>: r ==> P |] ==> P"; |
124 |
by (rtac (major RS UnE) 1); |
|
125 |
by (REPEAT (eresolve_tac (prems@[domainE,rangeE]) 1)) ; |
|
126 |
qed "fieldE"; |
|
0 | 127 |
|
2469 | 128 |
AddIs [fieldCI]; |
129 |
AddSEs [fieldE]; |
|
0 | 130 |
|
9180 | 131 |
Goal "field(A*B) <= A Un B"; |
132 |
by (Blast_tac 1) ; |
|
133 |
qed "field_subset"; |
|
2469 | 134 |
|
9211 | 135 |
Goalw [field_def] "domain(r) <= field(r)"; |
136 |
by (rtac Un_upper1 1) ; |
|
137 |
qed "domain_subset_field"; |
|
0 | 138 |
|
9211 | 139 |
Goalw [field_def] "range(r) <= field(r)"; |
140 |
by (rtac Un_upper2 1) ; |
|
141 |
qed "range_subset_field"; |
|
0 | 142 |
|
9180 | 143 |
Goal "r <= Sigma(A,B) ==> r <= domain(r)*range(r)"; |
144 |
by (Blast_tac 1) ; |
|
145 |
qed "domain_times_range"; |
|
0 | 146 |
|
9180 | 147 |
Goal "r <= Sigma(A,B) ==> r <= field(r)*field(r)"; |
148 |
by (Blast_tac 1) ; |
|
149 |
qed "field_times_field"; |
|
0 | 150 |
|
151 |
||
152 |
(*** Image of a set under a function/relation ***) |
|
153 |
||
9180 | 154 |
Goalw [image_def] "b : r``A <-> (EX x:A. <x,b>:r)"; |
155 |
by (Blast_tac 1); |
|
156 |
qed "image_iff"; |
|
0 | 157 |
|
9180 | 158 |
Goal "b : r``{a} <-> <a,b>:r"; |
159 |
by (rtac (image_iff RS iff_trans) 1); |
|
160 |
by (Blast_tac 1) ; |
|
161 |
qed "image_singleton_iff"; |
|
0 | 162 |
|
9211 | 163 |
Goalw [image_def] "[| <a,b>: r; a:A |] ==> b : r``A"; |
164 |
by (Blast_tac 1) ; |
|
165 |
qed "imageI"; |
|
0 | 166 |
|
9211 | 167 |
val major::prems= Goalw [image_def] |
168 |
"[| b: r``A; !!x.[| <x,b>: r; x:A |] ==> P |] ==> P"; |
|
169 |
by (rtac (major RS CollectE) 1); |
|
170 |
by (REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ; |
|
171 |
qed "imageE"; |
|
0 | 172 |
|
2469 | 173 |
AddIs [imageI]; |
174 |
AddSEs [imageE]; |
|
175 |
||
9180 | 176 |
Goal "r <= A*B ==> r``C <= B"; |
177 |
by (Blast_tac 1) ; |
|
178 |
qed "image_subset"; |
|
0 | 179 |
|
180 |
||
181 |
(*** Inverse image of a set under a function/relation ***) |
|
182 |
||
9211 | 183 |
Goalw [vimage_def,image_def,converse_def] |
184 |
"a : r-``B <-> (EX y:B. <a,y>:r)"; |
|
185 |
by (Blast_tac 1) ; |
|
186 |
qed "vimage_iff"; |
|
0 | 187 |
|
9180 | 188 |
Goal "a : r-``{b} <-> <a,b>:r"; |
189 |
by (rtac (vimage_iff RS iff_trans) 1); |
|
190 |
by (Blast_tac 1) ; |
|
191 |
qed "vimage_singleton_iff"; |
|
0 | 192 |
|
9211 | 193 |
Goalw [vimage_def] "[| <a,b>: r; b:B |] ==> a : r-``B"; |
194 |
by (Blast_tac 1) ; |
|
195 |
qed "vimageI"; |
|
0 | 196 |
|
9211 | 197 |
val major::prems= Goalw [vimage_def] |
198 |
"[| a: r-``B; !!x.[| <a,x>: r; x:B |] ==> P |] ==> P"; |
|
199 |
by (rtac (major RS imageE) 1); |
|
200 |
by (REPEAT (etac converseD 1 ORELSE ares_tac prems 1)) ; |
|
201 |
qed "vimageE"; |
|
0 | 202 |
|
9211 | 203 |
Goalw [vimage_def] "r <= A*B ==> r-``C <= A"; |
204 |
by (etac (converse_type RS image_subset) 1) ; |
|
205 |
qed "vimage_subset"; |
|
0 | 206 |
|
207 |
||
208 |
(** Theorem-proving for ZF set theory **) |
|
209 |
||
2469 | 210 |
AddIs [vimageI]; |
211 |
AddSEs [vimageE]; |
|
212 |
||
4091 | 213 |
val ZF_cs = claset() delrules [equalityI]; |
0 | 214 |
|
215 |
(** The Union of a set of relations is a relation -- Lemma for fun_Union **) |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5202
diff
changeset
|
216 |
Goal "(ALL x:S. EX A B. x <= A*B) ==> \ |
0 | 217 |
\ Union(S) <= domain(Union(S)) * range(Union(S))"; |
2877 | 218 |
by (Blast_tac 1); |
760 | 219 |
qed "rel_Union"; |
0 | 220 |
|
221 |
(** The Union of 2 relations is a relation (Lemma for fun_Un) **) |
|
9180 | 222 |
Goal "[| r <= A*B; s <= C*D |] ==> (r Un s) <= (A Un C) * (B Un D)"; |
223 |
by (Blast_tac 1) ; |
|
224 |
qed "rel_Un"; |
|
0 | 225 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5202
diff
changeset
|
226 |
Goal "[| <a,c> : r; c~=b |] ==> domain(r-{<a,b>}) = domain(r)"; |
2877 | 227 |
by (Blast_tac 1); |
2469 | 228 |
qed "domain_Diff_eq"; |
229 |
||
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5202
diff
changeset
|
230 |
Goal "[| <c,b> : r; c~=a |] ==> range(r-{<a,b>}) = range(r)"; |
2877 | 231 |
by (Blast_tac 1); |
2469 | 232 |
qed "range_Diff_eq"; |
233 |