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