author | schirmer |
Wed, 19 Dec 2007 16:32:12 +0100 | |
changeset 25705 | 45a2ffc5911e |
parent 24893 | b8ef7afe3a6b |
child 35762 | af3ff2ba4c54 |
permissions | -rw-r--r-- |
13165 | 1 |
(* Title: ZF/equalities |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
*) |
|
7 |
||
13356 | 8 |
header{*Basic Equalities and Inclusions*} |
9 |
||
16417 | 10 |
theory equalities imports pair begin |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
11 |
|
13356 | 12 |
text{*These cover union, intersection, converse, domain, range, etc. Philippe |
13 |
de Groote proved many of the inclusions.*} |
|
14 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
15 |
lemma in_mono: "A\<subseteq>B ==> x\<in>A --> x\<in>B" |
13259 | 16 |
by blast |
17 |
||
13174 | 18 |
lemma the_eq_0 [simp]: "(THE x. False) = 0" |
19 |
by (blast intro: the_0) |
|
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
20 |
|
13356 | 21 |
subsection{*Bounded Quantifiers*} |
22 |
text {* \medskip |
|
13178 | 23 |
|
24 |
The following are not added to the default simpset because |
|
13356 | 25 |
(a) they duplicate the body and (b) there are no similar rules for @{text Int}.*} |
13178 | 26 |
|
27 |
lemma ball_Un: "(\<forall>x \<in> A\<union>B. P(x)) <-> (\<forall>x \<in> A. P(x)) & (\<forall>x \<in> B. P(x))"; |
|
28 |
by blast |
|
29 |
||
30 |
lemma bex_Un: "(\<exists>x \<in> A\<union>B. P(x)) <-> (\<exists>x \<in> A. P(x)) | (\<exists>x \<in> B. P(x))"; |
|
31 |
by blast |
|
32 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
33 |
lemma ball_UN: "(\<forall>z \<in> (\<Union>x\<in>A. B(x)). P(z)) <-> (\<forall>x\<in>A. \<forall>z \<in> B(x). P(z))" |
13178 | 34 |
by blast |
35 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
36 |
lemma bex_UN: "(\<exists>z \<in> (\<Union>x\<in>A. B(x)). P(z)) <-> (\<exists>x\<in>A. \<exists>z\<in>B(x). P(z))" |
13178 | 37 |
by blast |
38 |
||
13356 | 39 |
subsection{*Converse of a Relation*} |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
40 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
41 |
lemma converse_iff [simp]: "<a,b>\<in> converse(r) <-> <b,a>\<in>r" |
13169 | 42 |
by (unfold converse_def, blast) |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
43 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
44 |
lemma converseI [intro!]: "<a,b>\<in>r ==> <b,a>\<in>converse(r)" |
13169 | 45 |
by (unfold converse_def, blast) |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
46 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
47 |
lemma converseD: "<a,b> \<in> converse(r) ==> <b,a> \<in> r" |
13169 | 48 |
by (unfold converse_def, blast) |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
49 |
|
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
50 |
lemma converseE [elim!]: |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
51 |
"[| yx \<in> converse(r); |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
52 |
!!x y. [| yx=<y,x>; <x,y>\<in>r |] ==> P |] |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
53 |
==> P" |
13174 | 54 |
by (unfold converse_def, blast) |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
55 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
56 |
lemma converse_converse: "r\<subseteq>Sigma(A,B) ==> converse(converse(r)) = r" |
13169 | 57 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
58 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
59 |
lemma converse_type: "r\<subseteq>A*B ==> converse(r)\<subseteq>B*A" |
13169 | 60 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
61 |
|
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
62 |
lemma converse_prod [simp]: "converse(A*B) = B*A" |
13169 | 63 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
64 |
|
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
65 |
lemma converse_empty [simp]: "converse(0) = 0" |
13169 | 66 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
67 |
|
13174 | 68 |
lemma converse_subset_iff: |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
69 |
"A \<subseteq> Sigma(X,Y) ==> converse(A) \<subseteq> converse(B) <-> A \<subseteq> B" |
13169 | 70 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
71 |
|
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
72 |
|
13356 | 73 |
subsection{*Finite Set Constructions Using @{term cons}*} |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
74 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
75 |
lemma cons_subsetI: "[| a\<in>C; B\<subseteq>C |] ==> cons(a,B) \<subseteq> C" |
13169 | 76 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
77 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
78 |
lemma subset_consI: "B \<subseteq> cons(a,B)" |
13169 | 79 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
80 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
81 |
lemma cons_subset_iff [iff]: "cons(a,B)\<subseteq>C <-> a\<in>C & B\<subseteq>C" |
13169 | 82 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
83 |
|
13259 | 84 |
(*A safe special case of subset elimination, adding no new variables |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
85 |
[| cons(a,B) \<subseteq> C; [| a \<in> C; B \<subseteq> C |] ==> R |] ==> R *) |
13259 | 86 |
lemmas cons_subsetE = cons_subset_iff [THEN iffD1, THEN conjE, standard] |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
87 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
88 |
lemma subset_empty_iff: "A\<subseteq>0 <-> A=0" |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
89 |
by blast |
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
90 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
91 |
lemma subset_cons_iff: "C\<subseteq>cons(a,B) <-> C\<subseteq>B | (a\<in>C & C-{a} \<subseteq> B)" |
13169 | 92 |
by blast |
13168
afcbca3498b0
converted domrange to Isar and merged with equalities
paulson
parents:
13165
diff
changeset
|
93 |
|
13165 | 94 |
(* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*) |
95 |
lemma cons_eq: "{a} Un B = cons(a,B)" |
|
96 |
by blast |
|
97 |
||
98 |
lemma cons_commute: "cons(a, cons(b, C)) = cons(b, cons(a, C))" |
|
99 |
by blast |
|
100 |
||
101 |
lemma cons_absorb: "a: B ==> cons(a,B) = B" |
|
102 |
by blast |
|
103 |
||
104 |
lemma cons_Diff: "a: B ==> cons(a, B-{a}) = B" |
|
105 |
by blast |
|
106 |
||
14071 | 107 |
lemma Diff_cons_eq: "cons(a,B) - C = (if a\<in>C then B-C else cons(a,B-C))" |
108 |
by auto |
|
109 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
110 |
lemma equal_singleton [rule_format]: "[| a: C; \<forall>y\<in>C. y=b |] ==> C = {b}" |
13165 | 111 |
by blast |
112 |
||
13172 | 113 |
lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)" |
114 |
by blast |
|
13165 | 115 |
|
13259 | 116 |
(** singletons **) |
117 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
118 |
lemma singleton_subsetI: "a\<in>C ==> {a} \<subseteq> C" |
13259 | 119 |
by blast |
120 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
121 |
lemma singleton_subsetD: "{a} \<subseteq> C ==> a\<in>C" |
13259 | 122 |
by blast |
123 |
||
124 |
||
13356 | 125 |
(** succ **) |
13259 | 126 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
127 |
lemma subset_succI: "i \<subseteq> succ(i)" |
13259 | 128 |
by blast |
129 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
130 |
(*But if j is an ordinal or is transitive, then i\<in>j implies i\<subseteq>j! |
13259 | 131 |
See ordinal/Ord_succ_subsetI*) |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
132 |
lemma succ_subsetI: "[| i\<in>j; i\<subseteq>j |] ==> succ(i)\<subseteq>j" |
13259 | 133 |
by (unfold succ_def, blast) |
13165 | 134 |
|
13259 | 135 |
lemma succ_subsetE: |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
136 |
"[| succ(i) \<subseteq> j; [| i\<in>j; i\<subseteq>j |] ==> P |] ==> P" |
13356 | 137 |
by (unfold succ_def, blast) |
13259 | 138 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
139 |
lemma succ_subset_iff: "succ(a) \<subseteq> B <-> (a \<subseteq> B & a \<in> B)" |
13259 | 140 |
by (unfold succ_def, blast) |
141 |
||
142 |
||
13356 | 143 |
subsection{*Binary Intersection*} |
13259 | 144 |
|
145 |
(** Intersection is the greatest lower bound of two sets **) |
|
146 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
147 |
lemma Int_subset_iff: "C \<subseteq> A Int B <-> C \<subseteq> A & C \<subseteq> B" |
13259 | 148 |
by blast |
149 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
150 |
lemma Int_lower1: "A Int B \<subseteq> A" |
13259 | 151 |
by blast |
152 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
153 |
lemma Int_lower2: "A Int B \<subseteq> B" |
13259 | 154 |
by blast |
155 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
156 |
lemma Int_greatest: "[| C\<subseteq>A; C\<subseteq>B |] ==> C \<subseteq> A Int B" |
13259 | 157 |
by blast |
158 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
159 |
lemma Int_cons: "cons(a,B) Int C \<subseteq> cons(a, B Int C)" |
13165 | 160 |
by blast |
161 |
||
162 |
lemma Int_absorb [simp]: "A Int A = A" |
|
163 |
by blast |
|
164 |
||
165 |
lemma Int_left_absorb: "A Int (A Int B) = A Int B" |
|
166 |
by blast |
|
167 |
||
168 |
lemma Int_commute: "A Int B = B Int A" |
|
169 |
by blast |
|
170 |
||
171 |
lemma Int_left_commute: "A Int (B Int C) = B Int (A Int C)" |
|
172 |
by blast |
|
173 |
||
174 |
lemma Int_assoc: "(A Int B) Int C = A Int (B Int C)" |
|
175 |
by blast |
|
176 |
||
177 |
(*Intersection is an AC-operator*) |
|
178 |
lemmas Int_ac= Int_assoc Int_left_absorb Int_commute Int_left_commute |
|
179 |
||
14046 | 180 |
lemma Int_absorb1: "B \<subseteq> A ==> A \<inter> B = B" |
181 |
by blast |
|
182 |
||
183 |
lemma Int_absorb2: "A \<subseteq> B ==> A \<inter> B = A" |
|
184 |
by blast |
|
185 |
||
13165 | 186 |
lemma Int_Un_distrib: "A Int (B Un C) = (A Int B) Un (A Int C)" |
187 |
by blast |
|
188 |
||
189 |
lemma Int_Un_distrib2: "(B Un C) Int A = (B Int A) Un (C Int A)" |
|
190 |
by blast |
|
191 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
192 |
lemma subset_Int_iff: "A\<subseteq>B <-> A Int B = A" |
13165 | 193 |
by (blast elim!: equalityE) |
194 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
195 |
lemma subset_Int_iff2: "A\<subseteq>B <-> B Int A = A" |
13165 | 196 |
by (blast elim!: equalityE) |
197 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
198 |
lemma Int_Diff_eq: "C\<subseteq>A ==> (A-B) Int C = C-B" |
13165 | 199 |
by blast |
200 |
||
14071 | 201 |
lemma Int_cons_left: |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
202 |
"cons(a,A) Int B = (if a \<in> B then cons(a, A Int B) else A Int B)" |
14071 | 203 |
by auto |
204 |
||
205 |
lemma Int_cons_right: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
206 |
"A Int cons(a, B) = (if a \<in> A then cons(a, A Int B) else A Int B)" |
14071 | 207 |
by auto |
208 |
||
14076 | 209 |
lemma cons_Int_distrib: "cons(x, A \<inter> B) = cons(x, A) \<inter> cons(x, B)" |
210 |
by auto |
|
211 |
||
13356 | 212 |
subsection{*Binary Union*} |
13259 | 213 |
|
214 |
(** Union is the least upper bound of two sets *) |
|
215 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
216 |
lemma Un_subset_iff: "A Un B \<subseteq> C <-> A \<subseteq> C & B \<subseteq> C" |
13259 | 217 |
by blast |
218 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
219 |
lemma Un_upper1: "A \<subseteq> A Un B" |
13259 | 220 |
by blast |
221 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
222 |
lemma Un_upper2: "B \<subseteq> A Un B" |
13259 | 223 |
by blast |
224 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
225 |
lemma Un_least: "[| A\<subseteq>C; B\<subseteq>C |] ==> A Un B \<subseteq> C" |
13259 | 226 |
by blast |
13165 | 227 |
|
228 |
lemma Un_cons: "cons(a,B) Un C = cons(a, B Un C)" |
|
229 |
by blast |
|
230 |
||
231 |
lemma Un_absorb [simp]: "A Un A = A" |
|
232 |
by blast |
|
233 |
||
234 |
lemma Un_left_absorb: "A Un (A Un B) = A Un B" |
|
235 |
by blast |
|
236 |
||
237 |
lemma Un_commute: "A Un B = B Un A" |
|
238 |
by blast |
|
239 |
||
240 |
lemma Un_left_commute: "A Un (B Un C) = B Un (A Un C)" |
|
241 |
by blast |
|
242 |
||
243 |
lemma Un_assoc: "(A Un B) Un C = A Un (B Un C)" |
|
244 |
by blast |
|
245 |
||
246 |
(*Union is an AC-operator*) |
|
247 |
lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute |
|
248 |
||
14046 | 249 |
lemma Un_absorb1: "A \<subseteq> B ==> A \<union> B = B" |
250 |
by blast |
|
251 |
||
252 |
lemma Un_absorb2: "B \<subseteq> A ==> A \<union> B = A" |
|
253 |
by blast |
|
254 |
||
13165 | 255 |
lemma Un_Int_distrib: "(A Int B) Un C = (A Un C) Int (B Un C)" |
256 |
by blast |
|
257 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
258 |
lemma subset_Un_iff: "A\<subseteq>B <-> A Un B = B" |
13165 | 259 |
by (blast elim!: equalityE) |
260 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
261 |
lemma subset_Un_iff2: "A\<subseteq>B <-> B Un A = B" |
13165 | 262 |
by (blast elim!: equalityE) |
263 |
||
264 |
lemma Un_empty [iff]: "(A Un B = 0) <-> (A = 0 & B = 0)" |
|
265 |
by blast |
|
266 |
||
267 |
lemma Un_eq_Union: "A Un B = Union({A, B})" |
|
268 |
by blast |
|
269 |
||
13356 | 270 |
subsection{*Set Difference*} |
13259 | 271 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
272 |
lemma Diff_subset: "A-B \<subseteq> A" |
13259 | 273 |
by blast |
274 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
275 |
lemma Diff_contains: "[| C\<subseteq>A; C Int B = 0 |] ==> C \<subseteq> A-B" |
13259 | 276 |
by blast |
277 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
278 |
lemma subset_Diff_cons_iff: "B \<subseteq> A - cons(c,C) <-> B\<subseteq>A-C & c ~: B" |
13259 | 279 |
by blast |
13165 | 280 |
|
281 |
lemma Diff_cancel: "A - A = 0" |
|
282 |
by blast |
|
283 |
||
284 |
lemma Diff_triv: "A Int B = 0 ==> A - B = A" |
|
285 |
by blast |
|
286 |
||
287 |
lemma empty_Diff [simp]: "0 - A = 0" |
|
288 |
by blast |
|
289 |
||
290 |
lemma Diff_0 [simp]: "A - 0 = A" |
|
291 |
by blast |
|
292 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
293 |
lemma Diff_eq_0_iff: "A - B = 0 <-> A \<subseteq> B" |
13165 | 294 |
by (blast elim: equalityE) |
295 |
||
296 |
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) |
|
297 |
lemma Diff_cons: "A - cons(a,B) = A - B - {a}" |
|
298 |
by blast |
|
299 |
||
300 |
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) |
|
301 |
lemma Diff_cons2: "A - cons(a,B) = A - {a} - B" |
|
302 |
by blast |
|
303 |
||
304 |
lemma Diff_disjoint: "A Int (B-A) = 0" |
|
305 |
by blast |
|
306 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
307 |
lemma Diff_partition: "A\<subseteq>B ==> A Un (B-A) = B" |
13165 | 308 |
by blast |
309 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
310 |
lemma subset_Un_Diff: "A \<subseteq> B Un (A - B)" |
13165 | 311 |
by blast |
312 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
313 |
lemma double_complement: "[| A\<subseteq>B; B\<subseteq>C |] ==> B-(C-A) = A" |
13165 | 314 |
by blast |
315 |
||
316 |
lemma double_complement_Un: "(A Un B) - (B-A) = A" |
|
317 |
by blast |
|
318 |
||
319 |
lemma Un_Int_crazy: |
|
320 |
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)" |
|
321 |
apply blast |
|
322 |
done |
|
323 |
||
324 |
lemma Diff_Un: "A - (B Un C) = (A-B) Int (A-C)" |
|
325 |
by blast |
|
326 |
||
327 |
lemma Diff_Int: "A - (B Int C) = (A-B) Un (A-C)" |
|
328 |
by blast |
|
329 |
||
330 |
lemma Un_Diff: "(A Un B) - C = (A - C) Un (B - C)" |
|
331 |
by blast |
|
332 |
||
333 |
lemma Int_Diff: "(A Int B) - C = A Int (B - C)" |
|
334 |
by blast |
|
335 |
||
336 |
lemma Diff_Int_distrib: "C Int (A-B) = (C Int A) - (C Int B)" |
|
337 |
by blast |
|
338 |
||
339 |
lemma Diff_Int_distrib2: "(A-B) Int C = (A Int C) - (B Int C)" |
|
340 |
by blast |
|
341 |
||
342 |
(*Halmos, Naive Set Theory, page 16.*) |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
343 |
lemma Un_Int_assoc_iff: "(A Int B) Un C = A Int (B Un C) <-> C\<subseteq>A" |
13165 | 344 |
by (blast elim!: equalityE) |
345 |
||
346 |
||
13356 | 347 |
subsection{*Big Union and Intersection*} |
13259 | 348 |
|
349 |
(** Big Union is the least upper bound of a set **) |
|
350 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
351 |
lemma Union_subset_iff: "Union(A) \<subseteq> C <-> (\<forall>x\<in>A. x \<subseteq> C)" |
13259 | 352 |
by blast |
353 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
354 |
lemma Union_upper: "B\<in>A ==> B \<subseteq> Union(A)" |
13259 | 355 |
by blast |
356 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
357 |
lemma Union_least: "[| !!x. x\<in>A ==> x\<subseteq>C |] ==> Union(A) \<subseteq> C" |
13259 | 358 |
by blast |
13165 | 359 |
|
360 |
lemma Union_cons [simp]: "Union(cons(a,B)) = a Un Union(B)" |
|
361 |
by blast |
|
362 |
||
363 |
lemma Union_Un_distrib: "Union(A Un B) = Union(A) Un Union(B)" |
|
364 |
by blast |
|
365 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
366 |
lemma Union_Int_subset: "Union(A Int B) \<subseteq> Union(A) Int Union(B)" |
13165 | 367 |
by blast |
368 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
369 |
lemma Union_disjoint: "Union(C) Int A = 0 <-> (\<forall>B\<in>C. B Int A = 0)" |
13165 | 370 |
by (blast elim!: equalityE) |
371 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
372 |
lemma Union_empty_iff: "Union(A) = 0 <-> (\<forall>B\<in>A. B=0)" |
13165 | 373 |
by blast |
374 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
375 |
lemma Int_Union2: "Union(B) Int A = (\<Union>C\<in>B. C Int A)" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
376 |
by blast |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
377 |
|
13259 | 378 |
(** Big Intersection is the greatest lower bound of a nonempty set **) |
379 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
380 |
lemma Inter_subset_iff: "A\<noteq>0 ==> C \<subseteq> Inter(A) <-> (\<forall>x\<in>A. C \<subseteq> x)" |
13259 | 381 |
by blast |
382 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
383 |
lemma Inter_lower: "B\<in>A ==> Inter(A) \<subseteq> B" |
13259 | 384 |
by blast |
385 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
386 |
lemma Inter_greatest: "[| A\<noteq>0; !!x. x\<in>A ==> C\<subseteq>x |] ==> C \<subseteq> Inter(A)" |
13259 | 387 |
by blast |
388 |
||
389 |
(** Intersection of a family of sets **) |
|
390 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
391 |
lemma INT_lower: "x\<in>A ==> (\<Inter>x\<in>A. B(x)) \<subseteq> B(x)" |
13259 | 392 |
by blast |
393 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
394 |
lemma INT_greatest: "[| A\<noteq>0; !!x. x\<in>A ==> C\<subseteq>B(x) |] ==> C \<subseteq> (\<Inter>x\<in>A. B(x))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
395 |
by force |
13259 | 396 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
397 |
lemma Inter_0 [simp]: "Inter(0) = 0" |
13165 | 398 |
by (unfold Inter_def, blast) |
399 |
||
13259 | 400 |
lemma Inter_Un_subset: |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
401 |
"[| z\<in>A; z\<in>B |] ==> Inter(A) Un Inter(B) \<subseteq> Inter(A Int B)" |
13165 | 402 |
by blast |
403 |
||
404 |
(* A good challenge: Inter is ill-behaved on the empty set *) |
|
405 |
lemma Inter_Un_distrib: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
406 |
"[| A\<noteq>0; B\<noteq>0 |] ==> Inter(A Un B) = Inter(A) Int Inter(B)" |
13165 | 407 |
by blast |
408 |
||
409 |
lemma Union_singleton: "Union({b}) = b" |
|
410 |
by blast |
|
411 |
||
412 |
lemma Inter_singleton: "Inter({b}) = b" |
|
413 |
by blast |
|
414 |
||
415 |
lemma Inter_cons [simp]: |
|
416 |
"Inter(cons(a,B)) = (if B=0 then a else a Int Inter(B))" |
|
417 |
by force |
|
418 |
||
13356 | 419 |
subsection{*Unions and Intersections of Families*} |
13259 | 420 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
421 |
lemma subset_UN_iff_eq: "A \<subseteq> (\<Union>i\<in>I. B(i)) <-> A = (\<Union>i\<in>I. A Int B(i))" |
13259 | 422 |
by (blast elim!: equalityE) |
423 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
424 |
lemma UN_subset_iff: "(\<Union>x\<in>A. B(x)) \<subseteq> C <-> (\<forall>x\<in>A. B(x) \<subseteq> C)" |
13259 | 425 |
by blast |
426 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
427 |
lemma UN_upper: "x\<in>A ==> B(x) \<subseteq> (\<Union>x\<in>A. B(x))" |
13259 | 428 |
by (erule RepFunI [THEN Union_upper]) |
429 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
430 |
lemma UN_least: "[| !!x. x\<in>A ==> B(x)\<subseteq>C |] ==> (\<Union>x\<in>A. B(x)) \<subseteq> C" |
13259 | 431 |
by blast |
13165 | 432 |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
433 |
lemma Union_eq_UN: "Union(A) = (\<Union>x\<in>A. x)" |
13165 | 434 |
by blast |
435 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
436 |
lemma Inter_eq_INT: "Inter(A) = (\<Inter>x\<in>A. x)" |
13165 | 437 |
by (unfold Inter_def, blast) |
438 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
439 |
lemma UN_0 [simp]: "(\<Union>i\<in>0. A(i)) = 0" |
13165 | 440 |
by blast |
441 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
442 |
lemma UN_singleton: "(\<Union>x\<in>A. {x}) = A" |
13165 | 443 |
by blast |
444 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
445 |
lemma UN_Un: "(\<Union>i\<in> A Un B. C(i)) = (\<Union>i\<in> A. C(i)) Un (\<Union>i\<in>B. C(i))" |
13165 | 446 |
by blast |
447 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
448 |
lemma INT_Un: "(\<Inter>i\<in>I Un J. A(i)) = |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
449 |
(if I=0 then \<Inter>j\<in>J. A(j) |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
450 |
else if J=0 then \<Inter>i\<in>I. A(i) |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
451 |
else ((\<Inter>i\<in>I. A(i)) Int (\<Inter>j\<in>J. A(j))))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
452 |
by (simp, blast intro!: equalityI) |
13165 | 453 |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
454 |
lemma UN_UN_flatten: "(\<Union>x \<in> (\<Union>y\<in>A. B(y)). C(x)) = (\<Union>y\<in>A. \<Union>x\<in> B(y). C(x))" |
13165 | 455 |
by blast |
456 |
||
457 |
(*Halmos, Naive Set Theory, page 35.*) |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
458 |
lemma Int_UN_distrib: "B Int (\<Union>i\<in>I. A(i)) = (\<Union>i\<in>I. B Int A(i))" |
13165 | 459 |
by blast |
460 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
461 |
lemma Un_INT_distrib: "I\<noteq>0 ==> B Un (\<Inter>i\<in>I. A(i)) = (\<Inter>i\<in>I. B Un A(i))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
462 |
by auto |
13165 | 463 |
|
464 |
lemma Int_UN_distrib2: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
465 |
"(\<Union>i\<in>I. A(i)) Int (\<Union>j\<in>J. B(j)) = (\<Union>i\<in>I. \<Union>j\<in>J. A(i) Int B(j))" |
13165 | 466 |
by blast |
467 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
468 |
lemma Un_INT_distrib2: "[| I\<noteq>0; J\<noteq>0 |] ==> |
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
469 |
(\<Inter>i\<in>I. A(i)) Un (\<Inter>j\<in>J. B(j)) = (\<Inter>i\<in>I. \<Inter>j\<in>J. A(i) Un B(j))" |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
470 |
by auto |
13165 | 471 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
472 |
lemma UN_constant [simp]: "(\<Union>y\<in>A. c) = (if A=0 then 0 else c)" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
473 |
by force |
13165 | 474 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
475 |
lemma INT_constant [simp]: "(\<Inter>y\<in>A. c) = (if A=0 then 0 else c)" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
476 |
by force |
13165 | 477 |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
478 |
lemma UN_RepFun [simp]: "(\<Union>y\<in> RepFun(A,f). B(y)) = (\<Union>x\<in>A. B(f(x)))" |
13165 | 479 |
by blast |
480 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
481 |
lemma INT_RepFun [simp]: "(\<Inter>x\<in>RepFun(A,f). B(x)) = (\<Inter>a\<in>A. B(f(a)))" |
13165 | 482 |
by (auto simp add: Inter_def) |
483 |
||
484 |
lemma INT_Union_eq: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
485 |
"0 ~: A ==> (\<Inter>x\<in> Union(A). B(x)) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B(x))" |
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
486 |
apply (subgoal_tac "\<forall>x\<in>A. x~=0") |
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
487 |
prefer 2 apply blast |
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
488 |
apply (force simp add: Inter_def ball_conj_distrib) |
13165 | 489 |
done |
490 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
491 |
lemma INT_UN_eq: |
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
492 |
"(\<forall>x\<in>A. B(x) ~= 0) |
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
493 |
==> (\<Inter>z\<in> (\<Union>x\<in>A. B(x)). C(z)) = (\<Inter>x\<in>A. \<Inter>z\<in> B(x). C(z))" |
13165 | 494 |
apply (subst INT_Union_eq, blast) |
495 |
apply (simp add: Inter_def) |
|
496 |
done |
|
497 |
||
498 |
||
499 |
(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: |
|
500 |
Union of a family of unions **) |
|
501 |
||
502 |
lemma UN_Un_distrib: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
503 |
"(\<Union>i\<in>I. A(i) Un B(i)) = (\<Union>i\<in>I. A(i)) Un (\<Union>i\<in>I. B(i))" |
13165 | 504 |
by blast |
505 |
||
506 |
lemma INT_Int_distrib: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
507 |
"I\<noteq>0 ==> (\<Inter>i\<in>I. A(i) Int B(i)) = (\<Inter>i\<in>I. A(i)) Int (\<Inter>i\<in>I. B(i))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
508 |
by (blast elim!: not_emptyE) |
13165 | 509 |
|
510 |
lemma UN_Int_subset: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
511 |
"(\<Union>z\<in>I Int J. A(z)) \<subseteq> (\<Union>z\<in>I. A(z)) Int (\<Union>z\<in>J. A(z))" |
13165 | 512 |
by blast |
513 |
||
514 |
(** Devlin, page 12, exercise 5: Complements **) |
|
515 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
516 |
lemma Diff_UN: "I\<noteq>0 ==> B - (\<Union>i\<in>I. A(i)) = (\<Inter>i\<in>I. B - A(i))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
517 |
by (blast elim!: not_emptyE) |
13165 | 518 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
519 |
lemma Diff_INT: "I\<noteq>0 ==> B - (\<Inter>i\<in>I. A(i)) = (\<Union>i\<in>I. B - A(i))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
520 |
by (blast elim!: not_emptyE) |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
521 |
|
13165 | 522 |
|
523 |
(** Unions and Intersections with General Sum **) |
|
524 |
||
525 |
(*Not suitable for rewriting: LOOPS!*) |
|
526 |
lemma Sigma_cons1: "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)" |
|
527 |
by blast |
|
528 |
||
529 |
(*Not suitable for rewriting: LOOPS!*) |
|
530 |
lemma Sigma_cons2: "A * cons(b,B) = A*{b} Un A*B" |
|
531 |
by blast |
|
532 |
||
533 |
lemma Sigma_succ1: "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)" |
|
534 |
by blast |
|
535 |
||
536 |
lemma Sigma_succ2: "A * succ(B) = A*{B} Un A*B" |
|
537 |
by blast |
|
538 |
||
539 |
lemma SUM_UN_distrib1: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
540 |
"(\<Sigma> x \<in> (\<Union>y\<in>A. C(y)). B(x)) = (\<Union>y\<in>A. \<Sigma> x\<in>C(y). B(x))" |
13165 | 541 |
by blast |
542 |
||
543 |
lemma SUM_UN_distrib2: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
544 |
"(\<Sigma> i\<in>I. \<Union>j\<in>J. C(i,j)) = (\<Union>j\<in>J. \<Sigma> i\<in>I. C(i,j))" |
13165 | 545 |
by blast |
546 |
||
547 |
lemma SUM_Un_distrib1: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
548 |
"(\<Sigma> i\<in>I Un J. C(i)) = (\<Sigma> i\<in>I. C(i)) Un (\<Sigma> j\<in>J. C(j))" |
13165 | 549 |
by blast |
550 |
||
551 |
lemma SUM_Un_distrib2: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
552 |
"(\<Sigma> i\<in>I. A(i) Un B(i)) = (\<Sigma> i\<in>I. A(i)) Un (\<Sigma> i\<in>I. B(i))" |
13165 | 553 |
by blast |
554 |
||
555 |
(*First-order version of the above, for rewriting*) |
|
556 |
lemma prod_Un_distrib2: "I * (A Un B) = I*A Un I*B" |
|
557 |
by (rule SUM_Un_distrib2) |
|
558 |
||
559 |
lemma SUM_Int_distrib1: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
560 |
"(\<Sigma> i\<in>I Int J. C(i)) = (\<Sigma> i\<in>I. C(i)) Int (\<Sigma> j\<in>J. C(j))" |
13165 | 561 |
by blast |
562 |
||
563 |
lemma SUM_Int_distrib2: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
564 |
"(\<Sigma> i\<in>I. A(i) Int B(i)) = (\<Sigma> i\<in>I. A(i)) Int (\<Sigma> i\<in>I. B(i))" |
13165 | 565 |
by blast |
566 |
||
567 |
(*First-order version of the above, for rewriting*) |
|
568 |
lemma prod_Int_distrib2: "I * (A Int B) = I*A Int I*B" |
|
569 |
by (rule SUM_Int_distrib2) |
|
570 |
||
571 |
(*Cf Aczel, Non-Well-Founded Sets, page 115*) |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
572 |
lemma SUM_eq_UN: "(\<Sigma> i\<in>I. A(i)) = (\<Union>i\<in>I. {i} * A(i))" |
13165 | 573 |
by blast |
574 |
||
13544 | 575 |
lemma times_subset_iff: |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
576 |
"(A'*B' \<subseteq> A*B) <-> (A' = 0 | B' = 0 | (A'\<subseteq>A) & (B'\<subseteq>B))" |
13544 | 577 |
by blast |
578 |
||
579 |
lemma Int_Sigma_eq: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14095
diff
changeset
|
580 |
"(\<Sigma> x \<in> A'. B'(x)) Int (\<Sigma> x \<in> A. B(x)) = (\<Sigma> x \<in> A' Int A. B'(x) Int B(x))" |
13544 | 581 |
by blast |
582 |
||
13165 | 583 |
(** Domain **) |
584 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
585 |
lemma domain_iff: "a: domain(r) <-> (EX y. <a,y>\<in> r)" |
13259 | 586 |
by (unfold domain_def, blast) |
587 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
588 |
lemma domainI [intro]: "<a,b>\<in> r ==> a: domain(r)" |
13259 | 589 |
by (unfold domain_def, blast) |
590 |
||
591 |
lemma domainE [elim!]: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
592 |
"[| a \<in> domain(r); !!y. <a,y>\<in> r ==> P |] ==> P" |
13259 | 593 |
by (unfold domain_def, blast) |
594 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
595 |
lemma domain_subset: "domain(Sigma(A,B)) \<subseteq> A" |
13259 | 596 |
by blast |
597 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
598 |
lemma domain_of_prod: "b\<in>B ==> domain(A*B) = A" |
13165 | 599 |
by blast |
600 |
||
601 |
lemma domain_0 [simp]: "domain(0) = 0" |
|
602 |
by blast |
|
603 |
||
604 |
lemma domain_cons [simp]: "domain(cons(<a,b>,r)) = cons(a, domain(r))" |
|
605 |
by blast |
|
606 |
||
607 |
lemma domain_Un_eq [simp]: "domain(A Un B) = domain(A) Un domain(B)" |
|
608 |
by blast |
|
609 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
610 |
lemma domain_Int_subset: "domain(A Int B) \<subseteq> domain(A) Int domain(B)" |
13165 | 611 |
by blast |
612 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
613 |
lemma domain_Diff_subset: "domain(A) - domain(B) \<subseteq> domain(A - B)" |
13165 | 614 |
by blast |
615 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
616 |
lemma domain_UN: "domain(\<Union>x\<in>A. B(x)) = (\<Union>x\<in>A. domain(B(x)))" |
13165 | 617 |
by blast |
618 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
619 |
lemma domain_Union: "domain(Union(A)) = (\<Union>x\<in>A. domain(x))" |
13165 | 620 |
by blast |
621 |
||
622 |
||
623 |
(** Range **) |
|
624 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
625 |
lemma rangeI [intro]: "<a,b>\<in> r ==> b \<in> range(r)" |
13259 | 626 |
apply (unfold range_def) |
627 |
apply (erule converseI [THEN domainI]) |
|
628 |
done |
|
629 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
630 |
lemma rangeE [elim!]: "[| b \<in> range(r); !!x. <x,b>\<in> r ==> P |] ==> P" |
13259 | 631 |
by (unfold range_def, blast) |
632 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
633 |
lemma range_subset: "range(A*B) \<subseteq> B" |
13259 | 634 |
apply (unfold range_def) |
635 |
apply (subst converse_prod) |
|
636 |
apply (rule domain_subset) |
|
637 |
done |
|
638 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
639 |
lemma range_of_prod: "a\<in>A ==> range(A*B) = B" |
13165 | 640 |
by blast |
641 |
||
642 |
lemma range_0 [simp]: "range(0) = 0" |
|
643 |
by blast |
|
644 |
||
645 |
lemma range_cons [simp]: "range(cons(<a,b>,r)) = cons(b, range(r))" |
|
646 |
by blast |
|
647 |
||
648 |
lemma range_Un_eq [simp]: "range(A Un B) = range(A) Un range(B)" |
|
649 |
by blast |
|
650 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
651 |
lemma range_Int_subset: "range(A Int B) \<subseteq> range(A) Int range(B)" |
13165 | 652 |
by blast |
653 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
654 |
lemma range_Diff_subset: "range(A) - range(B) \<subseteq> range(A - B)" |
13165 | 655 |
by blast |
656 |
||
13259 | 657 |
lemma domain_converse [simp]: "domain(converse(r)) = range(r)" |
658 |
by blast |
|
659 |
||
13165 | 660 |
lemma range_converse [simp]: "range(converse(r)) = domain(r)" |
661 |
by blast |
|
662 |
||
663 |
||
664 |
(** Field **) |
|
665 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
666 |
lemma fieldI1: "<a,b>\<in> r ==> a \<in> field(r)" |
13259 | 667 |
by (unfold field_def, blast) |
668 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
669 |
lemma fieldI2: "<a,b>\<in> r ==> b \<in> field(r)" |
13259 | 670 |
by (unfold field_def, blast) |
671 |
||
672 |
lemma fieldCI [intro]: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
673 |
"(~ <c,a>\<in>r ==> <a,b>\<in> r) ==> a \<in> field(r)" |
13259 | 674 |
apply (unfold field_def, blast) |
675 |
done |
|
676 |
||
677 |
lemma fieldE [elim!]: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
678 |
"[| a \<in> field(r); |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
679 |
!!x. <a,x>\<in> r ==> P; |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
680 |
!!x. <x,a>\<in> r ==> P |] ==> P" |
13259 | 681 |
by (unfold field_def, blast) |
682 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
683 |
lemma field_subset: "field(A*B) \<subseteq> A Un B" |
13259 | 684 |
by blast |
685 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
686 |
lemma domain_subset_field: "domain(r) \<subseteq> field(r)" |
13259 | 687 |
apply (unfold field_def) |
688 |
apply (rule Un_upper1) |
|
689 |
done |
|
690 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
691 |
lemma range_subset_field: "range(r) \<subseteq> field(r)" |
13259 | 692 |
apply (unfold field_def) |
693 |
apply (rule Un_upper2) |
|
694 |
done |
|
695 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
696 |
lemma domain_times_range: "r \<subseteq> Sigma(A,B) ==> r \<subseteq> domain(r)*range(r)" |
13259 | 697 |
by blast |
698 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
699 |
lemma field_times_field: "r \<subseteq> Sigma(A,B) ==> r \<subseteq> field(r)*field(r)" |
13259 | 700 |
by blast |
701 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
702 |
lemma relation_field_times_field: "relation(r) ==> r \<subseteq> field(r)*field(r)" |
13259 | 703 |
by (simp add: relation_def, blast) |
704 |
||
13165 | 705 |
lemma field_of_prod: "field(A*A) = A" |
706 |
by blast |
|
707 |
||
708 |
lemma field_0 [simp]: "field(0) = 0" |
|
709 |
by blast |
|
710 |
||
711 |
lemma field_cons [simp]: "field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))" |
|
712 |
by blast |
|
713 |
||
714 |
lemma field_Un_eq [simp]: "field(A Un B) = field(A) Un field(B)" |
|
715 |
by blast |
|
716 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
717 |
lemma field_Int_subset: "field(A Int B) \<subseteq> field(A) Int field(B)" |
13165 | 718 |
by blast |
719 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
720 |
lemma field_Diff_subset: "field(A) - field(B) \<subseteq> field(A - B)" |
13165 | 721 |
by blast |
722 |
||
723 |
lemma field_converse [simp]: "field(converse(r)) = field(r)" |
|
724 |
by blast |
|
725 |
||
13259 | 726 |
(** The Union of a set of relations is a relation -- Lemma for fun_Union **) |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
727 |
lemma rel_Union: "(\<forall>x\<in>S. EX A B. x \<subseteq> A*B) ==> |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
728 |
Union(S) \<subseteq> domain(Union(S)) * range(Union(S))" |
13259 | 729 |
by blast |
13165 | 730 |
|
13259 | 731 |
(** The Union of 2 relations is a relation (Lemma for fun_Un) **) |
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
732 |
lemma rel_Un: "[| r \<subseteq> A*B; s \<subseteq> C*D |] ==> (r Un s) \<subseteq> (A Un C) * (B Un D)" |
13259 | 733 |
by blast |
734 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
735 |
lemma domain_Diff_eq: "[| <a,c> \<in> r; c~=b |] ==> domain(r-{<a,b>}) = domain(r)" |
13259 | 736 |
by blast |
737 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
738 |
lemma range_Diff_eq: "[| <c,b> \<in> r; c~=a |] ==> range(r-{<a,b>}) = range(r)" |
13259 | 739 |
by blast |
740 |
||
741 |
||
13356 | 742 |
subsection{*Image of a Set under a Function or Relation*} |
13259 | 743 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
744 |
lemma image_iff: "b \<in> r``A <-> (\<exists>x\<in>A. <x,b>\<in>r)" |
13259 | 745 |
by (unfold image_def, blast) |
746 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
747 |
lemma image_singleton_iff: "b \<in> r``{a} <-> <a,b>\<in>r" |
13259 | 748 |
by (rule image_iff [THEN iff_trans], blast) |
749 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
750 |
lemma imageI [intro]: "[| <a,b>\<in> r; a\<in>A |] ==> b \<in> r``A" |
13259 | 751 |
by (unfold image_def, blast) |
752 |
||
753 |
lemma imageE [elim!]: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
754 |
"[| b: r``A; !!x.[| <x,b>\<in> r; x\<in>A |] ==> P |] ==> P" |
13259 | 755 |
by (unfold image_def, blast) |
756 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
757 |
lemma image_subset: "r \<subseteq> A*B ==> r``C \<subseteq> B" |
13259 | 758 |
by blast |
13165 | 759 |
|
760 |
lemma image_0 [simp]: "r``0 = 0" |
|
761 |
by blast |
|
762 |
||
763 |
lemma image_Un [simp]: "r``(A Un B) = (r``A) Un (r``B)" |
|
764 |
by blast |
|
765 |
||
14883 | 766 |
lemma image_UN: "r `` (\<Union>x\<in>A. B(x)) = (\<Union>x\<in>A. r `` B(x))" |
767 |
by blast |
|
768 |
||
769 |
lemma Collect_image_eq: |
|
770 |
"{z \<in> Sigma(A,B). P(z)} `` C = (\<Union>x \<in> A. {y \<in> B(x). x \<in> C & P(<x,y>)})" |
|
771 |
by blast |
|
772 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
773 |
lemma image_Int_subset: "r``(A Int B) \<subseteq> (r``A) Int (r``B)" |
13165 | 774 |
by blast |
775 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
776 |
lemma image_Int_square_subset: "(r Int A*A)``B \<subseteq> (r``B) Int A" |
13165 | 777 |
by blast |
778 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
779 |
lemma image_Int_square: "B\<subseteq>A ==> (r Int A*A)``B = (r``B) Int A" |
13165 | 780 |
by blast |
781 |
||
782 |
||
783 |
(*Image laws for special relations*) |
|
784 |
lemma image_0_left [simp]: "0``A = 0" |
|
785 |
by blast |
|
786 |
||
787 |
lemma image_Un_left: "(r Un s)``A = (r``A) Un (s``A)" |
|
788 |
by blast |
|
789 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
790 |
lemma image_Int_subset_left: "(r Int s)``A \<subseteq> (r``A) Int (s``A)" |
13165 | 791 |
by blast |
792 |
||
793 |
||
13356 | 794 |
subsection{*Inverse Image of a Set under a Function or Relation*} |
13259 | 795 |
|
796 |
lemma vimage_iff: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
797 |
"a \<in> r-``B <-> (\<exists>y\<in>B. <a,y>\<in>r)" |
13259 | 798 |
by (unfold vimage_def image_def converse_def, blast) |
799 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
800 |
lemma vimage_singleton_iff: "a \<in> r-``{b} <-> <a,b>\<in>r" |
13259 | 801 |
by (rule vimage_iff [THEN iff_trans], blast) |
802 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
803 |
lemma vimageI [intro]: "[| <a,b>\<in> r; b\<in>B |] ==> a \<in> r-``B" |
13259 | 804 |
by (unfold vimage_def, blast) |
805 |
||
806 |
lemma vimageE [elim!]: |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
807 |
"[| a: r-``B; !!x.[| <a,x>\<in> r; x\<in>B |] ==> P |] ==> P" |
13259 | 808 |
apply (unfold vimage_def, blast) |
809 |
done |
|
810 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
811 |
lemma vimage_subset: "r \<subseteq> A*B ==> r-``C \<subseteq> A" |
13259 | 812 |
apply (unfold vimage_def) |
813 |
apply (erule converse_type [THEN image_subset]) |
|
814 |
done |
|
13165 | 815 |
|
816 |
lemma vimage_0 [simp]: "r-``0 = 0" |
|
817 |
by blast |
|
818 |
||
819 |
lemma vimage_Un [simp]: "r-``(A Un B) = (r-``A) Un (r-``B)" |
|
820 |
by blast |
|
821 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
822 |
lemma vimage_Int_subset: "r-``(A Int B) \<subseteq> (r-``A) Int (r-``B)" |
13165 | 823 |
by blast |
824 |
||
825 |
(*NOT suitable for rewriting*) |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
826 |
lemma vimage_eq_UN: "f -``B = (\<Union>y\<in>B. f-``{y})" |
13165 | 827 |
by blast |
828 |
||
829 |
lemma function_vimage_Int: |
|
830 |
"function(f) ==> f-``(A Int B) = (f-``A) Int (f-``B)" |
|
831 |
by (unfold function_def, blast) |
|
832 |
||
833 |
lemma function_vimage_Diff: "function(f) ==> f-``(A-B) = (f-``A) - (f-``B)" |
|
834 |
by (unfold function_def, blast) |
|
835 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
836 |
lemma function_image_vimage: "function(f) ==> f `` (f-`` A) \<subseteq> A" |
13165 | 837 |
by (unfold function_def, blast) |
838 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
839 |
lemma vimage_Int_square_subset: "(r Int A*A)-``B \<subseteq> (r-``B) Int A" |
13165 | 840 |
by blast |
841 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
842 |
lemma vimage_Int_square: "B\<subseteq>A ==> (r Int A*A)-``B = (r-``B) Int A" |
13165 | 843 |
by blast |
844 |
||
845 |
||
846 |
||
847 |
(*Invese image laws for special relations*) |
|
848 |
lemma vimage_0_left [simp]: "0-``A = 0" |
|
849 |
by blast |
|
850 |
||
851 |
lemma vimage_Un_left: "(r Un s)-``A = (r-``A) Un (s-``A)" |
|
852 |
by blast |
|
853 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
854 |
lemma vimage_Int_subset_left: "(r Int s)-``A \<subseteq> (r-``A) Int (s-``A)" |
13165 | 855 |
by blast |
856 |
||
857 |
||
858 |
(** Converse **) |
|
859 |
||
860 |
lemma converse_Un [simp]: "converse(A Un B) = converse(A) Un converse(B)" |
|
861 |
by blast |
|
862 |
||
863 |
lemma converse_Int [simp]: "converse(A Int B) = converse(A) Int converse(B)" |
|
864 |
by blast |
|
865 |
||
866 |
lemma converse_Diff [simp]: "converse(A - B) = converse(A) - converse(B)" |
|
867 |
by blast |
|
868 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
869 |
lemma converse_UN [simp]: "converse(\<Union>x\<in>A. B(x)) = (\<Union>x\<in>A. converse(B(x)))" |
13165 | 870 |
by blast |
871 |
||
872 |
(*Unfolding Inter avoids using excluded middle on A=0*) |
|
873 |
lemma converse_INT [simp]: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13611
diff
changeset
|
874 |
"converse(\<Inter>x\<in>A. B(x)) = (\<Inter>x\<in>A. converse(B(x)))" |
13165 | 875 |
apply (unfold Inter_def, blast) |
876 |
done |
|
877 |
||
13356 | 878 |
|
879 |
subsection{*Powerset Operator*} |
|
13165 | 880 |
|
881 |
lemma Pow_0 [simp]: "Pow(0) = {0}" |
|
882 |
by blast |
|
883 |
||
884 |
lemma Pow_insert: "Pow (cons(a,A)) = Pow(A) Un {cons(a,X) . X: Pow(A)}" |
|
885 |
apply (rule equalityI, safe) |
|
886 |
apply (erule swap) |
|
887 |
apply (rule_tac a = "x-{a}" in RepFun_eqI, auto) |
|
888 |
done |
|
889 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
890 |
lemma Un_Pow_subset: "Pow(A) Un Pow(B) \<subseteq> Pow(A Un B)" |
13165 | 891 |
by blast |
892 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
893 |
lemma UN_Pow_subset: "(\<Union>x\<in>A. Pow(B(x))) \<subseteq> Pow(\<Union>x\<in>A. B(x))" |
13165 | 894 |
by blast |
895 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
896 |
lemma subset_Pow_Union: "A \<subseteq> Pow(Union(A))" |
13165 | 897 |
by blast |
898 |
||
899 |
lemma Union_Pow_eq [simp]: "Union(Pow(A)) = A" |
|
900 |
by blast |
|
901 |
||
14077 | 902 |
lemma Union_Pow_iff: "Union(A) \<in> Pow(B) <-> A \<in> Pow(Pow(B))" |
903 |
by blast |
|
904 |
||
13165 | 905 |
lemma Pow_Int_eq [simp]: "Pow(A Int B) = Pow(A) Int Pow(B)" |
906 |
by blast |
|
907 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
908 |
lemma Pow_INT_eq: "A\<noteq>0 ==> Pow(\<Inter>x\<in>A. B(x)) = (\<Inter>x\<in>A. Pow(B(x)))" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
909 |
by (blast elim!: not_emptyE) |
13165 | 910 |
|
13356 | 911 |
|
912 |
subsection{*RepFun*} |
|
13259 | 913 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
914 |
lemma RepFun_subset: "[| !!x. x\<in>A ==> f(x) \<in> B |] ==> {f(x). x\<in>A} \<subseteq> B" |
13259 | 915 |
by blast |
13165 | 916 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
917 |
lemma RepFun_eq_0_iff [simp]: "{f(x).x\<in>A}=0 <-> A=0" |
13165 | 918 |
by blast |
919 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
920 |
lemma RepFun_constant [simp]: "{c. x\<in>A} = (if A=0 then 0 else {c})" |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
921 |
by force |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
922 |
|
13165 | 923 |
|
13356 | 924 |
subsection{*Collect*} |
13259 | 925 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
926 |
lemma Collect_subset: "Collect(A,P) \<subseteq> A" |
13259 | 927 |
by blast |
2469 | 928 |
|
13165 | 929 |
lemma Collect_Un: "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)" |
930 |
by blast |
|
931 |
||
932 |
lemma Collect_Int: "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)" |
|
933 |
by blast |
|
934 |
||
935 |
lemma Collect_Diff: "Collect(A - B, P) = Collect(A,P) - Collect(B,P)" |
|
936 |
by blast |
|
937 |
||
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
938 |
lemma Collect_cons: "{x\<in>cons(a,B). P(x)} = |
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
939 |
(if P(a) then cons(a, {x\<in>B. P(x)}) else {x\<in>B. P(x)})" |
13165 | 940 |
by (simp, blast) |
941 |
||
942 |
lemma Int_Collect_self_eq: "A Int Collect(A,P) = Collect(A,P)" |
|
943 |
by blast |
|
944 |
||
945 |
lemma Collect_Collect_eq [simp]: |
|
946 |
"Collect(Collect(A,P), Q) = Collect(A, %x. P(x) & Q(x))" |
|
947 |
by blast |
|
948 |
||
949 |
lemma Collect_Int_Collect_eq: |
|
950 |
"Collect(A,P) Int Collect(A,Q) = Collect(A, %x. P(x) & Q(x))" |
|
951 |
by blast |
|
952 |
||
13203
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13178
diff
changeset
|
953 |
lemma Collect_Union_eq [simp]: |
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13178
diff
changeset
|
954 |
"Collect(\<Union>x\<in>A. B(x), P) = (\<Union>x\<in>A. Collect(B(x), P))" |
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13178
diff
changeset
|
955 |
by blast |
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13178
diff
changeset
|
956 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
957 |
lemma Collect_Int_left: "{x\<in>A. P(x)} Int B = {x \<in> A Int B. P(x)}" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
958 |
by blast |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
959 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
960 |
lemma Collect_Int_right: "A Int {x\<in>B. P(x)} = {x \<in> A Int B. P(x)}" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
961 |
by blast |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
962 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
963 |
lemma Collect_disj_eq: "{x\<in>A. P(x) | Q(x)} = Collect(A, P) Un Collect(A, Q)" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
964 |
by blast |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
965 |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
966 |
lemma Collect_conj_eq: "{x\<in>A. P(x) & Q(x)} = Collect(A, P) Int Collect(A, Q)" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
967 |
by blast |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14077
diff
changeset
|
968 |
|
13259 | 969 |
lemmas subset_SIs = subset_refl cons_subsetI subset_consI |
970 |
Union_least UN_least Un_least |
|
971 |
Inter_greatest Int_greatest RepFun_subset |
|
972 |
Un_upper1 Un_upper2 Int_lower1 Int_lower2 |
|
973 |
||
24893 | 974 |
ML {* |
975 |
val subset_cs = @{claset} |
|
976 |
delrules [@{thm subsetI}, @{thm subsetCE}] |
|
977 |
addSIs @{thms subset_SIs} |
|
978 |
addIs [@{thm Union_upper}, @{thm Inter_lower}] |
|
979 |
addSEs [@{thm cons_subsetE}]; |
|
13259 | 980 |
*} |
981 |
(*subset_cs is a claset for subset reasoning*) |
|
982 |
||
13165 | 983 |
ML |
984 |
{* |
|
24893 | 985 |
val ZF_cs = @{claset} delrules [@{thm equalityI}]; |
13165 | 986 |
*} |
14046 | 987 |
|
13165 | 988 |
end |
989 |