1461
|
1 |
(* Title: ZF/subset
|
0
|
2 |
ID: $Id$
|
1461
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
0
|
4 |
Copyright 1991 University of Cambridge
|
|
5 |
|
|
6 |
Derived rules involving subsets
|
|
7 |
Union and Intersection as lattice operations
|
|
8 |
*)
|
|
9 |
|
|
10 |
(*** cons ***)
|
|
11 |
|
2877
|
12 |
qed_goal "cons_subsetI" ZF.thy "!!a. [| a:C; B<=C |] ==> cons(a,B) <= C"
|
|
13 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
14 |
|
2877
|
15 |
qed_goal "subset_consI" ZF.thy "B <= cons(a,B)"
|
|
16 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
17 |
|
2877
|
18 |
qed_goal "cons_subset_iff" ZF.thy "cons(a,B)<=C <-> a:C & B<=C"
|
|
19 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
20 |
|
|
21 |
(*A safe special case of subset elimination, adding no new variables
|
|
22 |
[| cons(a,B) <= C; [| a : C; B <= C |] ==> R |] ==> R *)
|
2469
|
23 |
bind_thm ("cons_subsetE", cons_subset_iff RS iffD1 RS conjE);
|
0
|
24 |
|
2877
|
25 |
qed_goal "subset_empty_iff" ZF.thy "A<=0 <-> A=0"
|
|
26 |
(fn _=> [ (Blast_tac 1) ]);
|
0
|
27 |
|
2877
|
28 |
qed_goal "subset_cons_iff" ZF.thy
|
0
|
29 |
"C<=cons(a,B) <-> C<=B | (a:C & C-{a} <= B)"
|
2877
|
30 |
(fn _=> [ (Blast_tac 1) ]);
|
0
|
31 |
|
|
32 |
(*** succ ***)
|
|
33 |
|
2877
|
34 |
qed_goal "subset_succI" ZF.thy "i <= succ(i)"
|
|
35 |
(fn _=> [ (Blast_tac 1) ]);
|
0
|
36 |
|
|
37 |
(*But if j is an ordinal or is transitive, then i:j implies i<=j!
|
|
38 |
See ordinal/Ord_succ_subsetI*)
|
2877
|
39 |
qed_goalw "succ_subsetI" ZF.thy [succ_def]
|
2469
|
40 |
"!!i j. [| i:j; i<=j |] ==> succ(i)<=j"
|
2877
|
41 |
(fn _=> [ (Blast_tac 1) ]);
|
0
|
42 |
|
2877
|
43 |
qed_goalw "succ_subsetE" ZF.thy [succ_def]
|
0
|
44 |
"[| succ(i) <= j; [| i:j; i<=j |] ==> P \
|
|
45 |
\ |] ==> P"
|
|
46 |
(fn major::prems=>
|
|
47 |
[ (rtac (major RS cons_subsetE) 1),
|
|
48 |
(REPEAT (ares_tac prems 1)) ]);
|
|
49 |
|
|
50 |
(*** singletons ***)
|
|
51 |
|
2877
|
52 |
qed_goal "singleton_subsetI" ZF.thy "!!a c. a:C ==> {a} <= C"
|
|
53 |
(fn _=> [ (Blast_tac 1) ]);
|
0
|
54 |
|
2877
|
55 |
qed_goal "singleton_subsetD" ZF.thy "!!a C. {a} <= C ==> a:C"
|
|
56 |
(fn _=> [ (Blast_tac 1) ]);
|
0
|
57 |
|
|
58 |
(*** Big Union -- least upper bound of a set ***)
|
|
59 |
|
2877
|
60 |
qed_goal "Union_subset_iff" ZF.thy "Union(A) <= C <-> (ALL x:A. x <= C)"
|
|
61 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
62 |
|
2877
|
63 |
qed_goal "Union_upper" ZF.thy
|
2469
|
64 |
"!!B A. B:A ==> B <= Union(A)"
|
2877
|
65 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
66 |
|
2877
|
67 |
qed_goal "Union_least" ZF.thy
|
0
|
68 |
"[| !!x. x:A ==> x<=C |] ==> Union(A) <= C"
|
|
69 |
(fn [prem]=>
|
|
70 |
[ (rtac (ballI RS (Union_subset_iff RS iffD2)) 1),
|
|
71 |
(etac prem 1) ]);
|
|
72 |
|
|
73 |
(*** Union of a family of sets ***)
|
|
74 |
|
2877
|
75 |
goal ZF.thy "A <= (UN i:I. B(i)) <-> A = (UN i:I. A Int B(i))";
|
|
76 |
by (blast_tac (!claset addSEs [equalityE]) 1);
|
760
|
77 |
qed "subset_UN_iff_eq";
|
0
|
78 |
|
2877
|
79 |
qed_goal "UN_subset_iff" ZF.thy
|
0
|
80 |
"(UN x:A.B(x)) <= C <-> (ALL x:A. B(x) <= C)"
|
2877
|
81 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
82 |
|
2877
|
83 |
qed_goal "UN_upper" ZF.thy
|
0
|
84 |
"!!x A. x:A ==> B(x) <= (UN x:A.B(x))"
|
|
85 |
(fn _ => [ etac (RepFunI RS Union_upper) 1 ]);
|
|
86 |
|
2877
|
87 |
qed_goal "UN_least" ZF.thy
|
0
|
88 |
"[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A.B(x)) <= C"
|
|
89 |
(fn [prem]=>
|
|
90 |
[ (rtac (ballI RS (UN_subset_iff RS iffD2)) 1),
|
|
91 |
(etac prem 1) ]);
|
|
92 |
|
|
93 |
|
|
94 |
(*** Big Intersection -- greatest lower bound of a nonempty set ***)
|
|
95 |
|
2877
|
96 |
qed_goal "Inter_subset_iff" ZF.thy
|
0
|
97 |
"!!a A. a: A ==> C <= Inter(A) <-> (ALL x:A. C <= x)"
|
2877
|
98 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
99 |
|
2877
|
100 |
qed_goal "Inter_lower" ZF.thy "!!B A. B:A ==> Inter(A) <= B"
|
|
101 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
102 |
|
2877
|
103 |
qed_goal "Inter_greatest" ZF.thy
|
0
|
104 |
"[| a:A; !!x. x:A ==> C<=x |] ==> C <= Inter(A)"
|
|
105 |
(fn [prem1,prem2]=>
|
|
106 |
[ (rtac ([prem1, ballI] MRS (Inter_subset_iff RS iffD2)) 1),
|
|
107 |
(etac prem2 1) ]);
|
|
108 |
|
|
109 |
(*** Intersection of a family of sets ***)
|
|
110 |
|
2877
|
111 |
qed_goal "INT_lower" ZF.thy
|
2469
|
112 |
"!!x. x:A ==> (INT x:A.B(x)) <= B(x)"
|
2877
|
113 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
114 |
|
2877
|
115 |
qed_goal "INT_greatest" ZF.thy
|
0
|
116 |
"[| a:A; !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A.B(x))"
|
|
117 |
(fn [nonempty,prem] =>
|
|
118 |
[ rtac (nonempty RS RepFunI RS Inter_greatest) 1,
|
|
119 |
REPEAT (eresolve_tac [RepFunE, prem, ssubst] 1) ]);
|
|
120 |
|
|
121 |
|
|
122 |
(*** Finite Union -- the least upper bound of 2 sets ***)
|
|
123 |
|
2877
|
124 |
qed_goal "Un_subset_iff" ZF.thy "A Un B <= C <-> A <= C & B <= C"
|
|
125 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
126 |
|
2877
|
127 |
qed_goal "Un_upper1" ZF.thy "A <= A Un B"
|
|
128 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
129 |
|
2877
|
130 |
qed_goal "Un_upper2" ZF.thy "B <= A Un B"
|
|
131 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
132 |
|
2877
|
133 |
qed_goal "Un_least" ZF.thy "!!A B C. [| A<=C; B<=C |] ==> A Un B <= C"
|
|
134 |
(fn _ => [ Blast_tac 1 ]);
|
2469
|
135 |
|
0
|
136 |
|
|
137 |
(*** Finite Intersection -- the greatest lower bound of 2 sets *)
|
|
138 |
|
2877
|
139 |
qed_goal "Int_subset_iff" ZF.thy "C <= A Int B <-> C <= A & C <= B"
|
|
140 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
141 |
|
2877
|
142 |
qed_goal "Int_lower1" ZF.thy "A Int B <= A"
|
|
143 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
144 |
|
2877
|
145 |
qed_goal "Int_lower2" ZF.thy "A Int B <= B"
|
|
146 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
147 |
|
2877
|
148 |
qed_goal "Int_greatest" ZF.thy "!!A B C. [| C<=A; C<=B |] ==> C <= A Int B"
|
|
149 |
(fn _ => [ Blast_tac 1 ]);
|
2469
|
150 |
|
0
|
151 |
|
|
152 |
(*** Set difference *)
|
|
153 |
|
2877
|
154 |
qed_goal "Diff_subset" ZF.thy "A-B <= A"
|
|
155 |
(fn _ => [ Blast_tac 1 ]);
|
0
|
156 |
|
2877
|
157 |
qed_goal "Diff_contains" ZF.thy
|
2469
|
158 |
"!!C. [| C<=A; C Int B = 0 |] ==> C <= A-B"
|
2877
|
159 |
(fn _ => [ (blast_tac (!claset addSEs [equalityE]) 1) ]);
|
2469
|
160 |
|
0
|
161 |
|
|
162 |
(** Collect **)
|
|
163 |
|
2877
|
164 |
qed_goal "Collect_subset" ZF.thy "Collect(A,P) <= A"
|
|
165 |
(fn _ => [ Blast_tac 1 ]);
|
2469
|
166 |
|
0
|
167 |
|
|
168 |
(** RepFun **)
|
|
169 |
|
2877
|
170 |
val prems = goal ZF.thy "[| !!x. x:A ==> f(x): B |] ==> {f(x). x:A} <= B";
|
|
171 |
by (blast_tac (!claset addIs prems) 1);
|
760
|
172 |
qed "RepFun_subset";
|
689
|
173 |
|
2469
|
174 |
val subset_SIs =
|
|
175 |
[subset_refl, cons_subsetI, subset_consI,
|
|
176 |
Union_least, UN_least, Un_least,
|
|
177 |
Inter_greatest, Int_greatest, RepFun_subset,
|
|
178 |
Un_upper1, Un_upper2, Int_lower1, Int_lower2];
|
|
179 |
|
689
|
180 |
|
2469
|
181 |
(*A claset for subset reasoning*)
|
|
182 |
val subset_cs = !claset
|
|
183 |
delrules [subsetI, subsetCE]
|
|
184 |
addSIs subset_SIs
|
|
185 |
addIs [Union_upper, Inter_lower]
|
|
186 |
addSEs [cons_subsetE];
|
|
187 |
|