author | wenzelm |
Tue, 01 Oct 2024 20:39:16 +0200 | |
changeset 81091 | c007e6d9941d |
parent 81011 | 6d34c2bedaa3 |
child 81145 | c9f1e926d4ed |
permissions | -rw-r--r-- |
60770 | 1 |
section \<open>Extending FOL by a modified version of HOL set theory\<close> |
17456 | 2 |
|
3 |
theory Set |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
62143
diff
changeset
|
4 |
imports FOL |
17456 | 5 |
begin |
0 | 6 |
|
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
38499
diff
changeset
|
7 |
declare [[eta_contract]] |
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
38499
diff
changeset
|
8 |
|
17456 | 9 |
typedecl 'a set |
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
48475
diff
changeset
|
10 |
instance set :: ("term") "term" .. |
0 | 11 |
|
62143 | 12 |
|
13 |
subsection \<open>Set comprehension and membership\<close> |
|
14 |
||
15 |
axiomatization Collect :: "['a \<Rightarrow> o] \<Rightarrow> 'a set" |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
16 |
and mem :: "['a, 'a set] \<Rightarrow> o" (infixl \<open>:\<close> 50) |
62143 | 17 |
where mem_Collect_iff: "(a : Collect(P)) \<longleftrightarrow> P(a)" |
18 |
and set_extension: "A = B \<longleftrightarrow> (ALL x. x:A \<longleftrightarrow> x:B)" |
|
0 | 19 |
|
3935 | 20 |
syntax |
80917 | 21 |
"_Coll" :: "[idt, o] \<Rightarrow> 'a set" (\<open>(\<open>indent=1 notation=\<open>mixfix Collect\<close>\<close>{_./ _})\<close>) |
0 | 22 |
translations |
62143 | 23 |
"{x. P}" == "CONST Collect(\<lambda>x. P)" |
20140 | 24 |
|
58977 | 25 |
lemma CollectI: "P(a) \<Longrightarrow> a : {x. P(x)}" |
20140 | 26 |
apply (rule mem_Collect_iff [THEN iffD2]) |
27 |
apply assumption |
|
28 |
done |
|
29 |
||
58977 | 30 |
lemma CollectD: "a : {x. P(x)} \<Longrightarrow> P(a)" |
20140 | 31 |
apply (erule mem_Collect_iff [THEN iffD1]) |
32 |
done |
|
33 |
||
34 |
lemmas CollectE = CollectD [elim_format] |
|
35 |
||
58977 | 36 |
lemma set_ext: "(\<And>x. x:A \<longleftrightarrow> x:B) \<Longrightarrow> A = B" |
20140 | 37 |
apply (rule set_extension [THEN iffD2]) |
38 |
apply simp |
|
39 |
done |
|
40 |
||
41 |
||
60770 | 42 |
subsection \<open>Bounded quantifiers\<close> |
20140 | 43 |
|
62143 | 44 |
definition Ball :: "['a set, 'a \<Rightarrow> o] \<Rightarrow> o" |
45 |
where "Ball(A, P) == ALL x. x:A \<longrightarrow> P(x)" |
|
46 |
||
47 |
definition Bex :: "['a set, 'a \<Rightarrow> o] \<Rightarrow> o" |
|
48 |
where "Bex(A, P) == EX x. x:A \<and> P(x)" |
|
49 |
||
50 |
syntax |
|
80917 | 51 |
"_Ball" :: "[idt, 'a set, o] \<Rightarrow> o" (\<open>(\<open>notation=\<open>binder ALL:\<close>\<close>ALL _:_./ _)\<close> [0, 0, 0] 10) |
52 |
"_Bex" :: "[idt, 'a set, o] \<Rightarrow> o" (\<open>(\<open>notation=\<open>binder EX:\<close>\<close>EX _:_./ _)\<close> [0, 0, 0] 10) |
|
80761 | 53 |
syntax_consts |
54 |
"_Ball" == Ball and |
|
55 |
"_Bex" == Bex |
|
62143 | 56 |
translations |
57 |
"ALL x:A. P" == "CONST Ball(A, \<lambda>x. P)" |
|
58 |
"EX x:A. P" == "CONST Bex(A, \<lambda>x. P)" |
|
59 |
||
58977 | 60 |
lemma ballI: "(\<And>x. x:A \<Longrightarrow> P(x)) \<Longrightarrow> ALL x:A. P(x)" |
20140 | 61 |
by (simp add: Ball_def) |
62 |
||
58977 | 63 |
lemma bspec: "\<lbrakk>ALL x:A. P(x); x:A\<rbrakk> \<Longrightarrow> P(x)" |
20140 | 64 |
by (simp add: Ball_def) |
65 |
||
58977 | 66 |
lemma ballE: "\<lbrakk>ALL x:A. P(x); P(x) \<Longrightarrow> Q; \<not> x:A \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
20140 | 67 |
unfolding Ball_def by blast |
68 |
||
58977 | 69 |
lemma bexI: "\<lbrakk>P(x); x:A\<rbrakk> \<Longrightarrow> EX x:A. P(x)" |
20140 | 70 |
unfolding Bex_def by blast |
71 |
||
58977 | 72 |
lemma bexCI: "\<lbrakk>EX x:A. \<not>P(x) \<Longrightarrow> P(a); a:A\<rbrakk> \<Longrightarrow> EX x:A. P(x)" |
20140 | 73 |
unfolding Bex_def by blast |
74 |
||
58977 | 75 |
lemma bexE: "\<lbrakk>EX x:A. P(x); \<And>x. \<lbrakk>x:A; P(x)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
20140 | 76 |
unfolding Bex_def by blast |
77 |
||
78 |
(*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*) |
|
58977 | 79 |
lemma ball_rew: "(ALL x:A. True) \<longleftrightarrow> True" |
20140 | 80 |
by (blast intro: ballI) |
81 |
||
62143 | 82 |
subsubsection \<open>Congruence rules\<close> |
20140 | 83 |
|
84 |
lemma ball_cong: |
|
58977 | 85 |
"\<lbrakk>A = A'; \<And>x. x:A' \<Longrightarrow> P(x) \<longleftrightarrow> P'(x)\<rbrakk> \<Longrightarrow> |
86 |
(ALL x:A. P(x)) \<longleftrightarrow> (ALL x:A'. P'(x))" |
|
20140 | 87 |
by (blast intro: ballI elim: ballE) |
88 |
||
89 |
lemma bex_cong: |
|
58977 | 90 |
"\<lbrakk>A = A'; \<And>x. x:A' \<Longrightarrow> P(x) \<longleftrightarrow> P'(x)\<rbrakk> \<Longrightarrow> |
91 |
(EX x:A. P(x)) \<longleftrightarrow> (EX x:A'. P'(x))" |
|
20140 | 92 |
by (blast intro: bexI elim: bexE) |
93 |
||
94 |
||
62143 | 95 |
subsection \<open>Further operations\<close> |
96 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
97 |
definition subset :: "['a set, 'a set] \<Rightarrow> o" (infixl \<open><=\<close> 50) |
62143 | 98 |
where "A <= B == ALL x:A. x:B" |
99 |
||
100 |
definition mono :: "['a set \<Rightarrow> 'b set] \<Rightarrow> o" |
|
101 |
where "mono(f) == (ALL A B. A <= B \<longrightarrow> f(A) <= f(B))" |
|
102 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
103 |
definition singleton :: "'a \<Rightarrow> 'a set" (\<open>{_}\<close>) |
62143 | 104 |
where "{a} == {x. x=a}" |
105 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
106 |
definition empty :: "'a set" (\<open>{}\<close>) |
62143 | 107 |
where "{} == {x. False}" |
108 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
109 |
definition Un :: "['a set, 'a set] \<Rightarrow> 'a set" (infixl \<open>Un\<close> 65) |
62143 | 110 |
where "A Un B == {x. x:A | x:B}" |
111 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
112 |
definition Int :: "['a set, 'a set] \<Rightarrow> 'a set" (infixl \<open>Int\<close> 70) |
62143 | 113 |
where "A Int B == {x. x:A \<and> x:B}" |
114 |
||
115 |
definition Compl :: "('a set) \<Rightarrow> 'a set" |
|
116 |
where "Compl(A) == {x. \<not>x:A}" |
|
117 |
||
118 |
||
119 |
subsection \<open>Big Intersection / Union\<close> |
|
120 |
||
121 |
definition INTER :: "['a set, 'a \<Rightarrow> 'b set] \<Rightarrow> 'b set" |
|
122 |
where "INTER(A, B) == {y. ALL x:A. y: B(x)}" |
|
123 |
||
124 |
definition UNION :: "['a set, 'a \<Rightarrow> 'b set] \<Rightarrow> 'b set" |
|
125 |
where "UNION(A, B) == {y. EX x:A. y: B(x)}" |
|
126 |
||
127 |
syntax |
|
80917 | 128 |
"_INTER" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set" (\<open>(\<open>notation=\<open>binder INT:\<close>\<close>INT _:_./ _)\<close> [0, 0, 0] 10) |
129 |
"_UNION" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set" (\<open>(\<open>notation=\<open>binder UN:\<close>\<close>UN _:_./ _)\<close> [0, 0, 0] 10) |
|
80761 | 130 |
syntax_consts |
131 |
"_INTER" == INTER and |
|
132 |
"_UNION" == UNION |
|
62143 | 133 |
translations |
134 |
"INT x:A. B" == "CONST INTER(A, \<lambda>x. B)" |
|
135 |
"UN x:A. B" == "CONST UNION(A, \<lambda>x. B)" |
|
136 |
||
137 |
definition Inter :: "(('a set)set) \<Rightarrow> 'a set" |
|
138 |
where "Inter(S) == (INT x:S. x)" |
|
139 |
||
140 |
definition Union :: "(('a set)set) \<Rightarrow> 'a set" |
|
141 |
where "Union(S) == (UN x:S. x)" |
|
142 |
||
143 |
||
60770 | 144 |
subsection \<open>Rules for subsets\<close> |
20140 | 145 |
|
58977 | 146 |
lemma subsetI: "(\<And>x. x:A \<Longrightarrow> x:B) \<Longrightarrow> A <= B" |
20140 | 147 |
unfolding subset_def by (blast intro: ballI) |
148 |
||
149 |
(*Rule in Modus Ponens style*) |
|
58977 | 150 |
lemma subsetD: "\<lbrakk>A <= B; c:A\<rbrakk> \<Longrightarrow> c:B" |
20140 | 151 |
unfolding subset_def by (blast elim: ballE) |
152 |
||
153 |
(*Classical elimination rule*) |
|
58977 | 154 |
lemma subsetCE: "\<lbrakk>A <= B; \<not>(c:A) \<Longrightarrow> P; c:B \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
20140 | 155 |
by (blast dest: subsetD) |
156 |
||
157 |
lemma subset_refl: "A <= A" |
|
158 |
by (blast intro: subsetI) |
|
159 |
||
58977 | 160 |
lemma subset_trans: "\<lbrakk>A <= B; B <= C\<rbrakk> \<Longrightarrow> A <= C" |
20140 | 161 |
by (blast intro: subsetI dest: subsetD) |
162 |
||
163 |
||
60770 | 164 |
subsection \<open>Rules for equality\<close> |
20140 | 165 |
|
166 |
(*Anti-symmetry of the subset relation*) |
|
58977 | 167 |
lemma subset_antisym: "\<lbrakk>A <= B; B <= A\<rbrakk> \<Longrightarrow> A = B" |
20140 | 168 |
by (blast intro: set_ext dest: subsetD) |
169 |
||
170 |
lemmas equalityI = subset_antisym |
|
171 |
||
172 |
(* Equality rules from ZF set theory -- are they appropriate here? *) |
|
58977 | 173 |
lemma equalityD1: "A = B \<Longrightarrow> A<=B" |
174 |
and equalityD2: "A = B \<Longrightarrow> B<=A" |
|
20140 | 175 |
by (simp_all add: subset_refl) |
176 |
||
58977 | 177 |
lemma equalityE: "\<lbrakk>A = B; \<lbrakk>A <= B; B <= A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
20140 | 178 |
by (simp add: subset_refl) |
179 |
||
58977 | 180 |
lemma equalityCE: "\<lbrakk>A = B; \<lbrakk>c:A; c:B\<rbrakk> \<Longrightarrow> P; \<lbrakk>\<not> c:A; \<not> c:B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
20140 | 181 |
by (blast elim: equalityE subsetCE) |
182 |
||
183 |
lemma trivial_set: "{x. x:A} = A" |
|
184 |
by (blast intro: equalityI subsetI CollectI dest: CollectD) |
|
185 |
||
186 |
||
60770 | 187 |
subsection \<open>Rules for binary union\<close> |
20140 | 188 |
|
58977 | 189 |
lemma UnI1: "c:A \<Longrightarrow> c : A Un B" |
190 |
and UnI2: "c:B \<Longrightarrow> c : A Un B" |
|
20140 | 191 |
unfolding Un_def by (blast intro: CollectI)+ |
192 |
||
193 |
(*Classical introduction rule: no commitment to A vs B*) |
|
58977 | 194 |
lemma UnCI: "(\<not>c:B \<Longrightarrow> c:A) \<Longrightarrow> c : A Un B" |
20140 | 195 |
by (blast intro: UnI1 UnI2) |
196 |
||
58977 | 197 |
lemma UnE: "\<lbrakk>c : A Un B; c:A \<Longrightarrow> P; c:B \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
20140 | 198 |
unfolding Un_def by (blast dest: CollectD) |
199 |
||
200 |
||
60770 | 201 |
subsection \<open>Rules for small intersection\<close> |
20140 | 202 |
|
58977 | 203 |
lemma IntI: "\<lbrakk>c:A; c:B\<rbrakk> \<Longrightarrow> c : A Int B" |
20140 | 204 |
unfolding Int_def by (blast intro: CollectI) |
205 |
||
58977 | 206 |
lemma IntD1: "c : A Int B \<Longrightarrow> c:A" |
207 |
and IntD2: "c : A Int B \<Longrightarrow> c:B" |
|
20140 | 208 |
unfolding Int_def by (blast dest: CollectD)+ |
209 |
||
58977 | 210 |
lemma IntE: "\<lbrakk>c : A Int B; \<lbrakk>c:A; c:B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
20140 | 211 |
by (blast dest: IntD1 IntD2) |
212 |
||
213 |
||
60770 | 214 |
subsection \<open>Rules for set complement\<close> |
20140 | 215 |
|
58977 | 216 |
lemma ComplI: "(c:A \<Longrightarrow> False) \<Longrightarrow> c : Compl(A)" |
20140 | 217 |
unfolding Compl_def by (blast intro: CollectI) |
218 |
||
219 |
(*This form, with negated conclusion, works well with the Classical prover. |
|
220 |
Negated assumptions behave like formulae on the right side of the notional |
|
221 |
turnstile...*) |
|
58977 | 222 |
lemma ComplD: "c : Compl(A) \<Longrightarrow> \<not>c:A" |
20140 | 223 |
unfolding Compl_def by (blast dest: CollectD) |
224 |
||
225 |
lemmas ComplE = ComplD [elim_format] |
|
226 |
||
227 |
||
60770 | 228 |
subsection \<open>Empty sets\<close> |
20140 | 229 |
|
230 |
lemma empty_eq: "{x. False} = {}" |
|
231 |
by (simp add: empty_def) |
|
232 |
||
58977 | 233 |
lemma emptyD: "a : {} \<Longrightarrow> P" |
20140 | 234 |
unfolding empty_def by (blast dest: CollectD) |
235 |
||
236 |
lemmas emptyE = emptyD [elim_format] |
|
237 |
||
238 |
lemma not_emptyD: |
|
58977 | 239 |
assumes "\<not> A={}" |
20140 | 240 |
shows "EX x. x:A" |
241 |
proof - |
|
242 |
have "\<not> (EX x. x:A) \<Longrightarrow> A = {}" |
|
243 |
by (rule equalityI) (blast intro!: subsetI elim!: emptyD)+ |
|
41526 | 244 |
with assms show ?thesis by blast |
20140 | 245 |
qed |
246 |
||
247 |
||
60770 | 248 |
subsection \<open>Singleton sets\<close> |
20140 | 249 |
|
250 |
lemma singletonI: "a : {a}" |
|
251 |
unfolding singleton_def by (blast intro: CollectI) |
|
252 |
||
58977 | 253 |
lemma singletonD: "b : {a} \<Longrightarrow> b=a" |
20140 | 254 |
unfolding singleton_def by (blast dest: CollectD) |
255 |
||
256 |
lemmas singletonE = singletonD [elim_format] |
|
257 |
||
258 |
||
60770 | 259 |
subsection \<open>Unions of families\<close> |
20140 | 260 |
|
261 |
(*The order of the premises presupposes that A is rigid; b may be flexible*) |
|
58977 | 262 |
lemma UN_I: "\<lbrakk>a:A; b: B(a)\<rbrakk> \<Longrightarrow> b: (UN x:A. B(x))" |
20140 | 263 |
unfolding UNION_def by (blast intro: bexI CollectI) |
264 |
||
58977 | 265 |
lemma UN_E: "\<lbrakk>b : (UN x:A. B(x)); \<And>x. \<lbrakk>x:A; b: B(x)\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
20140 | 266 |
unfolding UNION_def by (blast dest: CollectD elim: bexE) |
267 |
||
58977 | 268 |
lemma UN_cong: "\<lbrakk>A = B; \<And>x. x:B \<Longrightarrow> C(x) = D(x)\<rbrakk> \<Longrightarrow> (UN x:A. C(x)) = (UN x:B. D(x))" |
20140 | 269 |
by (simp add: UNION_def cong: bex_cong) |
270 |
||
271 |
||
60770 | 272 |
subsection \<open>Intersections of families\<close> |
20140 | 273 |
|
58977 | 274 |
lemma INT_I: "(\<And>x. x:A \<Longrightarrow> b: B(x)) \<Longrightarrow> b : (INT x:A. B(x))" |
20140 | 275 |
unfolding INTER_def by (blast intro: CollectI ballI) |
276 |
||
58977 | 277 |
lemma INT_D: "\<lbrakk>b : (INT x:A. B(x)); a:A\<rbrakk> \<Longrightarrow> b: B(a)" |
20140 | 278 |
unfolding INTER_def by (blast dest: CollectD bspec) |
279 |
||
280 |
(*"Classical" elimination rule -- does not require proving X:C *) |
|
58977 | 281 |
lemma INT_E: "\<lbrakk>b : (INT x:A. B(x)); b: B(a) \<Longrightarrow> R; \<not> a:A \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
20140 | 282 |
unfolding INTER_def by (blast dest: CollectD bspec) |
283 |
||
58977 | 284 |
lemma INT_cong: "\<lbrakk>A = B; \<And>x. x:B \<Longrightarrow> C(x) = D(x)\<rbrakk> \<Longrightarrow> (INT x:A. C(x)) = (INT x:B. D(x))" |
20140 | 285 |
by (simp add: INTER_def cong: ball_cong) |
286 |
||
287 |
||
60770 | 288 |
subsection \<open>Rules for Unions\<close> |
20140 | 289 |
|
290 |
(*The order of the premises presupposes that C is rigid; A may be flexible*) |
|
58977 | 291 |
lemma UnionI: "\<lbrakk>X:C; A:X\<rbrakk> \<Longrightarrow> A : Union(C)" |
20140 | 292 |
unfolding Union_def by (blast intro: UN_I) |
293 |
||
58977 | 294 |
lemma UnionE: "\<lbrakk>A : Union(C); \<And>X. \<lbrakk> A:X; X:C\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
20140 | 295 |
unfolding Union_def by (blast elim: UN_E) |
296 |
||
297 |
||
60770 | 298 |
subsection \<open>Rules for Inter\<close> |
20140 | 299 |
|
58977 | 300 |
lemma InterI: "(\<And>X. X:C \<Longrightarrow> A:X) \<Longrightarrow> A : Inter(C)" |
20140 | 301 |
unfolding Inter_def by (blast intro: INT_I) |
302 |
||
303 |
(*A "destruct" rule -- every X in C contains A as an element, but |
|
304 |
A:X can hold when X:C does not! This rule is analogous to "spec". *) |
|
58977 | 305 |
lemma InterD: "\<lbrakk>A : Inter(C); X:C\<rbrakk> \<Longrightarrow> A:X" |
20140 | 306 |
unfolding Inter_def by (blast dest: INT_D) |
307 |
||
308 |
(*"Classical" elimination rule -- does not require proving X:C *) |
|
58977 | 309 |
lemma InterE: "\<lbrakk>A : Inter(C); A:X \<Longrightarrow> R; \<not> X:C \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
20140 | 310 |
unfolding Inter_def by (blast elim: INT_E) |
311 |
||
312 |
||
60770 | 313 |
section \<open>Derived rules involving subsets; Union and Intersection as lattice operations\<close> |
20140 | 314 |
|
60770 | 315 |
subsection \<open>Big Union -- least upper bound of a set\<close> |
20140 | 316 |
|
58977 | 317 |
lemma Union_upper: "B:A \<Longrightarrow> B <= Union(A)" |
20140 | 318 |
by (blast intro: subsetI UnionI) |
319 |
||
58977 | 320 |
lemma Union_least: "(\<And>X. X:A \<Longrightarrow> X<=C) \<Longrightarrow> Union(A) <= C" |
20140 | 321 |
by (blast intro: subsetI dest: subsetD elim: UnionE) |
322 |
||
323 |
||
60770 | 324 |
subsection \<open>Big Intersection -- greatest lower bound of a set\<close> |
20140 | 325 |
|
58977 | 326 |
lemma Inter_lower: "B:A \<Longrightarrow> Inter(A) <= B" |
20140 | 327 |
by (blast intro: subsetI dest: InterD) |
328 |
||
58977 | 329 |
lemma Inter_greatest: "(\<And>X. X:A \<Longrightarrow> C<=X) \<Longrightarrow> C <= Inter(A)" |
20140 | 330 |
by (blast intro: subsetI InterI dest: subsetD) |
331 |
||
332 |
||
60770 | 333 |
subsection \<open>Finite Union -- the least upper bound of 2 sets\<close> |
20140 | 334 |
|
335 |
lemma Un_upper1: "A <= A Un B" |
|
336 |
by (blast intro: subsetI UnI1) |
|
337 |
||
338 |
lemma Un_upper2: "B <= A Un B" |
|
339 |
by (blast intro: subsetI UnI2) |
|
340 |
||
58977 | 341 |
lemma Un_least: "\<lbrakk>A<=C; B<=C\<rbrakk> \<Longrightarrow> A Un B <= C" |
20140 | 342 |
by (blast intro: subsetI elim: UnE dest: subsetD) |
343 |
||
344 |
||
60770 | 345 |
subsection \<open>Finite Intersection -- the greatest lower bound of 2 sets\<close> |
20140 | 346 |
|
347 |
lemma Int_lower1: "A Int B <= A" |
|
348 |
by (blast intro: subsetI elim: IntE) |
|
349 |
||
350 |
lemma Int_lower2: "A Int B <= B" |
|
351 |
by (blast intro: subsetI elim: IntE) |
|
352 |
||
58977 | 353 |
lemma Int_greatest: "\<lbrakk>C<=A; C<=B\<rbrakk> \<Longrightarrow> C <= A Int B" |
20140 | 354 |
by (blast intro: subsetI IntI dest: subsetD) |
355 |
||
356 |
||
60770 | 357 |
subsection \<open>Monotonicity\<close> |
20140 | 358 |
|
58977 | 359 |
lemma monoI: "(\<And>A B. A <= B \<Longrightarrow> f(A) <= f(B)) \<Longrightarrow> mono(f)" |
20140 | 360 |
unfolding mono_def by blast |
361 |
||
58977 | 362 |
lemma monoD: "\<lbrakk>mono(f); A <= B\<rbrakk> \<Longrightarrow> f(A) <= f(B)" |
20140 | 363 |
unfolding mono_def by blast |
364 |
||
58977 | 365 |
lemma mono_Un: "mono(f) \<Longrightarrow> f(A) Un f(B) <= f(A Un B)" |
20140 | 366 |
by (blast intro: Un_least dest: monoD intro: Un_upper1 Un_upper2) |
367 |
||
58977 | 368 |
lemma mono_Int: "mono(f) \<Longrightarrow> f(A Int B) <= f(A) Int f(B)" |
20140 | 369 |
by (blast intro: Int_greatest dest: monoD intro: Int_lower1 Int_lower2) |
370 |
||
371 |
||
60770 | 372 |
subsection \<open>Automated reasoning setup\<close> |
20140 | 373 |
|
374 |
lemmas [intro!] = ballI subsetI InterI INT_I CollectI ComplI IntI UnCI singletonI |
|
375 |
and [intro] = bexI UnionI UN_I |
|
376 |
and [elim!] = bexE UnionE UN_E CollectE ComplE IntE UnE emptyE singletonE |
|
377 |
and [elim] = ballE InterD InterE INT_D INT_E subsetD subsetCE |
|
378 |
||
379 |
lemma mem_rews: |
|
58977 | 380 |
"(a : A Un B) \<longleftrightarrow> (a:A | a:B)" |
381 |
"(a : A Int B) \<longleftrightarrow> (a:A \<and> a:B)" |
|
382 |
"(a : Compl(B)) \<longleftrightarrow> (\<not>a:B)" |
|
383 |
"(a : {b}) \<longleftrightarrow> (a=b)" |
|
384 |
"(a : {}) \<longleftrightarrow> False" |
|
385 |
"(a : {x. P(x)}) \<longleftrightarrow> P(a)" |
|
20140 | 386 |
by blast+ |
387 |
||
388 |
lemmas [simp] = trivial_set empty_eq mem_rews |
|
389 |
and [cong] = ball_cong bex_cong INT_cong UN_cong |
|
390 |
||
391 |
||
60770 | 392 |
section \<open>Equalities involving union, intersection, inclusion, etc.\<close> |
20140 | 393 |
|
60770 | 394 |
subsection \<open>Binary Intersection\<close> |
20140 | 395 |
|
396 |
lemma Int_absorb: "A Int A = A" |
|
397 |
by (blast intro: equalityI) |
|
398 |
||
399 |
lemma Int_commute: "A Int B = B Int A" |
|
400 |
by (blast intro: equalityI) |
|
401 |
||
402 |
lemma Int_assoc: "(A Int B) Int C = A Int (B Int C)" |
|
403 |
by (blast intro: equalityI) |
|
404 |
||
405 |
lemma Int_Un_distrib: "(A Un B) Int C = (A Int C) Un (B Int C)" |
|
406 |
by (blast intro: equalityI) |
|
407 |
||
58977 | 408 |
lemma subset_Int_eq: "(A<=B) \<longleftrightarrow> (A Int B = A)" |
20140 | 409 |
by (blast intro: equalityI elim: equalityE) |
410 |
||
411 |
||
60770 | 412 |
subsection \<open>Binary Union\<close> |
20140 | 413 |
|
414 |
lemma Un_absorb: "A Un A = A" |
|
415 |
by (blast intro: equalityI) |
|
416 |
||
417 |
lemma Un_commute: "A Un B = B Un A" |
|
418 |
by (blast intro: equalityI) |
|
419 |
||
420 |
lemma Un_assoc: "(A Un B) Un C = A Un (B Un C)" |
|
421 |
by (blast intro: equalityI) |
|
422 |
||
423 |
lemma Un_Int_distrib: "(A Int B) Un C = (A Un C) Int (B Un C)" |
|
424 |
by (blast intro: equalityI) |
|
425 |
||
426 |
lemma Un_Int_crazy: |
|
427 |
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)" |
|
428 |
by (blast intro: equalityI) |
|
429 |
||
58977 | 430 |
lemma subset_Un_eq: "(A<=B) \<longleftrightarrow> (A Un B = B)" |
20140 | 431 |
by (blast intro: equalityI elim: equalityE) |
432 |
||
433 |
||
62020 | 434 |
subsection \<open>Simple properties of \<open>Compl\<close> -- complement of a set\<close> |
20140 | 435 |
|
436 |
lemma Compl_disjoint: "A Int Compl(A) = {x. False}" |
|
437 |
by (blast intro: equalityI) |
|
438 |
||
439 |
lemma Compl_partition: "A Un Compl(A) = {x. True}" |
|
440 |
by (blast intro: equalityI) |
|
441 |
||
442 |
lemma double_complement: "Compl(Compl(A)) = A" |
|
443 |
by (blast intro: equalityI) |
|
444 |
||
445 |
lemma Compl_Un: "Compl(A Un B) = Compl(A) Int Compl(B)" |
|
446 |
by (blast intro: equalityI) |
|
447 |
||
448 |
lemma Compl_Int: "Compl(A Int B) = Compl(A) Un Compl(B)" |
|
449 |
by (blast intro: equalityI) |
|
450 |
||
451 |
lemma Compl_UN: "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))" |
|
452 |
by (blast intro: equalityI) |
|
453 |
||
454 |
lemma Compl_INT: "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))" |
|
455 |
by (blast intro: equalityI) |
|
456 |
||
457 |
(*Halmos, Naive Set Theory, page 16.*) |
|
58977 | 458 |
lemma Un_Int_assoc_eq: "((A Int B) Un C = A Int (B Un C)) \<longleftrightarrow> (C<=A)" |
20140 | 459 |
by (blast intro: equalityI elim: equalityE) |
460 |
||
461 |
||
60770 | 462 |
subsection \<open>Big Union and Intersection\<close> |
20140 | 463 |
|
464 |
lemma Union_Un_distrib: "Union(A Un B) = Union(A) Un Union(B)" |
|
465 |
by (blast intro: equalityI) |
|
466 |
||
467 |
lemma Union_disjoint: |
|
58977 | 468 |
"(Union(C) Int A = {x. False}) \<longleftrightarrow> (ALL B:C. B Int A = {x. False})" |
20140 | 469 |
by (blast intro: equalityI elim: equalityE) |
470 |
||
471 |
lemma Inter_Un_distrib: "Inter(A Un B) = Inter(A) Int Inter(B)" |
|
472 |
by (blast intro: equalityI) |
|
473 |
||
474 |
||
60770 | 475 |
subsection \<open>Unions and Intersections of Families\<close> |
20140 | 476 |
|
477 |
lemma UN_eq: "(UN x:A. B(x)) = Union({Y. EX x:A. Y=B(x)})" |
|
478 |
by (blast intro: equalityI) |
|
479 |
||
480 |
(*Look: it has an EXISTENTIAL quantifier*) |
|
481 |
lemma INT_eq: "(INT x:A. B(x)) = Inter({Y. EX x:A. Y=B(x)})" |
|
482 |
by (blast intro: equalityI) |
|
483 |
||
484 |
lemma Int_Union_image: "A Int Union(B) = (UN C:B. A Int C)" |
|
485 |
by (blast intro: equalityI) |
|
486 |
||
487 |
lemma Un_Inter_image: "A Un Inter(B) = (INT C:B. A Un C)" |
|
488 |
by (blast intro: equalityI) |
|
489 |
||
490 |
||
60770 | 491 |
section \<open>Monotonicity of various operations\<close> |
20140 | 492 |
|
58977 | 493 |
lemma Union_mono: "A<=B \<Longrightarrow> Union(A) <= Union(B)" |
20140 | 494 |
by blast |
495 |
||
58977 | 496 |
lemma Inter_anti_mono: "B <= A \<Longrightarrow> Inter(A) <= Inter(B)" |
20140 | 497 |
by blast |
498 |
||
58977 | 499 |
lemma UN_mono: "\<lbrakk>A <= B; \<And>x. x:A \<Longrightarrow> f(x)<=g(x)\<rbrakk> \<Longrightarrow> (UN x:A. f(x)) <= (UN x:B. g(x))" |
20140 | 500 |
by blast |
501 |
||
58977 | 502 |
lemma INT_anti_mono: "\<lbrakk>B <= A; \<And>x. x:A \<Longrightarrow> f(x) <= g(x)\<rbrakk> \<Longrightarrow> (INT x:A. f(x)) <= (INT x:A. g(x))" |
20140 | 503 |
by blast |
504 |
||
58977 | 505 |
lemma Un_mono: "\<lbrakk>A <= C; B <= D\<rbrakk> \<Longrightarrow> A Un B <= C Un D" |
20140 | 506 |
by blast |
507 |
||
58977 | 508 |
lemma Int_mono: "\<lbrakk>A <= C; B <= D\<rbrakk> \<Longrightarrow> A Int B <= C Int D" |
20140 | 509 |
by blast |
510 |
||
58977 | 511 |
lemma Compl_anti_mono: "A <= B \<Longrightarrow> Compl(B) <= Compl(A)" |
20140 | 512 |
by blast |
0 | 513 |
|
514 |
end |