| author | wenzelm | 
| Wed, 12 Feb 1997 15:43:50 +0100 | |
| changeset 2607 | a224a2865e05 | 
| parent 2519 | 761e3094e32f | 
| child 2891 | d8f254ad1ab9 | 
| permissions | -rw-r--r-- | 
| 1465 | 1  | 
(* Title: HOL/equalities  | 
| 923 | 2  | 
ID: $Id$  | 
| 1465 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 923 | 4  | 
Copyright 1994 University of Cambridge  | 
5  | 
||
6  | 
Equalities involving union, intersection, inclusion, etc.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
writeln"File HOL/equalities";  | 
|
10  | 
||
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
11  | 
AddSIs [equalityI];  | 
| 
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
12  | 
|
| 1548 | 13  | 
section "{}";
 | 
14  | 
||
| 1531 | 15  | 
goal Set.thy "{x.False} = {}";
 | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
16  | 
by (Fast_tac 1);  | 
| 1531 | 17  | 
qed "Collect_False_empty";  | 
18  | 
Addsimps [Collect_False_empty];  | 
|
19  | 
||
20  | 
goal Set.thy "(A <= {}) = (A = {})";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
21  | 
by (Fast_tac 1);  | 
| 1531 | 22  | 
qed "subset_empty";  | 
23  | 
Addsimps [subset_empty];  | 
|
24  | 
||
| 1548 | 25  | 
section "insert";  | 
| 923 | 26  | 
|
| 1531 | 27  | 
(*NOT SUITABLE FOR REWRITING since {a} == insert a {}*)
 | 
28  | 
goal Set.thy "insert a A = {a} Un A";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
29  | 
by (Fast_tac 1);  | 
| 1531 | 30  | 
qed "insert_is_Un";  | 
31  | 
||
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
32  | 
goal Set.thy "insert a A ~= {}";
 | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
33  | 
by (fast_tac (!claset addEs [equalityCE]) 1);  | 
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
34  | 
qed"insert_not_empty";  | 
| 1531 | 35  | 
Addsimps[insert_not_empty];  | 
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
36  | 
|
| 
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
37  | 
bind_thm("empty_not_insert",insert_not_empty RS not_sym);
 | 
| 1531 | 38  | 
Addsimps[empty_not_insert];  | 
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
39  | 
|
| 923 | 40  | 
goal Set.thy "!!a. a:A ==> insert a A = A";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
41  | 
by (Fast_tac 1);  | 
| 923 | 42  | 
qed "insert_absorb";  | 
43  | 
||
| 1531 | 44  | 
goal Set.thy "insert x (insert x A) = insert x A";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
45  | 
by (Fast_tac 1);  | 
| 1531 | 46  | 
qed "insert_absorb2";  | 
47  | 
Addsimps [insert_absorb2];  | 
|
48  | 
||
| 1879 | 49  | 
goal Set.thy "insert x (insert y A) = insert y (insert x A)";  | 
50  | 
by (Fast_tac 1);  | 
|
51  | 
qed "insert_commute";  | 
|
52  | 
||
| 923 | 53  | 
goal Set.thy "(insert x A <= B) = (x:B & A <= B)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
54  | 
by (Fast_tac 1);  | 
| 923 | 55  | 
qed "insert_subset";  | 
| 1531 | 56  | 
Addsimps[insert_subset];  | 
57  | 
||
58  | 
(* use new B rather than (A-{a}) to avoid infinite unfolding *)
 | 
|
59  | 
goal Set.thy "!!a. a:A ==> ? B. A = insert a B & a ~: B";  | 
|
| 1553 | 60  | 
by (res_inst_tac [("x","A-{a}")] exI 1);
 | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
61  | 
by (Fast_tac 1);  | 
| 1531 | 62  | 
qed "mk_disjoint_insert";  | 
| 923 | 63  | 
|
| 
1843
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
64  | 
goal Set.thy  | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
65  | 
    "!!A. A~={} ==> (UN x:A. insert a (B x)) = insert a (UN x:A. B x)";
 | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
66  | 
by (Fast_tac 1);  | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
67  | 
qed "UN_insert_distrib";  | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
68  | 
|
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
69  | 
goal Set.thy "(UN x. insert a (B x)) = insert a (UN x. B x)";  | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
70  | 
by (Fast_tac 1);  | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
71  | 
qed "UN1_insert_distrib";  | 
| 
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
72  | 
|
| 1660 | 73  | 
section "``";  | 
| 923 | 74  | 
|
75  | 
goal Set.thy "f``{} = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
76  | 
by (Fast_tac 1);  | 
| 923 | 77  | 
qed "image_empty";  | 
| 1531 | 78  | 
Addsimps[image_empty];  | 
| 923 | 79  | 
|
80  | 
goal Set.thy "f``insert a B = insert (f a) (f``B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
81  | 
by (Fast_tac 1);  | 
| 923 | 82  | 
qed "image_insert";  | 
| 1531 | 83  | 
Addsimps[image_insert];  | 
| 923 | 84  | 
|
| 1660 | 85  | 
qed_goal "ball_image" Set.thy "(!y:F``S. P y) = (!x:S. P (F x))"  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
86  | 
(fn _ => [Fast_tac 1]);  | 
| 1660 | 87  | 
|
| 1884 | 88  | 
goal Set.thy "!!x. x:A ==> insert (f x) (f``A) = f``A";  | 
89  | 
by (Fast_tac 1);  | 
|
90  | 
qed "insert_image";  | 
|
91  | 
Addsimps [insert_image];  | 
|
92  | 
||
| 1748 | 93  | 
goalw Set.thy [image_def]  | 
| 1763 | 94  | 
"(%x. if P x then f x else g x) `` S \  | 
| 1748 | 95  | 
\ = (f `` ({x.x:S & P x})) Un (g `` ({x.x:S & ~(P x)}))";
 | 
| 2031 | 96  | 
by (split_tac [expand_if] 1);  | 
97  | 
by (Fast_tac 1);  | 
|
| 1748 | 98  | 
qed "if_image_distrib";  | 
99  | 
Addsimps[if_image_distrib];  | 
|
100  | 
||
101  | 
||
| 1660 | 102  | 
section "range";  | 
103  | 
||
104  | 
qed_goal "ball_range" Set.thy "(!y:range f. P y) = (!x. P (f x))"  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
105  | 
(fn _ => [Fast_tac 1]);  | 
| 1660 | 106  | 
|
| 1884 | 107  | 
qed_goalw "image_range" Set.thy [image_def]  | 
108  | 
"f``range g = range (%x. f (g x))"  | 
|
109  | 
(fn _ => [rtac Collect_cong 1, Fast_tac 1]);  | 
|
| 1660 | 110  | 
|
| 1548 | 111  | 
section "Int";  | 
| 923 | 112  | 
|
113  | 
goal Set.thy "A Int A = A";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
114  | 
by (Fast_tac 1);  | 
| 923 | 115  | 
qed "Int_absorb";  | 
| 1531 | 116  | 
Addsimps[Int_absorb];  | 
| 923 | 117  | 
|
118  | 
goal Set.thy "A Int B = B Int A";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
119  | 
by (Fast_tac 1);  | 
| 923 | 120  | 
qed "Int_commute";  | 
121  | 
||
122  | 
goal Set.thy "(A Int B) Int C = A Int (B Int C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
123  | 
by (Fast_tac 1);  | 
| 923 | 124  | 
qed "Int_assoc";  | 
125  | 
||
126  | 
goal Set.thy "{} Int B = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
127  | 
by (Fast_tac 1);  | 
| 923 | 128  | 
qed "Int_empty_left";  | 
| 1531 | 129  | 
Addsimps[Int_empty_left];  | 
| 923 | 130  | 
|
131  | 
goal Set.thy "A Int {} = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
132  | 
by (Fast_tac 1);  | 
| 923 | 133  | 
qed "Int_empty_right";  | 
| 1531 | 134  | 
Addsimps[Int_empty_right];  | 
135  | 
||
136  | 
goal Set.thy "UNIV Int B = B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
137  | 
by (Fast_tac 1);  | 
| 1531 | 138  | 
qed "Int_UNIV_left";  | 
139  | 
Addsimps[Int_UNIV_left];  | 
|
140  | 
||
141  | 
goal Set.thy "A Int UNIV = A";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
142  | 
by (Fast_tac 1);  | 
| 1531 | 143  | 
qed "Int_UNIV_right";  | 
144  | 
Addsimps[Int_UNIV_right];  | 
|
| 923 | 145  | 
|
146  | 
goal Set.thy "A Int (B Un C) = (A Int B) Un (A Int C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
147  | 
by (Fast_tac 1);  | 
| 923 | 148  | 
qed "Int_Un_distrib";  | 
149  | 
||
| 1618 | 150  | 
goal Set.thy "(B Un C) Int A = (B Int A) Un (C Int A)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
151  | 
by (Fast_tac 1);  | 
| 1618 | 152  | 
qed "Int_Un_distrib2";  | 
153  | 
||
| 923 | 154  | 
goal Set.thy "(A<=B) = (A Int B = A)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
155  | 
by (fast_tac (!claset addSEs [equalityE]) 1);  | 
| 923 | 156  | 
qed "subset_Int_eq";  | 
157  | 
||
| 1531 | 158  | 
goal Set.thy "(A Int B = UNIV) = (A = UNIV & B = UNIV)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
159  | 
by (fast_tac (!claset addEs [equalityCE]) 1);  | 
| 1531 | 160  | 
qed "Int_UNIV";  | 
161  | 
Addsimps[Int_UNIV];  | 
|
162  | 
||
| 1548 | 163  | 
section "Un";  | 
| 923 | 164  | 
|
165  | 
goal Set.thy "A Un A = A";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
166  | 
by (Fast_tac 1);  | 
| 923 | 167  | 
qed "Un_absorb";  | 
| 1531 | 168  | 
Addsimps[Un_absorb];  | 
| 923 | 169  | 
|
170  | 
goal Set.thy "A Un B = B Un A";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
171  | 
by (Fast_tac 1);  | 
| 923 | 172  | 
qed "Un_commute";  | 
173  | 
||
174  | 
goal Set.thy "(A Un B) Un C = A Un (B Un C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
175  | 
by (Fast_tac 1);  | 
| 923 | 176  | 
qed "Un_assoc";  | 
177  | 
||
178  | 
goal Set.thy "{} Un B = B";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
179  | 
by (Fast_tac 1);  | 
| 923 | 180  | 
qed "Un_empty_left";  | 
| 1531 | 181  | 
Addsimps[Un_empty_left];  | 
| 923 | 182  | 
|
183  | 
goal Set.thy "A Un {} = A";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
184  | 
by (Fast_tac 1);  | 
| 923 | 185  | 
qed "Un_empty_right";  | 
| 1531 | 186  | 
Addsimps[Un_empty_right];  | 
187  | 
||
188  | 
goal Set.thy "UNIV Un B = UNIV";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
189  | 
by (Fast_tac 1);  | 
| 1531 | 190  | 
qed "Un_UNIV_left";  | 
191  | 
Addsimps[Un_UNIV_left];  | 
|
192  | 
||
193  | 
goal Set.thy "A Un UNIV = UNIV";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
194  | 
by (Fast_tac 1);  | 
| 1531 | 195  | 
qed "Un_UNIV_right";  | 
196  | 
Addsimps[Un_UNIV_right];  | 
|
| 923 | 197  | 
|
| 
1843
 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 
paulson 
parents: 
1786 
diff
changeset
 | 
198  | 
goal Set.thy "(insert a B) Un C = insert a (B Un C)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
199  | 
by (Fast_tac 1);  | 
| 923 | 200  | 
qed "Un_insert_left";  | 
201  | 
||
| 1917 | 202  | 
goal Set.thy "A Un (insert a B) = insert a (A Un B)";  | 
203  | 
by (Fast_tac 1);  | 
|
204  | 
qed "Un_insert_right";  | 
|
205  | 
||
| 923 | 206  | 
goal Set.thy "(A Int B) Un C = (A Un C) Int (B Un C)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
207  | 
by (Fast_tac 1);  | 
| 923 | 208  | 
qed "Un_Int_distrib";  | 
209  | 
||
210  | 
goal Set.thy  | 
|
211  | 
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
212  | 
by (Fast_tac 1);  | 
| 923 | 213  | 
qed "Un_Int_crazy";  | 
214  | 
||
215  | 
goal Set.thy "(A<=B) = (A Un B = B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
216  | 
by (fast_tac (!claset addSEs [equalityE]) 1);  | 
| 923 | 217  | 
qed "subset_Un_eq";  | 
218  | 
||
219  | 
goal Set.thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
220  | 
by (Fast_tac 1);  | 
| 923 | 221  | 
qed "subset_insert_iff";  | 
222  | 
||
223  | 
goal Set.thy "(A Un B = {}) = (A = {} & B = {})";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
224  | 
by (fast_tac (!claset addEs [equalityCE]) 1);  | 
| 923 | 225  | 
qed "Un_empty";  | 
| 1531 | 226  | 
Addsimps[Un_empty];  | 
| 923 | 227  | 
|
| 1548 | 228  | 
section "Compl";  | 
| 923 | 229  | 
|
230  | 
goal Set.thy "A Int Compl(A) = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
231  | 
by (Fast_tac 1);  | 
| 923 | 232  | 
qed "Compl_disjoint";  | 
| 1531 | 233  | 
Addsimps[Compl_disjoint];  | 
| 923 | 234  | 
|
| 1531 | 235  | 
goal Set.thy "A Un Compl(A) = UNIV";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
236  | 
by (Fast_tac 1);  | 
| 923 | 237  | 
qed "Compl_partition";  | 
238  | 
||
239  | 
goal Set.thy "Compl(Compl(A)) = A";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
240  | 
by (Fast_tac 1);  | 
| 923 | 241  | 
qed "double_complement";  | 
| 1531 | 242  | 
Addsimps[double_complement];  | 
| 923 | 243  | 
|
244  | 
goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
245  | 
by (Fast_tac 1);  | 
| 923 | 246  | 
qed "Compl_Un";  | 
247  | 
||
248  | 
goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
249  | 
by (Fast_tac 1);  | 
| 923 | 250  | 
qed "Compl_Int";  | 
251  | 
||
252  | 
goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
253  | 
by (Fast_tac 1);  | 
| 923 | 254  | 
qed "Compl_UN";  | 
255  | 
||
256  | 
goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
257  | 
by (Fast_tac 1);  | 
| 923 | 258  | 
qed "Compl_INT";  | 
259  | 
||
260  | 
(*Halmos, Naive Set Theory, page 16.*)  | 
|
261  | 
||
262  | 
goal Set.thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
263  | 
by (fast_tac (!claset addSEs [equalityE]) 1);  | 
| 923 | 264  | 
qed "Un_Int_assoc_eq";  | 
265  | 
||
266  | 
||
| 1548 | 267  | 
section "Union";  | 
| 923 | 268  | 
|
269  | 
goal Set.thy "Union({}) = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
270  | 
by (Fast_tac 1);  | 
| 923 | 271  | 
qed "Union_empty";  | 
| 1531 | 272  | 
Addsimps[Union_empty];  | 
273  | 
||
274  | 
goal Set.thy "Union(UNIV) = UNIV";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
275  | 
by (Fast_tac 1);  | 
| 1531 | 276  | 
qed "Union_UNIV";  | 
277  | 
Addsimps[Union_UNIV];  | 
|
| 923 | 278  | 
|
279  | 
goal Set.thy "Union(insert a B) = a Un Union(B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
280  | 
by (Fast_tac 1);  | 
| 923 | 281  | 
qed "Union_insert";  | 
| 1531 | 282  | 
Addsimps[Union_insert];  | 
| 923 | 283  | 
|
284  | 
goal Set.thy "Union(A Un B) = Union(A) Un Union(B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
285  | 
by (Fast_tac 1);  | 
| 923 | 286  | 
qed "Union_Un_distrib";  | 
| 1531 | 287  | 
Addsimps[Union_Un_distrib];  | 
| 923 | 288  | 
|
289  | 
goal Set.thy "Union(A Int B) <= Union(A) Int Union(B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
290  | 
by (Fast_tac 1);  | 
| 923 | 291  | 
qed "Union_Int_subset";  | 
292  | 
||
293  | 
val prems = goal Set.thy  | 
|
294  | 
   "(Union(C) Int A = {}) = (! B:C. B Int A = {})";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
295  | 
by (fast_tac (!claset addSEs [equalityE]) 1);  | 
| 923 | 296  | 
qed "Union_disjoint";  | 
297  | 
||
| 1548 | 298  | 
section "Inter";  | 
299  | 
||
| 1531 | 300  | 
goal Set.thy "Inter({}) = UNIV";
 | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
301  | 
by (Fast_tac 1);  | 
| 1531 | 302  | 
qed "Inter_empty";  | 
303  | 
Addsimps[Inter_empty];  | 
|
304  | 
||
305  | 
goal Set.thy "Inter(UNIV) = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
306  | 
by (Fast_tac 1);  | 
| 1531 | 307  | 
qed "Inter_UNIV";  | 
308  | 
Addsimps[Inter_UNIV];  | 
|
309  | 
||
310  | 
goal Set.thy "Inter(insert a B) = a Int Inter(B)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
311  | 
by (Fast_tac 1);  | 
| 1531 | 312  | 
qed "Inter_insert";  | 
313  | 
Addsimps[Inter_insert];  | 
|
314  | 
||
| 
1564
 
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
 
paulson 
parents: 
1553 
diff
changeset
 | 
315  | 
goal Set.thy "Inter(A) Un Inter(B) <= Inter(A Int B)";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
316  | 
by (Fast_tac 1);  | 
| 
1564
 
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
 
paulson 
parents: 
1553 
diff
changeset
 | 
317  | 
qed "Inter_Un_subset";  | 
| 1531 | 318  | 
|
| 923 | 319  | 
goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)";  | 
| 
1786
 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 
berghofe 
parents: 
1763 
diff
changeset
 | 
320  | 
by (best_tac (!claset) 1);  | 
| 923 | 321  | 
qed "Inter_Un_distrib";  | 
322  | 
||
| 1548 | 323  | 
section "UN and INT";  | 
| 923 | 324  | 
|
325  | 
(*Basic identities*)  | 
|
326  | 
||
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
327  | 
goal Set.thy "(UN x:{}. B x) = {}";
 | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
328  | 
by (Fast_tac 1);  | 
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
329  | 
qed "UN_empty";  | 
| 1531 | 330  | 
Addsimps[UN_empty];  | 
331  | 
||
332  | 
goal Set.thy "(UN x:UNIV. B x) = (UN x. B x)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
333  | 
by (Fast_tac 1);  | 
| 1531 | 334  | 
qed "UN_UNIV";  | 
335  | 
Addsimps[UN_UNIV];  | 
|
336  | 
||
337  | 
goal Set.thy "(INT x:{}. B x) = UNIV";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
338  | 
by (Fast_tac 1);  | 
| 1531 | 339  | 
qed "INT_empty";  | 
340  | 
Addsimps[INT_empty];  | 
|
341  | 
||
342  | 
goal Set.thy "(INT x:UNIV. B x) = (INT x. B x)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
343  | 
by (Fast_tac 1);  | 
| 1531 | 344  | 
qed "INT_UNIV";  | 
345  | 
Addsimps[INT_UNIV];  | 
|
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
346  | 
|
| 
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
347  | 
goal Set.thy "(UN x:insert a A. B x) = B a Un UNION A B";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
348  | 
by (Fast_tac 1);  | 
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
349  | 
qed "UN_insert";  | 
| 1531 | 350  | 
Addsimps[UN_insert];  | 
351  | 
||
352  | 
goal Set.thy "(INT x:insert a A. B x) = B a Int INTER A B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
353  | 
by (Fast_tac 1);  | 
| 1531 | 354  | 
qed "INT_insert";  | 
355  | 
Addsimps[INT_insert];  | 
|
| 
1179
 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 
nipkow 
parents: 
923 
diff
changeset
 | 
356  | 
|
| 2021 | 357  | 
goal Set.thy  | 
358  | 
    "!!A. A~={} ==> (INT x:A. insert a (B x)) = insert a (INT x:A. B x)";
 | 
|
359  | 
by (Fast_tac 1);  | 
|
360  | 
qed "INT_insert_distrib";  | 
|
361  | 
||
362  | 
goal Set.thy "(INT x. insert a (B x)) = insert a (INT x. B x)";  | 
|
363  | 
by (Fast_tac 1);  | 
|
364  | 
qed "INT1_insert_distrib";  | 
|
365  | 
||
| 923 | 366  | 
goal Set.thy "Union(range(f)) = (UN x.f(x))";  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
367  | 
by (Fast_tac 1);  | 
| 923 | 368  | 
qed "Union_range_eq";  | 
369  | 
||
370  | 
goal Set.thy "Inter(range(f)) = (INT x.f(x))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
371  | 
by (Fast_tac 1);  | 
| 923 | 372  | 
qed "Inter_range_eq";  | 
373  | 
||
374  | 
goal Set.thy "Union(B``A) = (UN x:A. B(x))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
375  | 
by (Fast_tac 1);  | 
| 923 | 376  | 
qed "Union_image_eq";  | 
377  | 
||
378  | 
goal Set.thy "Inter(B``A) = (INT x:A. B(x))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
379  | 
by (Fast_tac 1);  | 
| 923 | 380  | 
qed "Inter_image_eq";  | 
381  | 
||
382  | 
goal Set.thy "!!A. a: A ==> (UN y:A. c) = c";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
383  | 
by (Fast_tac 1);  | 
| 923 | 384  | 
qed "UN_constant";  | 
385  | 
||
386  | 
goal Set.thy "!!A. a: A ==> (INT y:A. c) = c";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
387  | 
by (Fast_tac 1);  | 
| 923 | 388  | 
qed "INT_constant";  | 
389  | 
||
390  | 
goal Set.thy "(UN x.B) = B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
391  | 
by (Fast_tac 1);  | 
| 923 | 392  | 
qed "UN1_constant";  | 
| 1531 | 393  | 
Addsimps[UN1_constant];  | 
| 923 | 394  | 
|
395  | 
goal Set.thy "(INT x.B) = B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
396  | 
by (Fast_tac 1);  | 
| 923 | 397  | 
qed "INT1_constant";  | 
| 1531 | 398  | 
Addsimps[INT1_constant];  | 
| 923 | 399  | 
|
400  | 
goal Set.thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
401  | 
by (Fast_tac 1);  | 
| 923 | 402  | 
qed "UN_eq";  | 
403  | 
||
404  | 
(*Look: it has an EXISTENTIAL quantifier*)  | 
|
405  | 
goal Set.thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
406  | 
by (Fast_tac 1);  | 
| 923 | 407  | 
qed "INT_eq";  | 
408  | 
||
409  | 
(*Distributive laws...*)  | 
|
410  | 
||
411  | 
goal Set.thy "A Int Union(B) = (UN C:B. A Int C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
412  | 
by (Fast_tac 1);  | 
| 923 | 413  | 
qed "Int_Union";  | 
414  | 
||
415  | 
(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:  | 
|
416  | 
Union of a family of unions **)  | 
|
417  | 
goal Set.thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
418  | 
by (Fast_tac 1);  | 
| 923 | 419  | 
qed "Un_Union_image";  | 
420  | 
||
421  | 
(*Equivalent version*)  | 
|
422  | 
goal Set.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
423  | 
by (Fast_tac 1);  | 
| 923 | 424  | 
qed "UN_Un_distrib";  | 
425  | 
||
426  | 
goal Set.thy "A Un Inter(B) = (INT C:B. A Un C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
427  | 
by (Fast_tac 1);  | 
| 923 | 428  | 
qed "Un_Inter";  | 
429  | 
||
430  | 
goal Set.thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)";  | 
|
| 
1786
 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 
berghofe 
parents: 
1763 
diff
changeset
 | 
431  | 
by (best_tac (!claset) 1);  | 
| 923 | 432  | 
qed "Int_Inter_image";  | 
433  | 
||
434  | 
(*Equivalent version*)  | 
|
435  | 
goal Set.thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
436  | 
by (Fast_tac 1);  | 
| 923 | 437  | 
qed "INT_Int_distrib";  | 
438  | 
||
439  | 
(*Halmos, Naive Set Theory, page 35.*)  | 
|
440  | 
goal Set.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
441  | 
by (Fast_tac 1);  | 
| 923 | 442  | 
qed "Int_UN_distrib";  | 
443  | 
||
444  | 
goal Set.thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
445  | 
by (Fast_tac 1);  | 
| 923 | 446  | 
qed "Un_INT_distrib";  | 
447  | 
||
448  | 
goal Set.thy  | 
|
449  | 
"(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
450  | 
by (Fast_tac 1);  | 
| 923 | 451  | 
qed "Int_UN_distrib2";  | 
452  | 
||
453  | 
goal Set.thy  | 
|
454  | 
"(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
455  | 
by (Fast_tac 1);  | 
| 923 | 456  | 
qed "Un_INT_distrib2";  | 
457  | 
||
| 2512 | 458  | 
|
459  | 
section"Bounded quantifiers";  | 
|
460  | 
||
| 2519 | 461  | 
(** These are not added to the default simpset because (a) they duplicate the  | 
462  | 
body and (b) there are no similar rules for Int. **)  | 
|
| 2512 | 463  | 
|
| 2519 | 464  | 
goal Set.thy "(ALL x:A Un B.P x) = ((ALL x:A.P x) & (ALL x:B.P x))";  | 
| 2512 | 465  | 
by (Fast_tac 1);  | 
| 2519 | 466  | 
qed "ball_Un";  | 
467  | 
||
468  | 
goal Set.thy "(EX x:A Un B.P x) = ((EX x:A.P x) | (EX x:B.P x))";  | 
|
469  | 
by (Fast_tac 1);  | 
|
470  | 
qed "bex_Un";  | 
|
| 2512 | 471  | 
|
472  | 
||
| 1548 | 473  | 
section "-";  | 
| 923 | 474  | 
|
475  | 
goal Set.thy "A-A = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
476  | 
by (Fast_tac 1);  | 
| 923 | 477  | 
qed "Diff_cancel";  | 
| 1531 | 478  | 
Addsimps[Diff_cancel];  | 
| 923 | 479  | 
|
480  | 
goal Set.thy "{}-A = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
481  | 
by (Fast_tac 1);  | 
| 923 | 482  | 
qed "empty_Diff";  | 
| 1531 | 483  | 
Addsimps[empty_Diff];  | 
| 923 | 484  | 
|
485  | 
goal Set.thy "A-{} = A";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
486  | 
by (Fast_tac 1);  | 
| 923 | 487  | 
qed "Diff_empty";  | 
| 1531 | 488  | 
Addsimps[Diff_empty];  | 
489  | 
||
490  | 
goal Set.thy "A-UNIV = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
491  | 
by (Fast_tac 1);  | 
| 1531 | 492  | 
qed "Diff_UNIV";  | 
493  | 
Addsimps[Diff_UNIV];  | 
|
494  | 
||
495  | 
goal Set.thy "!!x. x~:A ==> A - insert x B = A-B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
496  | 
by (Fast_tac 1);  | 
| 1531 | 497  | 
qed "Diff_insert0";  | 
498  | 
Addsimps [Diff_insert0];  | 
|
| 923 | 499  | 
|
500  | 
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
 | 
|
501  | 
goal Set.thy "A - insert a B = A - B - {a}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
502  | 
by (Fast_tac 1);  | 
| 923 | 503  | 
qed "Diff_insert";  | 
504  | 
||
505  | 
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
 | 
|
506  | 
goal Set.thy "A - insert a B = A - {a} - B";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
507  | 
by (Fast_tac 1);  | 
| 923 | 508  | 
qed "Diff_insert2";  | 
509  | 
||
| 1531 | 510  | 
goal Set.thy "insert x A - B = (if x:B then A-B else insert x (A-B))";  | 
| 1553 | 511  | 
by (simp_tac (!simpset setloop split_tac[expand_if]) 1);  | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
512  | 
by (Fast_tac 1);  | 
| 1531 | 513  | 
qed "insert_Diff_if";  | 
514  | 
||
515  | 
goal Set.thy "!!x. x:B ==> insert x A - B = A-B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
516  | 
by (Fast_tac 1);  | 
| 1531 | 517  | 
qed "insert_Diff1";  | 
518  | 
Addsimps [insert_Diff1];  | 
|
519  | 
||
| 923 | 520  | 
val prems = goal Set.thy "a:A ==> insert a (A-{a}) = A";
 | 
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
521  | 
by (fast_tac (!claset addSIs prems) 1);  | 
| 923 | 522  | 
qed "insert_Diff";  | 
523  | 
||
524  | 
goal Set.thy "A Int (B-A) = {}";
 | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
525  | 
by (Fast_tac 1);  | 
| 923 | 526  | 
qed "Diff_disjoint";  | 
| 1531 | 527  | 
Addsimps[Diff_disjoint];  | 
| 923 | 528  | 
|
529  | 
goal Set.thy "!!A. A<=B ==> A Un (B-A) = B";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
530  | 
by (Fast_tac 1);  | 
| 923 | 531  | 
qed "Diff_partition";  | 
532  | 
||
533  | 
goal Set.thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
534  | 
by (Fast_tac 1);  | 
| 923 | 535  | 
qed "double_diff";  | 
536  | 
||
537  | 
goal Set.thy "A - (B Un C) = (A-B) Int (A-C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
538  | 
by (Fast_tac 1);  | 
| 923 | 539  | 
qed "Diff_Un";  | 
540  | 
||
541  | 
goal Set.thy "A - (B Int C) = (A-B) Un (A-C)";  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1748 
diff
changeset
 | 
542  | 
by (Fast_tac 1);  | 
| 923 | 543  | 
qed "Diff_Int";  | 
544  | 
||
| 1531 | 545  | 
Addsimps[subset_UNIV, empty_subsetI, subset_refl];  | 
| 2021 | 546  | 
|
547  | 
||
548  | 
(** Miniscoping: pushing in big Unions and Intersections **)  | 
|
549  | 
local  | 
|
550  | 
fun prover s = prove_goal Set.thy s (fn _ => [Fast_tac 1])  | 
|
551  | 
in  | 
|
552  | 
val UN1_simps = map prover  | 
|
| 2031 | 553  | 
["(UN x. insert a (B x)) = insert a (UN x. B x)",  | 
554  | 
"(UN x. A x Int B) = ((UN x.A x) Int B)",  | 
|
555  | 
"(UN x. A Int B x) = (A Int (UN x.B x))",  | 
|
556  | 
"(UN x. A x Un B) = ((UN x.A x) Un B)",  | 
|
557  | 
"(UN x. A Un B x) = (A Un (UN x.B x))",  | 
|
558  | 
"(UN x. A x - B) = ((UN x.A x) - B)",  | 
|
559  | 
"(UN x. A - B x) = (A - (INT x.B x))"];  | 
|
| 2021 | 560  | 
|
561  | 
val INT1_simps = map prover  | 
|
| 2031 | 562  | 
["(INT x. insert a (B x)) = insert a (INT x. B x)",  | 
563  | 
"(INT x. A x Int B) = ((INT x.A x) Int B)",  | 
|
564  | 
"(INT x. A Int B x) = (A Int (INT x.B x))",  | 
|
565  | 
"(INT x. A x Un B) = ((INT x.A x) Un B)",  | 
|
566  | 
"(INT x. A Un B x) = (A Un (INT x.B x))",  | 
|
567  | 
"(INT x. A x - B) = ((INT x.A x) - B)",  | 
|
568  | 
"(INT x. A - B x) = (A - (UN x.B x))"];  | 
|
| 2021 | 569  | 
|
| 
2513
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
570  | 
val UN_simps = map prover  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
571  | 
["(UN x:C. A x Int B) = ((UN x:C.A x) Int B)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
572  | 
"(UN x:C. A Int B x) = (A Int (UN x:C.B x))",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
573  | 
"(UN x:C. A x - B) = ((UN x:C.A x) - B)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
574  | 
"(UN x:C. A - B x) = (A - (INT x:C.B x))"];  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
575  | 
|
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
576  | 
val INT_simps = map prover  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
577  | 
["(INT x:C. insert a (B x)) = insert a (INT x:C. B x)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
578  | 
"(INT x:C. A x Un B) = ((INT x:C.A x) Un B)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
579  | 
"(INT x:C. A Un B x) = (A Un (INT x:C.B x))"];  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
580  | 
|
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
581  | 
(*The missing laws for bounded Unions and Intersections are conditional  | 
| 2021 | 582  | 
on the index set's being non-empty. Thus they are probably NOT worth  | 
583  | 
adding as default rewrites.*)  | 
|
| 
2513
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
584  | 
|
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
585  | 
val ball_simps = map prover  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
586  | 
["(ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
587  | 
"(ALL x:A. P | Q x) = (P | (ALL x:A. Q x))",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
588  | 
     "(ALL x:{}. P x) = True",
 | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
589  | 
"(ALL x:insert a B. P x) = (P(a) & (ALL x:B. P x))",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
590  | 
"(ALL x:Union(A). P x) = (ALL y:A. ALL x:y. P x)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
591  | 
"(ALL x:Collect Q. P x) = (ALL x. Q x --> P x)"];  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
592  | 
|
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
593  | 
val ball_conj_distrib =  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
594  | 
prover "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))";  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
595  | 
|
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
596  | 
val bex_simps = map prover  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
597  | 
["(EX x:A. P x & Q) = ((EX x:A. P x) & Q)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
598  | 
"(EX x:A. P & Q x) = (P & (EX x:A. Q x))",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
599  | 
     "(EX x:{}. P x) = False",
 | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
600  | 
"(EX x:insert a B. P x) = (P(a) | (EX x:B. P x))",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
601  | 
"(EX x:Union(A). P x) = (EX y:A. EX x:y. P x)",  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
602  | 
"(EX x:Collect Q. P x) = (EX x. Q x & P x)"];  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
603  | 
|
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
604  | 
val bex_conj_distrib =  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
605  | 
prover "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))";  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
606  | 
|
| 2021 | 607  | 
end;  | 
608  | 
||
| 
2513
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
609  | 
Addsimps (UN1_simps @ INT1_simps @ UN_simps @ INT_simps @  | 
| 
 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 
paulson 
parents: 
2512 
diff
changeset
 | 
610  | 
ball_simps @ bex_simps);  |